diff --git a/src/Init/Core.lean b/src/Init/Core.lean index 0db8d81ea5c5..85e342e10ccf 100644 --- a/src/Init/Core.lean +++ b/src/Init/Core.lean @@ -1053,7 +1053,7 @@ set_option linter.missingDocs false in /-- Similar to `decide`, but uses an explicit instance -/ @[inline] def toBoolUsing {p : Prop} (d : Decidable p) : Bool := - decide (h := d) + @decide p d theorem toBoolUsing_eq_true {p : Prop} (d : Decidable p) (h : p) : toBoolUsing d = true := decide_eq_true (inst := d) h @@ -1117,34 +1117,34 @@ end Decidable section variable {p q : Prop} /-- Transfer a decidability proof across an equivalence of propositions. -/ -@[inline] def decidable_of_decidable_of_iff [Decidable p] (h : p ↔ q) : Decidable q := - if hp : p then - isTrue (Iff.mp h hp) - else - isFalse fun hq => absurd (Iff.mpr h hq) hp +@[inline] def decidable_of_decidable_of_iff [dp : Decidable p] (h : p ↔ q) : Decidable q where + decide := decide p + of_decide := + match dp with + | isTrue hp => Iff.mp h hp + | isFalse hp => fun hq => absurd (Iff.mpr h hq) hp /-- Transfer a decidability proof across an equality of propositions. -/ @[inline] def decidable_of_decidable_of_eq [Decidable p] (h : p = q) : Decidable q := decidable_of_decidable_of_iff (p := p) (h ▸ Iff.rfl) end -@[macro_inline] instance {p q} [Decidable p] [Decidable q] : Decidable (p → q) := - if hp : p then - if hq : q then isTrue (fun _ => hq) - else isFalse (fun h => absurd (h hp) hq) - else isTrue (fun h => absurd h hp) - -instance {p q} [Decidable p] [Decidable q] : Decidable (p ↔ q) := - if hp : p then - if hq : q then - isTrue ⟨fun _ => hq, fun _ => hp⟩ - else - isFalse fun h => hq (h.1 hp) - else - if hq : q then - isFalse fun h => hp (h.2 hq) - else - isTrue ⟨fun h => absurd h hp, fun h => absurd h hq⟩ +@[macro_inline] instance {p q} [dp : Decidable p] [dq : Decidable q] : Decidable (p → q) where + decide := !p || q + of_decide := + match dp, dq with + | isTrue _, isTrue hq => fun _ => hq + | isTrue hp, isFalse hq => fun h => absurd (h hp) hq + | isFalse hp, _ => fun h => absurd h hp + +instance {p q} [dp : Decidable p] [dq : Decidable q] : Decidable (p ↔ q) where + decide := (decide p).beq (decide q) + of_decide := + match dp, dq with + | isTrue hp, isTrue hq => ⟨fun _ => hq, fun _ => hp⟩ + | isTrue hp, isFalse hq => fun h => hq (h.1 hp) + | isFalse hp, isTrue hq => fun h => hp (h.2 hq) + | isFalse hp, isFalse hq => ⟨fun h => absurd h hp, fun h => absurd h hq⟩ /-! # if-then-else expression theorems -/ @@ -1192,18 +1192,16 @@ instance {c : Prop} {t : c → Prop} {e : ¬c → Prop} [dC : Decidable c] [dT : | isFalse hc => dE hc /-- Auxiliary definition for generating compact `noConfusion` for enumeration types -/ -abbrev noConfusionTypeEnum {α : Sort u} {β : Sort v} [inst : DecidableEq β] (f : α → β) (P : Sort w) (x y : α) : Sort w := - (inst (f x) (f y)).casesOn - (fun _ => P) - (fun _ => P → P) +abbrev noConfusionTypeEnum {α : Sort u} (f : α → Nat) (P : Sort w) (x y : α) : Sort w := + ((f x).beq (f y)).casesOn P (P → P) /-- Auxiliary definition for generating compact `noConfusion` for enumeration types -/ -abbrev noConfusionEnum {α : Sort u} {β : Sort v} [inst : DecidableEq β] (f : α → β) {P : Sort w} {x y : α} (h : x = y) : noConfusionTypeEnum f P x y := - Decidable.casesOn - (motive := fun (inst : Decidable (f x = f y)) => Decidable.casesOn (motive := fun _ => Sort w) inst (fun _ => P) (fun _ => P → P)) - (inst (f x) (f y)) - (fun h' => False.elim (h' (congrArg f h))) - (fun _ => fun x => x) +abbrev noConfusionEnum {α : Sort u} (f : α → Nat) {P : Sort w} {x y : α} (h : x = y) : noConfusionTypeEnum f P x y := + ((f x).beq (f y)).casesOn + (motive := fun b => (f x).beq (f y) = b → b.casesOn P (P → P)) + (fun h' => False.elim (Nat.ne_of_beq_eq_false h' (congrArg f h))) + (fun _ a => a) + rfl /-! # Inhabited -/ @@ -1271,7 +1269,7 @@ theorem recSubsingleton {h₂ : ¬p → Sort u} [h₃ : ∀ (h : p), Subsingleton (h₁ h)] [h₄ : ∀ (h : ¬p), Subsingleton (h₂ h)] - : Subsingleton (h.casesOn h₂ h₁) := + : Subsingleton (h.falseTrueCases h₂ h₁) := match h with | isTrue h => h₃ h | isFalse h => h₄ h diff --git a/src/Init/Data/List/Basic.lean b/src/Init/Data/List/Basic.lean index 13fee0b134f2..453c44a517de 100644 --- a/src/Init/Data/List/Basic.lean +++ b/src/Init/Data/List/Basic.lean @@ -117,15 +117,6 @@ theorem of_concat_eq_concat {as bs : List α} {a b : α} (h : as.concat a = bs.c /-! ## Equality -/ -/-- -Checks whether two lists have the same length and their elements are pairwise `BEq`. Normally used -via the `==` operator. --/ -protected def beq [BEq α] : List α → List α → Bool - | [], [] => true - | a::as, b::bs => a == b && List.beq as bs - | _, _ => false - @[simp] theorem beq_nil_nil [BEq α] : List.beq ([] : List α) ([] : List α) = true := rfl @[simp] theorem beq_cons_nil [BEq α] {a : α} {as : List α} : List.beq (a::as) [] = false := rfl @[simp] theorem beq_nil_cons [BEq α] {a : α} {as : List α} : List.beq [] (a::as) = false := rfl diff --git a/src/Init/Data/Vector/Extract.lean b/src/Init/Data/Vector/Extract.lean index 93bc8e14ccd0..cdf2be2dd41f 100644 --- a/src/Init/Data/Vector/Extract.lean +++ b/src/Init/Data/Vector/Extract.lean @@ -9,6 +9,8 @@ prelude import Init.Data.Vector.Lemmas import Init.Data.Array.Extract +set_option maxHeartbeats 20000000 + /-! # Lemmas about `Vector.extract` -/ diff --git a/src/Init/Prelude.lean b/src/Init/Prelude.lean index 8ed2a9b3455c..53ccce525388 100644 --- a/src/Init/Prelude.lean +++ b/src/Init/Prelude.lean @@ -886,6 +886,96 @@ theorem ULift.up_down {α : Type u} (b : ULift.{v} α) : Eq (up (down b)) b := r /-- Bijection between `α` and `ULift.{v} α` -/ theorem ULift.down_up {α : Type u} (a : α) : Eq (down (up.{v} a)) a := rfl +/-! # Boolean operators -/ + +/-- +The conditional function. + +`cond c x y` is the same as `if c then x else y`, but optimized for a Boolean condition rather than +a decidable proposition. It can also be written using the notation `bif c then x else y`. + +Just like `ite`, `cond` is declared `@[macro_inline]`, which causes applications of `cond` to be +unfolded. As a result, `x` and `y` are not evaluated at runtime until one of them is selected, and +only the selected branch is evaluated. +-/ +@[macro_inline] def cond {α : Sort u} (c : Bool) (x y : α) : α := + match c with + | true => x + | false => y + +/-- +The dependent conditional function, in which each branch is provided with a local assumption about +the condition's value. This allows the value to be used in proofs as well as for control flow. + +`dcond c (fun h => x) (fun h => y)` is the same as `if h : c then x else y`, but optimized for a +Boolean condition rather than a decidable proposition. Unlike the non-dependent version `cond`, +there is no special notation for `dcond`. + +Just like `ite`, `dite`, and `cond`, `dcond` is declared `@[macro_inline]`, which causes +applications of `dcond` to be unfolded. As a result, `x` and `y` are not evaluated at runtime until +one of them is selected, and only the selected branch is evaluated. `dcond` is intended for +metaprogramming use, rather than for use in verified programs, so behavioral lemmas are not +provided. +-/ +@[macro_inline] +protected def Bool.dcond {α : Sort u} (c : Bool) (x : Eq c true → α) (y : Eq c false → α) : α := + match c with + | true => x rfl + | false => y rfl + +/-- +Boolean “or”, also known as disjunction. `or x y` can be written `x || y`. + +The corresponding propositional connective is `Or : Prop → Prop → Prop`, written with the `∨` +operator. + +The Boolean `or` is a `@[macro_inline]` function in order to give it short-circuiting evaluation: +if `x` is `true` then `y` is not evaluated at runtime. +-/ +@[macro_inline] def Bool.or (x y : Bool) : Bool := + match x with + | true => true + | false => y + +/-- +Boolean “and”, also known as conjunction. `and x y` can be written `x && y`. + +The corresponding propositional connective is `And : Prop → Prop → Prop`, written with the `∧` +operator. + +The Boolean `and` is a `@[macro_inline]` function in order to give it short-circuiting evaluation: +if `x` is `false` then `y` is not evaluated at runtime. +-/ +@[macro_inline] def Bool.and (x y : Bool) : Bool := + match x with + | false => false + | true => y + +/-- +Boolean negation, also known as Boolean complement. `not x` can be written `!x`. + +This is a function that maps the value `true` to `false` and the value `false` to `true`. The +propositional connective is `Not : Prop → Prop`. +-/ +@[inline] def Bool.not : Bool → Bool + | true => false + | false => true + +export Bool (or and not) + +/-- +Boolean equality. This should usually be used via the `a == b` notation. + +This function returns `true` if the booleans are equal or `false` otherwise. The propositional +connective is `Iff : Prop → Prop → Prop`, written with the `↔` operator. + +This function comes with two definional equalities: `(a == true) = a` and `(a == false) = !a`. +-/ +@[inline] def Bool.beq (a b : Bool) : Bool := + match b with + | false => not a + | true => a + /-- Either a proof that `p` is true or a proof that `p` is false. This is equivalent to a `Bool` paired with a proof that the `Bool` is `true` if and only if `p` is true. @@ -900,20 +990,38 @@ Because `Decidable` carries data, when writing `@[simp]` lemmas which include a on the LHS, it is best to use `{_ : Decidable p}` rather than `[Decidable p]` so that non-canonical instances can be found via unification rather than instance synthesis. -/ -class inductive Decidable (p : Prop) where - /-- Proves that `p` is decidable by supplying a proof of `¬p` -/ - | isFalse (h : Not p) : Decidable p - /-- Proves that `p` is decidable by supplying a proof of `p` -/ - | isTrue (h : p) : Decidable p +class Decidable (p : Prop) where + /-- + Introduces `Decidable p` using a boolean `decide` and a proof of `¬p` if `decide` is `false` + or `p` if `decide` is `true`. + -/ + intro :: + /-- + Converts a decidable proposition into a `Bool`. -/-- -Converts a decidable proposition into a `Bool`. + If `p : Prop` is decidable, then `decide p : Bool` is the Boolean value + that is `true` if `p` is true and `false` if `p` is false. + -/ + decide : Bool + /-- + A proof of `¬p` or `p`, depending on the value of `decide`. + -/ + of_decide (p) : cond decide p (Not p) + +/-- Proves that `p` is decidable by supplying a proof of `¬p` -/ +@[match_pattern] abbrev Decidable.isFalse (h : Not p) : Decidable p := ⟨false, h⟩ + +/-- Proves that `p` is decidable by supplying a proof of `p` -/ +@[match_pattern] abbrev Decidable.isTrue (h : p) : Decidable p := ⟨true, h⟩ -If `p : Prop` is decidable, then `decide p : Bool` is the Boolean value -that is `true` if `p` is true and `false` if `p` is false. +/-- +The default eliminator for `Decidable`, using `isFalse` and `isTrue`. -/ -@[inline_if_reduce, nospecialize] def Decidable.decide (p : Prop) [h : Decidable p] : Bool := - h.casesOn (fun _ => false) (fun _ => true) +@[cases_eliminator, induction_eliminator, elab_as_elim] +def Decidable.falseTrueCases {p : Prop} {motive : Decidable p → Sort u} + (isFalse : ∀ h, motive (isFalse h)) (isTrue : ∀ h, motive (isTrue h)) + (t : Decidable p) : motive t := + t.casesOn fun x => x.casesOn isFalse isTrue export Decidable (isTrue isFalse decide) @@ -972,21 +1080,48 @@ theorem of_decide_eq_self_eq_true [inst : DecidableEq α] (a : α) : Eq (decide | isTrue _ => rfl | isFalse h₁ => absurd rfl h₁ +@[macro_inline] instance {p q} [dp : Decidable p] [dq : Decidable q] : Decidable (And p q) where + decide := and (decide p) (decide q) + of_decide := + match dp, dq with + | isTrue hp, isTrue hq => ⟨hp, hq⟩ + | isTrue _, isFalse hq => fun h => hq h.2 + | isFalse hp, _ => fun h => hp (And.left h) + +@[macro_inline] instance [dp : Decidable p] [dq : Decidable q] : Decidable (Or p q) where + decide := or (decide p) (decide q) + of_decide := + match dp, dq with + | isTrue hp, _ => .inl hp + | isFalse _, isTrue hq => .inr hq + | isFalse hp, isFalse hq => fun + | Or.inl h => hp h + | Or.inr h => hq h + +@[inline] instance [dp : Decidable p] : Decidable (Not p) where + decide := not (decide p) + of_decide := + match dp with + | isTrue hp => absurd hp + | isFalse hp => hp + /-- Decides whether two Booleans are equal. This function should normally be called via the `DecidableEq Bool` instance that it exists to support. -/ -@[inline] def Bool.decEq (a b : Bool) : Decidable (Eq a b) := - match a, b with - | false, false => isTrue rfl - | false, true => isFalse (fun h => Bool.noConfusion h) - | true, false => isFalse (fun h => Bool.noConfusion h) - | true, true => isTrue rfl +@[inline] def Bool.decEq (a b : Bool) : Decidable (Eq a b) where + decide := a.beq b + of_decide := + match a, b with + | false, false => rfl + | false, true => Bool.noConfusion + | true, false => Bool.noConfusion + | true, true => rfl @[inline] instance : DecidableEq Bool := - Bool.decEq + Bool.decEq /-- `BEq α` is a typeclass for supplying a boolean-valued equality relation on @@ -1024,7 +1159,7 @@ lifted the check into an explicit `if`, but we could also use this proof multipl or derive `i < arr.size` from some other proposition that we are checking in the `if`.) -/ @[macro_inline] def dite {α : Sort u} (c : Prop) [h : Decidable c] (t : c → α) (e : Not c → α) : α := - h.casesOn e t + h.decide.casesOn e t h.of_decide /-! # if-then-else -/ @@ -1049,110 +1184,7 @@ the expected "lazy" behavior of `if`: the `t` and `e` arguments delay evaluation until `c` is known. -/ @[macro_inline] def ite {α : Sort u} (c : Prop) [h : Decidable c] (t e : α) : α := - h.casesOn (fun _ => e) (fun _ => t) - -@[macro_inline] instance {p q} [dp : Decidable p] [dq : Decidable q] : Decidable (And p q) := - match dp with - | isTrue hp => - match dq with - | isTrue hq => isTrue ⟨hp, hq⟩ - | isFalse hq => isFalse (fun h => hq (And.right h)) - | isFalse hp => - isFalse (fun h => hp (And.left h)) - -@[macro_inline] instance [dp : Decidable p] [dq : Decidable q] : Decidable (Or p q) := - match dp with - | isTrue hp => isTrue (Or.inl hp) - | isFalse hp => - match dq with - | isTrue hq => isTrue (Or.inr hq) - | isFalse hq => - isFalse fun h => match h with - | Or.inl h => hp h - | Or.inr h => hq h - -instance [dp : Decidable p] : Decidable (Not p) := - match dp with - | isTrue hp => isFalse (absurd hp) - | isFalse hp => isTrue hp - -/-! # Boolean operators -/ - -/-- -The conditional function. - -`cond c x y` is the same as `if c then x else y`, but optimized for a Boolean condition rather than -a decidable proposition. It can also be written using the notation `bif c then x else y`. - -Just like `ite`, `cond` is declared `@[macro_inline]`, which causes applications of `cond` to be -unfolded. As a result, `x` and `y` are not evaluated at runtime until one of them is selected, and -only the selected branch is evaluated. --/ -@[macro_inline] def cond {α : Sort u} (c : Bool) (x y : α) : α := - match c with - | true => x - | false => y - - -/-- -The dependent conditional function, in which each branch is provided with a local assumption about -the condition's value. This allows the value to be used in proofs as well as for control flow. - -`dcond c (fun h => x) (fun h => y)` is the same as `if h : c then x else y`, but optimized for a -Boolean condition rather than a decidable proposition. Unlike the non-dependent version `cond`, -there is no special notation for `dcond`. - -Just like `ite`, `dite`, and `cond`, `dcond` is declared `@[macro_inline]`, which causes -applications of `dcond` to be unfolded. As a result, `x` and `y` are not evaluated at runtime until -one of them is selected, and only the selected branch is evaluated. `dcond` is intended for -metaprogramming use, rather than for use in verified programs, so behavioral lemmas are not -provided. --/ -@[macro_inline] -protected def Bool.dcond {α : Sort u} (c : Bool) (x : Eq c true → α) (y : Eq c false → α) : α := - match c with - | true => x rfl - | false => y rfl - -/-- -Boolean “or”, also known as disjunction. `or x y` can be written `x || y`. - -The corresponding propositional connective is `Or : Prop → Prop → Prop`, written with the `∨` -operator. - -The Boolean `or` is a `@[macro_inline]` function in order to give it short-circuiting evaluation: -if `x` is `true` then `y` is not evaluated at runtime. --/ -@[macro_inline] def Bool.or (x y : Bool) : Bool := - match x with - | true => true - | false => y - -/-- -Boolean “and”, also known as conjunction. `and x y` can be written `x && y`. - -The corresponding propositional connective is `And : Prop → Prop → Prop`, written with the `∧` -operator. - -The Boolean `and` is a `@[macro_inline]` function in order to give it short-circuiting evaluation: -if `x` is `false` then `y` is not evaluated at runtime. --/ -@[macro_inline] def Bool.and (x y : Bool) : Bool := - match x with - | false => false - | true => y - -/-- -Boolean negation, also known as Boolean complement. `not x` can be written `!x`. - -This is a function that maps the value `true` to `false` and the value `false` to `true`. The -propositional connective is `Not : Prop → Prop`. --/ -@[inline] def Bool.not : Bool → Bool - | true => false - | false => true - -export Bool (or and not) + dite c (fun _ => t) (fun _ => e) /-- The natural numbers, starting at zero. @@ -1727,10 +1759,12 @@ Examples: * `show 12 = 12 by decide` -/ @[reducible, extern "lean_nat_dec_eq"] -protected def Nat.decEq (n m : @& Nat) : Decidable (Eq n m) := - match h:beq n m with - | true => isTrue (eq_of_beq_eq_true h) - | false => isFalse (ne_of_beq_eq_false h) +protected def Nat.decEq (n m : @& Nat) : Decidable (Eq n m) where + decide := beq n m + of_decide := + match h : beq n m with + | true => eq_of_beq_eq_true h + | false => ne_of_beq_eq_false h @[inline] instance : DecidableEq Nat := Nat.decEq @@ -1915,8 +1949,12 @@ Examples: * `show 5 ≤ 12 by decide` -/ @[extern "lean_nat_dec_le"] -instance Nat.decLe (n m : @& Nat) : Decidable (LE.le n m) := - dite (Eq (Nat.ble n m) true) (fun h => isTrue (Nat.le_of_ble_eq_true h)) (fun h => isFalse (Nat.not_le_of_not_ble_eq_true h)) +instance Nat.decLe (n m : @& Nat) : Decidable (LE.le n m) where + decide := ble n m + of_decide := + match h : ble n m with + | true => Nat.le_of_ble_eq_true h + | false => Nat.not_le_of_not_ble_eq_true (fun h' => Bool.noConfusion (h' ▸ h :)) /-- A decision procedure for strict inequality of natural numbers, usually accessed via the @@ -2009,11 +2047,13 @@ theorem Fin.eq_of_val_eq {n} : ∀ {i j : Fin n}, Eq i.val j.val → Eq i j theorem Fin.val_eq_of_eq {n} {i j : Fin n} (h : Eq i j) : Eq i.val j.val := h ▸ rfl -instance (n : Nat) : DecidableEq (Fin n) := - fun i j => - match decEq i.val j.val with - | isTrue h => isTrue (Fin.eq_of_val_eq h) - | isFalse h => isFalse (fun h' => absurd (Fin.val_eq_of_eq h') h) +instance (n : Nat) : DecidableEq (Fin n) := fun i j => { + decide := decide (Eq i.val j.val) + of_decide := + match instDecidableEqNat i.val j.val with + | isTrue h => Fin.eq_of_val_eq h + | isFalse h => fun h' => absurd (Fin.val_eq_of_eq h') h +} instance {n} : LT (Fin n) where lt a b := LT.lt a.val b.val @@ -2046,12 +2086,12 @@ This should be used via the instance `DecidableEq (BitVec w)`. -- We manually derive the `DecidableEq` instances for `BitVec` because -- we want to have builtin support for bit-vector literals, and we -- need a name for this function to implement `canUnfoldAtMatcher` at `WHNF.lean`. -def BitVec.decEq (x y : BitVec w) : Decidable (Eq x y) := - match x, y with - | ⟨n⟩, ⟨m⟩ => - dite (Eq n m) - (fun h => isTrue (h ▸ rfl)) - (fun h => isFalse (fun h' => BitVec.noConfusion h' (fun h' => absurd h' h))) +def BitVec.decEq (x y : BitVec w) : Decidable (Eq x y) where + decide := decide (Eq x.toFin y.toFin) + of_decide := + match instDecidableEqFin _ x.toFin y.toFin with + | isTrue h => congrArg BitVec.ofFin h + | isFalse h => fun h' => BitVec.noConfusion h' h instance : DecidableEq (BitVec w) := BitVec.decEq @@ -2120,12 +2160,12 @@ Examples: * `show (7 : UInt8) = 7 by decide` -/ @[extern "lean_uint8_dec_eq"] -def UInt8.decEq (a b : UInt8) : Decidable (Eq a b) := - match a, b with - | ⟨n⟩, ⟨m⟩ => - dite (Eq n m) - (fun h => isTrue (h ▸ rfl)) - (fun h => isFalse (fun h' => UInt8.noConfusion h' (fun h' => absurd h' h))) +def UInt8.decEq (a b : UInt8) : Decidable (Eq a b) where + decide := decide (Eq a.toBitVec b.toBitVec) + of_decide := + match instDecidableEqBitVec a.toBitVec b.toBitVec with + | isTrue h => congrArg UInt8.ofBitVec h + | isFalse h => fun h' => UInt8.noConfusion h' h instance : DecidableEq UInt8 := UInt8.decEq @@ -2178,12 +2218,12 @@ Examples: * `show (7 : UInt16) = 7 by decide` -/ @[extern "lean_uint16_dec_eq"] -def UInt16.decEq (a b : UInt16) : Decidable (Eq a b) := - match a, b with - | ⟨n⟩, ⟨m⟩ => - dite (Eq n m) - (fun h => isTrue (h ▸ rfl)) - (fun h => isFalse (fun h' => UInt16.noConfusion h' (fun h' => absurd h' h))) +def UInt16.decEq (a b : UInt16) : Decidable (Eq a b) where + decide := decide (Eq a.toBitVec b.toBitVec) + of_decide := + match instDecidableEqBitVec a.toBitVec b.toBitVec with + | isTrue h => congrArg UInt16.ofBitVec h + | isFalse h => fun h' => UInt16.noConfusion h' h instance : DecidableEq UInt16 := UInt16.decEq @@ -2243,10 +2283,12 @@ Examples: * `show (7 : UInt32) = 7 by decide` -/ @[extern "lean_uint32_dec_eq"] -def UInt32.decEq (a b : UInt32) : Decidable (Eq a b) := - match a, b with - | ⟨n⟩, ⟨m⟩ => - dite (Eq n m) (fun h => isTrue (h ▸ rfl)) (fun h => isFalse (fun h' => UInt32.noConfusion h' (fun h' => absurd h' h))) +def UInt32.decEq (a b : UInt32) : Decidable (Eq a b) where + decide := decide (Eq a.toBitVec b.toBitVec) + of_decide := + match instDecidableEqBitVec a.toBitVec b.toBitVec with + | isTrue h => congrArg UInt32.ofBitVec h + | isFalse h => fun h' => UInt32.noConfusion h' h instance : DecidableEq UInt32 := UInt32.decEq @@ -2341,12 +2383,12 @@ Examples: * `show (7 : UInt64) = 7 by decide` -/ @[extern "lean_uint64_dec_eq"] -def UInt64.decEq (a b : UInt64) : Decidable (Eq a b) := - match a, b with - | ⟨n⟩, ⟨m⟩ => - dite (Eq n m) - (fun h => isTrue (h ▸ rfl)) - (fun h => isFalse (fun h' => UInt64.noConfusion h' (fun h' => absurd h' h))) +def UInt64.decEq (a b : UInt64) : Decidable (Eq a b) where + decide := decide (Eq a.toBitVec b.toBitVec) + of_decide := + match instDecidableEqBitVec a.toBitVec b.toBitVec with + | isTrue h => congrArg UInt64.ofBitVec h + | isFalse h => fun h' => UInt64.noConfusion h' h instance : DecidableEq UInt64 := UInt64.decEq @@ -2411,12 +2453,12 @@ Examples: * `show (7 : USize) = 7 by decide` -/ @[extern "lean_usize_dec_eq"] -def USize.decEq (a b : USize) : Decidable (Eq a b) := - match a, b with - | ⟨n⟩, ⟨m⟩ => - dite (Eq n m) - (fun h => isTrue (h ▸ rfl)) - (fun h => isFalse (fun h' => USize.noConfusion h' (fun h' => absurd h' h))) +def USize.decEq (a b : USize) : Decidable (Eq a b) where + decide := decide (Eq a.toBitVec b.toBitVec) + of_decide := + match instDecidableEqBitVec a.toBitVec b.toBitVec with + | isTrue h => congrArg USize.ofBitVec h + | isFalse h => fun h' => USize.noConfusion h' h instance : DecidableEq USize := USize.decEq @@ -2481,11 +2523,13 @@ theorem Char.ne_of_val_ne {c d : Char} (h : Not (Eq c.val d.val)) : Not (Eq c d) theorem Char.val_ne_of_ne {c d : Char} (h : Not (Eq c d)) : Not (Eq c.val d.val) := fun h' => absurd (eq_of_val_eq h') h -instance : DecidableEq Char := - fun c d => - match decEq c.val d.val with - | isTrue h => isTrue (Char.eq_of_val_eq h) - | isFalse h => isFalse (Char.ne_of_val_ne h) +instance : DecidableEq Char := fun c d => { + decide := decide (Eq c.val d.val) + of_decide := + match instDecidableEqUInt32 c.val d.val with + | isTrue h => Char.eq_of_val_eq h + | isFalse h => Char.ne_of_val_ne h +} /-- Returns the number of bytes required to encode this `Char` in UTF-8. -/ def Char.utf8Size (c : Char) : Nat := @@ -2566,18 +2610,32 @@ inductive List (α : Type u) where instance {α} : Inhabited (List α) where default := List.nil +/-- +Checks whether two lists have the same length and their elements are pairwise `BEq`. Normally used +via the `==` operator. +-/ +protected def List.beq {α : Type u} [BEq α] : (a b : List α) → Bool + | nil, nil => true + | cons _ _, nil => false + | nil, cons _ _ => false + | cons a as, cons b bs => and (beq a b) (List.beq as bs) + +set_option pp.notation false + /-- Implements decidable equality for `List α`, assuming `α` has decidable equality. -/ -protected def List.hasDecEq {α : Type u} [DecidableEq α] : (a b : List α) → Decidable (Eq a b) - | nil, nil => isTrue rfl - | cons _ _, nil => isFalse (fun h => List.noConfusion h) - | nil, cons _ _ => isFalse (fun h => List.noConfusion h) - | cons a as, cons b bs => - match decEq a b with - | isTrue hab => - match List.hasDecEq as bs with - | isTrue habs => isTrue (hab ▸ habs ▸ rfl) - | isFalse nabs => isFalse (fun h => List.noConfusion h (fun _ habs => absurd habs nabs)) - | isFalse nab => isFalse (fun h => List.noConfusion h (fun hab _ => absurd hab nab)) +protected def List.hasDecEq {α : Type u} [d : DecidableEq α] (a b : List α) : Decidable (Eq a b) where + decide := a.beq b + of_decide := go a b +where + go : (a b : List α) → cond (a.beq b) (Eq a b) (Not (Eq a b)) + | nil, nil => rfl + | cons _ _, nil => List.noConfusion + | nil, cons _ _ => List.noConfusion + | cons a as, cons b bs => show cond (and (decide (Eq a b)) (List.beq as bs)) _ _ from + match d a b, List.beq as bs, go as bs with + | isFalse h, _, _ => fun h' => List.noConfusion h' fun h1 _ => absurd h1 h + | isTrue _, false, h => fun h' => List.noConfusion h' fun _ h2 => absurd h2 h + | isTrue h, true, h' => h ▸ h' ▸ rfl instance {α : Type u} [DecidableEq α] : DecidableEq (List α) := List.hasDecEq @@ -2693,10 +2751,12 @@ Decides whether two strings are equal. Normally used via the `DecidableEq String At runtime, this function is overridden with an efficient native implementation. -/ @[extern "lean_string_dec_eq"] -def String.decEq (s₁ s₂ : @& String) : Decidable (Eq s₁ s₂) := - match s₁, s₂ with - | ⟨s₁⟩, ⟨s₂⟩ => - dite (Eq s₁ s₂) (fun h => isTrue (congrArg _ h)) (fun h => isFalse (fun h' => String.noConfusion h' (fun h' => absurd h' h))) +def String.decEq (s₁ s₂ : @& String) : Decidable (Eq s₁ s₂) where + decide := decide (Eq s₁.data s₂.data) + of_decide := + match instDecidableEqList s₁.data s₂.data with + | isTrue h => congrArg String.mk h + | isFalse h => fun h' => String.noConfusion h' h instance : DecidableEq String := String.decEq @@ -2717,10 +2777,13 @@ structure String.Pos where instance : Inhabited String.Pos where default := {} -instance : DecidableEq String.Pos := - fun ⟨a⟩ ⟨b⟩ => match decEq a b with - | isTrue h => isTrue (h ▸ rfl) - | isFalse h => isFalse (fun he => String.Pos.noConfusion he fun he => absurd he h) +instance : DecidableEq String.Pos := fun a b => { + decide := decide (Eq a.byteIdx b.byteIdx) + of_decide := + match instDecidableEqNat a.byteIdx b.byteIdx with + | isTrue h => congrArg String.Pos.mk h + | isFalse h => fun h' => String.Pos.noConfusion h' h +} /-- A region or slice of some underlying string. diff --git a/src/Init/PropLemmas.lean b/src/Init/PropLemmas.lean index 0f925e5b08aa..a5b7a0b67f10 100644 --- a/src/Init/PropLemmas.lean +++ b/src/Init/PropLemmas.lean @@ -609,9 +609,12 @@ instance [DecidablePred p] : DecidablePred (p ∘ f) := /-- Prove that `a` is decidable by constructing a boolean `b` and a proof that `b ↔ a`. (This is sometimes taken as an alternate definition of decidability.) -/ -def decidable_of_bool : ∀ (b : Bool), (b ↔ a) → Decidable a - | true, h => isTrue (h.1 rfl) - | false, h => isFalse (mt h.2 Bool.noConfusion) +def decidable_of_bool (b : Bool) (h : b ↔ a) : Decidable a where + decide := b + of_decide := + match b, h with + | true, h => h.1 rfl + | false, h => mt h.2 Bool.noConfusion protected theorem Decidable.not_forall {p : α → Prop} [Decidable (∃ x, ¬p x)] [∀ x, Decidable (p x)] : (¬∀ x, p x) ↔ ∃ x, ¬p x := diff --git a/src/Lean/Compiler/LCNF/MonoTypes.lean b/src/Lean/Compiler/LCNF/MonoTypes.lean index 44d82eb96d04..19bf4dfcb0be 100644 --- a/src/Lean/Compiler/LCNF/MonoTypes.lean +++ b/src/Lean/Compiler/LCNF/MonoTypes.lean @@ -92,7 +92,6 @@ where else return anyExpr | .const ``lcAny _ => return anyExpr - | .const ``Decidable _ => return mkConst ``Bool | .const declName us => if let some info ← hasTrivialStructure? declName then let ctorType ← getOtherDeclBaseType info.ctorName [] diff --git a/src/Lean/Compiler/LCNF/Passes.lean b/src/Lean/Compiler/LCNF/Passes.lean index 8a6ad50a9230..2abd1941acaf 100644 --- a/src/Lean/Compiler/LCNF/Passes.lean +++ b/src/Lean/Compiler/LCNF/Passes.lean @@ -18,6 +18,7 @@ import Lean.Compiler.LCNF.LambdaLifting import Lean.Compiler.LCNF.FloatLetIn import Lean.Compiler.LCNF.ReduceArity import Lean.Compiler.LCNF.ElimDeadBranches +import Lean.Compiler.LCNF.StructProjCases namespace Lean.Compiler.LCNF @@ -76,6 +77,7 @@ def builtinPassManager : PassManager := { lambdaLifting, extendJoinPointContext (phase := .mono) (occurrence := 1), simp (occurrence := 5) (phase := .mono), + structProjCases, cse (occurrence := 2) (phase := .mono), saveMono -- End of mono phase ] diff --git a/src/Lean/Compiler/LCNF/PullLetDecls.lean b/src/Lean/Compiler/LCNF/PullLetDecls.lean index ec48511a6920..cb07a1bcfa15 100644 --- a/src/Lean/Compiler/LCNF/PullLetDecls.lean +++ b/src/Lean/Compiler/LCNF/PullLetDecls.lean @@ -69,14 +69,9 @@ mutual partial def pullDecls (code : Code) : PullM Code := do match code with | .cases c => - -- At the present time, we can't correctly enforce the dependencies required for lifting - -- out of a cases expression on Decidable, so we disable this optimization. - if c.typeName == ``Decidable then - return code - else - withCheckpoint do - let alts ← c.alts.mapMonoM pullAlt - return code.updateAlts! alts + withCheckpoint do + let alts ← c.alts.mapMonoM pullAlt + return code.updateAlts! alts | .let decl k => if (← shouldPull decl) then pullDecls k diff --git a/src/Lean/Compiler/LCNF/Simp/InlineCandidate.lean b/src/Lean/Compiler/LCNF/Simp/InlineCandidate.lean index ac6f1f8d2d74..f08efa22bddc 100644 --- a/src/Lean/Compiler/LCNF/Simp/InlineCandidate.lean +++ b/src/Lean/Compiler/LCNF/Simp/InlineCandidate.lean @@ -52,15 +52,7 @@ def inlineCandidate? (e : LetValue) : SimpM (Option InlineCandidateInfo) := do We assume that at the base phase these annotations are for the instance methods that have been lambda lifted. -/ if (← inBasePhase <&&> Meta.isInstance decl.name) then - unless decl.name == ``instDecidableEqBool do - /- - TODO: remove this hack after we refactor `Decidable` as suggested by Gabriel. - Recall that the current `Decidable` class is special case since it is an inductive datatype which is not a - structure like all other type classes. This is bad since it prevents us from treating all classes in a uniform - way. After we change `Decidable` to a structure as suggested by Gabriel, we should only accept type classes - that are structures. Moreover, we should reject instances that have only one exit point producing an explicit structure. - -/ - return false + return false if decl.alwaysInlineAttr then return true -- TODO: check inlining quota if decl.inlineAttr || decl.inlineIfReduceAttr then return true diff --git a/src/Lean/Compiler/LCNF/StructProjCases.lean b/src/Lean/Compiler/LCNF/StructProjCases.lean new file mode 100644 index 000000000000..69198d4bae79 --- /dev/null +++ b/src/Lean/Compiler/LCNF/StructProjCases.lean @@ -0,0 +1,132 @@ +/- +Copyright (c) 2025 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Cameron Zwarich +-/ +prelude +import Lean.Compiler.LCNF.Basic +import Lean.Compiler.LCNF.InferType +import Lean.Compiler.LCNF.MonoTypes +import Lean.Compiler.LCNF.PassManager +import Lean.Compiler.LCNF.PrettyPrinter + +namespace Lean.Compiler.LCNF +namespace StructProjCases + +def findStructCtorInfo? (typeName : Name) : CoreM (Option ConstructorVal) := do + let .inductInfo info ← getConstInfo typeName | return none + let [ctorName] := info.ctors | return none + let some (.ctorInfo ctorInfo) := (← getEnv).find? ctorName | return none + return ctorInfo + +def mkFieldParamsForCtorType (e : Expr) (numParams : Nat): CompilerM (Array Param) := do + let rec loop (params : Array Param) (e : Expr) (numParams : Nat): CompilerM (Array Param) := do + match e with + | .forallE name type body _ => + if numParams == 0 then + let param ← mkParam name (← toMonoType type) false + loop (params.push param) body numParams + else + loop params body (numParams - 1) + | _ => return params + loop #[] e numParams + +structure StructProjState where + projMap : Std.HashMap FVarId (Array FVarId) := {} + fvarMap : Std.HashMap FVarId FVarId := {} + +abbrev M := StateRefT StructProjState CompilerM + +def M.run (x : M α) : CompilerM α := do + x.run' {} + +def remapFVar (fvarId : FVarId) : M FVarId := do + match (← get).fvarMap[fvarId]? with + | some newFvarId => return newFvarId + | none => return fvarId + +mutual + +partial def visitCode (code : Code) : M Code := do + match code with + | .let decl k => + match decl.value with + | .proj typeName i base => + eraseLetDecl decl + let base ← remapFVar base + if let some projVars := (← get).projMap[base]? then + modify fun s => { s with fvarMap := s.fvarMap.insert decl.fvarId projVars[i]! } + visitCode k + else + let some ctorInfo ← findStructCtorInfo? typeName | panic! "expected struct constructor" + let params ← mkFieldParamsForCtorType ctorInfo.type ctorInfo.numParams + assert! params.size == ctorInfo.numFields + let fvars := params.map (·.fvarId) + modify fun s => { s with projMap := s.projMap.insert base fvars, + fvarMap := s.fvarMap.insert decl.fvarId fvars[i]! } + let k ← visitCode k + modify fun s => { s with projMap := s.projMap.erase base } + let resultType ← toMonoType (← k.inferType) + let alts := #[.alt ctorInfo.name params k] + return .cases { typeName, resultType, discr := base, alts } + | _ => return code.updateLet! (← decl.updateValue (← visitLetValue decl.value)) (← visitCode k) + | .fun decl k => + let decl ← decl.updateValue (← visitCode decl.value) + return code.updateFun! decl (← visitCode k) + | .jp decl k => + let decl ← decl.updateValue (← visitCode decl.value) + return code.updateFun! decl (← visitCode k) + | .jmp fvarId args => + return code.updateJmp! (← remapFVar fvarId) (← args.mapM visitArg) + | .cases cases => + let discr ← remapFVar cases.discr + if let #[.alt ctorName params k] := cases.alts then + if let some projVars := (← get).projMap[discr]? then + assert! projVars.size == params.size + for param in params, projVar in projVars do + modify fun s => { s with fvarMap := s.fvarMap.insert param.fvarId projVar } + eraseParam param + visitCode k + else + let fvars := params.map (·.fvarId) + modify fun s => { s with projMap := s.projMap.insert discr fvars } + let k ← visitCode k + modify fun s => { s with projMap := s.projMap.erase discr } + -- TODO: Should this also preserve the alt allocation? + return code.updateCases! cases.resultType discr #[.alt ctorName params k] + else + let alts ← cases.alts.mapM (visitAlt ·) + return code.updateCases! cases.resultType discr alts + | .return fvarId => return code.updateReturn! (← remapFVar fvarId) + | .unreach .. => return code + +partial def visitLetValue (v : LetValue) : M LetValue := do + match v with + | .const _ _ args => + return v.updateArgs! (← args.mapM visitArg) + | .fvar fvarId args => + return v.updateFVar! (← remapFVar fvarId) (← args.mapM visitArg) + | .value _ | .erased => return v + -- Projections should be handled directly by `visitCode`. + | .proj .. => unreachable! + +partial def visitAlt (alt : Alt) : M Alt := do + return alt.updateCode (← visitCode alt.getCode) + +partial def visitArg (arg : Arg) : M Arg := + match arg with + | .fvar fvarId => return arg.updateFVar! (← remapFVar fvarId) + | .type _ | .erased => return arg + +end + +def visitDecl (decl : Decl) : M Decl := do + let value ← decl.value.mapCodeM (visitCode ·) + return { decl with value } + +end StructProjCases + +def structProjCases : Pass := + .mkPerDeclaration `structProjCases (StructProjCases.visitDecl · |>.run) .mono + +end Lean.Compiler.LCNF diff --git a/src/Lean/Compiler/LCNF/ToMono.lean b/src/Lean/Compiler/LCNF/ToMono.lean index 97eb5dddfd95..8d340860459d 100644 --- a/src/Lean/Compiler/LCNF/ToMono.lean +++ b/src/Lean/Compiler/LCNF/ToMono.lean @@ -46,15 +46,7 @@ partial def LetValue.toMono (e : LetValue) : ToMonoM LetValue := do match e with | .erased | .value .. => return e | .const declName _ args => - if declName == ``Decidable.isTrue then - return .const ``Bool.true [] #[] - else if declName == ``Decidable.isFalse then - return .const ``Bool.false [] #[] - else if declName == ``Decidable.decide then - -- Decidable.decide is the identity function since Decidable - -- and Bool have the same runtime representation. - return args[1]!.toLetValue - else if let some e' ← isTrivialConstructorApp? declName args then + if let some e' ← isTrivialConstructorApp? declName args then e'.toMono else if let some (.ctorInfo ctorInfo) := (← getEnv).find? declName then ctorAppToMono ctorInfo args @@ -89,18 +81,6 @@ partial def FunDecl.toMono (decl : FunDecl) : ToMonoM FunDecl := do let value ← decl.value.toMono decl.update type params value -/-- Convert `cases` `Decidable` => `Bool` -/ -partial def decToMono (c : Cases) (_ : c.typeName == ``Decidable) : ToMonoM Code := do - let resultType ← toMonoType c.resultType - let alts ← c.alts.mapM fun alt => do - match alt with - | .default k => return alt.updateCode (← k.toMono) - | .alt ctorName ps k => - eraseParams ps - let ctorName := if ctorName == ``Decidable.isTrue then ``Bool.true else ``Bool.false - return .alt ctorName #[] (← k.toMono) - return .cases { c with resultType, alts, typeName := ``Bool } - /-- Eliminate `cases` for `Nat`. -/ partial def casesNatToMono (c: Cases) (_ : c.typeName == ``Nat) : ToMonoM Code := do let resultType ← toMonoType c.resultType @@ -223,9 +203,7 @@ partial def Code.toMono (code : Code) : ToMonoM Code := do | .unreach type => return .unreach (← toMonoType type) | .return .. | .jmp .. => return code | .cases c => - if h : c.typeName == ``Decidable then - decToMono c h - else if h : c.typeName == ``Nat then + if h : c.typeName == ``Nat then casesNatToMono c h else if h : c.typeName == ``Int then casesIntToMono c h diff --git a/src/Lean/CoreM.lean b/src/Lean/CoreM.lean index 1f1c729d5bc3..6a7f4cafb42f 100644 --- a/src/Lean/CoreM.lean +++ b/src/Lean/CoreM.lean @@ -561,7 +561,7 @@ private def checkUnsupported [Monad m] [MonadEnv m] [MonadError m] (decl : Decla | _ => pure () register_builtin_option compiler.enableNew : Bool := { - defValue := false + defValue := true -- false group := "compiler" descr := "(compiler) enable the new code generator, this should have no significant effect on your code but it does help to test the new code generator; unset to only use the old code generator instead" } diff --git a/src/Lean/Elab/Tactic/ElabTerm.lean b/src/Lean/Elab/Tactic/ElabTerm.lean index 05df41787b08..a8f7aeea8108 100644 --- a/src/Lean/Elab/Tactic/ElabTerm.lean +++ b/src/Lean/Elab/Tactic/ElabTerm.lean @@ -366,26 +366,26 @@ private def preprocessPropToDecide (expectedType : Expr) : TermElabM Expr := do return expectedType /-- -Given the decidable instance `inst`, reduces it and returns a decidable instance expression -in whnf that can be regarded as the reason for the failure of `inst` to fully reduce. +Given the boolean expression `bool`, reduces it and returns a boolean expression in whnf that can +be regarded as the reason for the failure of `bool` to fully reduce. -/ -private partial def blameDecideReductionFailure (inst : Expr) : MetaM Expr := withIncRecDepth do - let inst ← whnf inst - -- If it's the Decidable recursor, then blame the major premise. - if inst.isAppOfArity ``Decidable.rec 5 then - return ← blameDecideReductionFailure inst.appArg! +private partial def blameDecideReductionFailure (bool : Expr) : MetaM Expr := withIncRecDepth do + let bool ← whnf bool + -- If it's the Bool recursor, then blame the major premise. + if bool.isAppOfArity ``Bool.rec 4 then + return ← blameDecideReductionFailure bool.appArg! -- If it is a matcher, look for a discriminant that's a Decidable instance to blame. - if let .const c _ := inst.getAppFn then + if let .const c _ := bool.getAppFn then if let some info ← getMatcherInfo? c then - if inst.getAppNumArgs == info.arity then - let args := inst.getAppArgs + if bool.getAppNumArgs == info.arity then + let args := bool.getAppArgs for i in [0:info.numDiscrs] do - let inst' := args[info.numParams + 1 + i]! - if (← Meta.isClass? (← inferType inst')) == ``Decidable then - let inst'' ← whnf inst' - if !(inst''.isAppOf ``isTrue || inst''.isAppOf ``isFalse) then + let bool' := args[info.numParams + 1 + i]! + if (← inferType bool').isConstOf ``Bool then + let inst'' ← whnf bool' + if !(inst''.isAppOf ``Bool.true || inst''.isAppOf ``Bool.false) then return ← blameDecideReductionFailure inst'' - return inst + return bool private unsafe def elabNativeDecideCoreUnsafe (tacticName : Name) (expectedType : Expr) : TacticM Expr := do let d ← mkDecide expectedType @@ -451,18 +451,18 @@ def evalDecideCore (tacticName : Name) (cfg : Parser.Tactic.DecideConfig) : Tact doElab expectedType where doElab (expectedType : Expr) : TacticM Expr := do - let pf ← mkDecideProof expectedType + let dec ← mkDecide expectedType -- Get instance from `pf` - let s := pf.appFn!.appArg! - let r ← withAtLeastTransparency .default <| whnf s - if r.isAppOf ``isTrue then + let inst := dec.appArg! + -- reduce `dec` + let r ← withAtLeastTransparency .default <| whnf dec + if r.isConstOf ``Bool.true then -- Success! - -- While we have a proof from reduction, we do not embed it in the proof term, - -- and instead we let the kernel recompute it during type checking from the following more - -- efficient term. The kernel handles the unification `e =?= true` specially. - return pf + let eq := mkApp3 (mkConst ``Eq [1]) (mkConst ``Bool) dec (mkConst ``Bool.true) + let refl := mkApp2 (.const ``id [0]) eq reflBoolTrue + return mkApp3 (mkConst ``of_decide_eq_true) expectedType inst refl else - diagnose expectedType s r + diagnose expectedType dec r doKernel (expectedType : Expr) : TacticM Expr := do let pf ← mkDecideProof expectedType -- Get instance from `pf` @@ -479,16 +479,16 @@ where mkAuxLemma lemmaLevels expectedType pf return mkConst lemmaName (lemmaLevels.map .param) catch _ => - diagnose expectedType s none - diagnose {α : Type} (expectedType s : Expr) (r? : Option Expr) : TacticM α := + diagnose expectedType (mkApp2 (mkConst ``Decidable.decide) expectedType s) none + diagnose {α : Type} (expectedType b : Expr) (r? : Option Expr) : TacticM α := -- Diagnose the failure, lazily so that there is no performance impact if `decide` isn't being used interactively. throwError MessageData.ofLazyM (es := #[expectedType]) do - let r ← r?.getDM (withAtLeastTransparency .default <| whnf s) - if r.isAppOf ``isTrue then + let r ← r?.getDM (withAtLeastTransparency .default <| whnf b) + if r.isConstOf ``Bool.true then return m!"\ tactic '{tacticName}' failed. internal error: the elaborator is able to reduce the \ '{.ofConstName ``Decidable}' instance, but the kernel is not able to" - else if r.isAppOf ``isFalse then + else if r.isConstOf ``Bool.false then return m!"\ tactic '{tacticName}' proved that the proposition\ {indentExpr expectedType}\n\ @@ -496,7 +496,7 @@ where -- Re-reduce the instance and collect diagnostics, to get all unfolded Decidable instances let (reason, unfoldedInsts) ← withoutModifyingState <| withOptions (fun opt => diagnostics.set opt true) do modifyDiag (fun _ => {}) - let reason ← withAtLeastTransparency .default <| blameDecideReductionFailure s + let reason ← withAtLeastTransparency .default <| blameDecideReductionFailure b let unfolded := (← get).diag.unfoldCounter.foldl (init := #[]) fun cs n _ => cs.push n let unfoldedInsts ← unfolded |>.qsort Name.lt |>.filterMapM fun n => do let e ← mkConstWithLevelParams n @@ -507,11 +507,11 @@ where return (reason, unfoldedInsts) let stuckMsg := if unfoldedInsts.isEmpty then - m!"Reduction got stuck at the '{.ofConstName ``Decidable}' instance{indentExpr reason}" + m!"Reduction got stuck at the boolean{indentExpr reason}" else let instances := if unfoldedInsts.size == 1 then "instance" else "instances" m!"After unfolding the {instances} {.andList unfoldedInsts.toList}, \ - reduction got stuck at the '{.ofConstName ``Decidable}' instance{indentExpr reason}" + reduction got stuck at the boolean{indentExpr reason}" let hint := if reason.isAppOf ``Eq.rec then m!"\n\n\ @@ -534,9 +534,9 @@ where return m!"\ tactic '{tacticName}' failed for proposition\ {indentExpr expectedType}\n\ - since its '{.ofConstName ``Decidable}' instance\ - {indentExpr s}\n\ - did not reduce to '{.ofConstName ``isTrue}' or '{.ofConstName ``isFalse}'.\n\n\ + since\ + {indentExpr b}\n\ + did not reduce to '{.ofConstName ``Bool.true}' or '{.ofConstName ``Bool.false}'.\n\n\ {stuckMsg}{hint}" declare_config_elab elabDecideConfig Parser.Tactic.DecideConfig diff --git a/src/Lean/Elab/Tactic/FalseOrByContra.lean b/src/Lean/Elab/Tactic/FalseOrByContra.lean index 825d259d4ac2..7b62a79b8a8a 100644 --- a/src/Lean/Elab/Tactic/FalseOrByContra.lean +++ b/src/Lean/Elab/Tactic/FalseOrByContra.lean @@ -54,12 +54,12 @@ partial def falseOrByContra (g : MVarId) (useClassical : Option Bool := none) : match gs with | some [] => return none | some [g] => return some (← g.intro1).2 - | some _ => panic! "expected at most one sugoal" + | some _ => panic! "expected at most one subgoal" | none => match (← g.applyConst ``False.elim) with | [] => return none | [g] => return some g - | _ => panic! "expected at most one sugoal" + | _ => panic! "expected at most one subgoal" @[builtin_tactic Lean.Parser.Tactic.falseOrByContra] def elabFalseOrByContra : Tactic diff --git a/src/Lean/Meta/Constructions/NoConfusion.lean b/src/Lean/Meta/Constructions/NoConfusion.lean index 52cf24dbcfe3..275f5fe46f37 100644 --- a/src/Lean/Meta/Constructions/NoConfusion.lean +++ b/src/Lean/Meta/Constructions/NoConfusion.lean @@ -104,7 +104,7 @@ where withLocalDecl `y BinderInfo.implicit enumType fun y => do withLocalDeclD `h (← mkEq x y) fun h => do let declType ← mkForallFVars #[P, x, y, h] (mkApp3 noConfusionType P x y) - let declValue ← mkLambdaFVars #[P, x, y, h] (← mkAppOptM ``noConfusionEnum #[none, none, none, toCtorIdx, P, x, y, h]) + let declValue ← mkLambdaFVars #[P, x, y, h] (← mkAppOptM ``noConfusionEnum #[none, toCtorIdx, P, x, y, h]) let declName := Name.mkStr enumName "noConfusion" addAndCompile <| Declaration.defnDecl { name := declName diff --git a/src/Lean/Meta/Tactic/Cases.lean b/src/Lean/Meta/Tactic/Cases.lean index a2c24fc403de..14f77483c237 100644 --- a/src/Lean/Meta/Tactic/Cases.lean +++ b/src/Lean/Meta/Tactic/Cases.lean @@ -252,9 +252,13 @@ private def inductionCasesOn (mvarId : MVarId) (majorFVarId : FVarId) (givenName let majorType ← inferType (mkFVar majorFVarId) let (us, params) ← getInductiveUniverseAndParams majorType let mut casesOn := mkCasesOnName ctx.inductiveVal.name + let mut ctors := ctx.inductiveVal.ctors.toArray if useNatCasesAuxOn && ctx.inductiveVal.name == ``Nat && (← getEnv).contains ``Nat.casesAuxOn then casesOn := ``Nat.casesAuxOn - let ctors := ctx.inductiveVal.ctors.toArray + else if useNatCasesAuxOn && ctx.inductiveVal.name == ``Decidable && (← getEnv).contains `Decidable.falseTrueCases then + -- hack: allow `byCasesDec` to use the normal `cases` frontend + casesOn := `Decidable.falseTrueCases + ctors := #[``Decidable.isFalse, ``Decidable.isTrue] let s ← mvarId.induction majorFVarId casesOn givenNames return toCasesSubgoals s ctors majorFVarId us params @@ -351,7 +355,7 @@ Given `dec : Decidable p`, split the goal in two subgoals: one containing the hy def _root_.Lean.MVarId.byCasesDec (mvarId : MVarId) (p : Expr) (dec : Expr) (hName : Name := `h) : MetaM (ByCasesSubgoal × ByCasesSubgoal) := do let mvarId ← mvarId.assert `hByCases (mkApp (mkConst ``Decidable) p) dec let (fvarId, mvarId) ← mvarId.intro1 - let #[s₁, s₂] ← mvarId.cases fvarId #[{ varNames := [hName] }, { varNames := [hName] }] | + let #[s₁, s₂] ← mvarId.cases fvarId #[{ varNames := [hName] }, { varNames := [hName] }] true | throwError "'byCasesDec' tactic failed, unexpected number of subgoals" -- We flip `s₁` and `s₂` because `isFalse` is the first constructor. return ((← toByCasesSubgoal s₂), (← toByCasesSubgoal s₁)) diff --git a/src/Lean/Meta/Tactic/Simp/BuiltinSimprocs/Core.lean b/src/Lean/Meta/Tactic/Simp/BuiltinSimprocs/Core.lean index 38060e97f155..8b1a24ecb070 100644 --- a/src/Lean/Meta/Tactic/Simp/BuiltinSimprocs/Core.lean +++ b/src/Lean/Meta/Tactic/Simp/BuiltinSimprocs/Core.lean @@ -56,9 +56,10 @@ builtin_dsimproc ↓ [simp, seval] dreduceIte (ite _ _ _) := fun e => do -/ let r ← simp c if r.expr.isTrue || r.expr.isFalse then - match_expr (← whnfD i) with - | Decidable.isTrue _ _ => return .visit tb - | Decidable.isFalse _ _ => return .visit eb + let dec := mkApp2 (mkConst ``Decidable.decide) c i + match_expr (← whnfD dec) with + | Bool.true => return .visit tb + | Bool.false => return .visit eb | _ => return .continue return .continue @@ -70,9 +71,10 @@ builtin_dsimproc ↓ [simp, seval] dreduceDIte (dite _ _ _) := fun e => do -- See comment at `dreduceIte` let r ← simp c if r.expr.isTrue || r.expr.isFalse then - match_expr (← whnfD i) with - | Decidable.isTrue _ h => return .visit (mkApp tb h).headBeta - | Decidable.isFalse _ h => return .visit (mkApp eb h).headBeta + let dec := mkApp2 (mkConst ``Decidable.decide) c i + match_expr (← whnfD dec) with + | Bool.true => return .visit (mkApp tb (mkApp2 (mkConst `Decidable.of_decide) c i)).headBeta + | Bool.false => return .visit (mkApp eb (mkApp2 (mkConst `Decidable.of_decide) c i)).headBeta | _ => return .continue return .continue diff --git a/src/Lean/Util/FoldConsts.lean b/src/Lean/Util/FoldConsts.lean index 107e5f79a634..f9da96ddf72f 100644 --- a/src/Lean/Util/FoldConsts.lean +++ b/src/Lean/Util/FoldConsts.lean @@ -13,8 +13,8 @@ namespace Expr namespace FoldConstsImpl unsafe structure State where - visited : PtrSet Expr := mkPtrSet - visitedConsts : NameHashSet := {} + visited : PtrSet Expr := mkPtrSet + visitedConsts : NameHashSet := {} unsafe abbrev FoldM := StateM State @@ -22,7 +22,7 @@ unsafe def fold {α : Type} (f : Name → α → α) (e : Expr) (acc : α) : Fol let rec visit (e : Expr) (acc : α) : FoldM α := do if (← get).visited.contains e then return acc - modify fun s => { s with visited := s.visited.insert e } + modify fun s => let a := s.1; let b := s.2; ⟨a.insert e, b⟩ match e with | .forallE _ d b _ => visit b (← visit d acc) | .lam _ d b _ => visit b (← visit d acc) @@ -34,7 +34,7 @@ unsafe def fold {α : Type} (f : Name → α → α) (e : Expr) (acc : α) : Fol if (← get).visitedConsts.contains c then return acc else - modify fun s => { s with visitedConsts := s.visitedConsts.insert c }; + modify fun s => let a := s.1; let b := s.2; ⟨a, b.insert c⟩ return f c acc | _ => return acc visit e acc diff --git a/src/Std/Data/DTreeMap/Internal/Balancing.lean b/src/Std/Data/DTreeMap/Internal/Balancing.lean index d79e7ba71623..77c832b698a4 100644 --- a/src/Std/Data/DTreeMap/Internal/Balancing.lean +++ b/src/Std/Data/DTreeMap/Internal/Balancing.lean @@ -654,93 +654,69 @@ theorem balanced_rotateL (k v l rs rk rv rl rr) (hl : l.Balanced) (hlr : BalanceLErasePrecond l.size rs ∨ BalanceLErasePrecond rs l.size) (hh : rs > delta * l.size) : (rotateL k v l rk rv rl rr : Impl α β).Balanced := by - fun_cases rotateL k v l rk rv rl rr <;> dsimp only [rotateL] - · split - · next h => - exact balanced_singleL _ _ _ _ _ _ _ _ hl hr hlr hh h - · contradiction - · rw [if_neg ‹_›] - tree_tac - · rw [if_neg ‹_›] - exact balanced_doubleL k v _ _ _ _ _ _ _ _ _ _ hl hr hlr hh ‹_› + induction k, v, l, rk, rv, rl, rr using rotateL.fun_cases_unfolding + · exact balanced_singleL _ _ _ _ _ _ _ _ hl hr hlr hh ‹_› + · tree_tac + · exact balanced_doubleL _ _ _ _ _ _ _ _ _ _ _ _ hl hr hlr hh ‹_› theorem balanced_rotateR (k v ls lk lv ll lr r) (hl : (Impl.inner ls lk lv ll lr).Balanced) (hr : r.Balanced) (hlr : BalanceLErasePrecond ls r.size ∨ BalanceLErasePrecond r.size ls) (hh : ls > delta * r.size) : (rotateR k v lk lv ll lr r : Impl α β).Balanced := by - fun_cases rotateR k v lk lv ll lr r <;> dsimp only [rotateR] - · split - · next h => - exact balanced_singleR k v _ _ _ _ _ _ hl hr hlr hh h - · contradiction - · rw [if_neg ‹_›] - tree_tac - · rw [if_neg ‹_›] - exact balanced_doubleR k v _ _ _ _ _ _ _ _ _ _ hl hr hlr hh ‹_› + induction k, v, lk, lv, ll, lr, r using rotateR.fun_cases_unfolding + · exact balanced_singleR _ _ _ _ _ _ _ _ hl hr hlr hh ‹_› + · tree_tac + · exact balanced_doubleR _ _ _ _ _ _ _ _ _ _ _ _ hl hr hlr hh ‹_› theorem balanceL_eq_balanceLErase {k : α} {v : β k} {l r : Impl α β} {hlb hrb hlr} : balanceL k v l r hlb hrb hlr = balanceLErase k v l r hlb hrb hlr.erase := by - fun_cases balanceL k v l r hlb hrb hlr - all_goals dsimp only [balanceL, balanceLErase] - split - · split <;> contradiction - · rfl + induction k, v, l, r, hlb, hrb, hlr using balanceL.fun_cases_unfolding + all_goals try rfl + all_goals simp only [balanceLErase, *, ↓reduceDIte, ↓reduceIte] theorem balanceLErase_eq_balanceL! {k : α} {v : β k} {l r : Impl α β} {hlb hrb hlr} : balanceLErase k v l r hlb hrb hlr = balanceL! k v l r := by - fun_cases balanceL! k v l r - all_goals dsimp only [balanceLErase, balanceL!] - all_goals simp only [*] - all_goals dsimp only [dreduceDIte, dreduceIte] + induction k, v, l, r using balanceL!.fun_cases_unfolding + all_goals try rfl + all_goals simp only [balanceLErase, *, ↓reduceDIte, ↓reduceIte] all_goals contradiction theorem balanceL!_eq_balance! {k : α} {v : β k} {l r : Impl α β} (hlb : l.Balanced) (hrb : r.Balanced) (hlr : BalanceLErasePrecond l.size r.size) : balanceL! k v l r = balance! k v l r := by - cases k, v, l, r using balance!.fun_cases - all_goals dsimp only [balanceL!, balance!] - all_goals try simp only [*] - all_goals try dsimp only [dreduceDIte, dreduceIte] + induction k, v, l, r using balance!.fun_cases_unfolding all_goals try rfl + all_goals simp only [balanceL!, *, ↓reduceDIte, ↓reduceIte] all_goals try contradiction all_goals try (exfalso; tree_tac; done) congr; tree_tac theorem balanceR_eq_balanceRErase {k : α} {v : β k} {l r : Impl α β} {hlb hrb hlr} : balanceR k v l r hlb hrb hlr = balanceRErase k v l r hlb hrb hlr.erase := by - fun_cases balanceR k v l r hlb hrb hlr - all_goals dsimp only [balanceR, balanceRErase] - split - · split <;> contradiction - · rfl + induction k, v, l, r, hlb, hrb, hlr using balanceR.fun_cases_unfolding + all_goals try rfl + all_goals simp only [balanceRErase, *, ↓reduceDIte, ↓reduceIte] theorem balanceRErase_eq_balanceR! {k : α} {v : β k} {l r : Impl α β} {hlb hrb hlr} : balanceRErase k v l r hlb hrb hlr = balanceR! k v l r := by - fun_cases balanceR! k v l r - all_goals dsimp only [balanceRErase, balanceR!] - all_goals simp only [*] - all_goals dsimp only [dreduceDIte, dreduceIte] - all_goals contradiction + induction k, v, l, r, hlb, hrb, hlr using balanceRErase.fun_cases_unfolding + all_goals try rfl + all_goals simp only [balanceR!, *, ↓reduceIte] theorem balanceR!_eq_balance! {k : α} {v : β k} {l r : Impl α β} (hlb : l.Balanced) (hrb : r.Balanced) (hlr : BalanceLErasePrecond r.size l.size) : balanceR! k v l r = balance! k v l r := by - cases k, v, l, r using balance!.fun_cases - all_goals dsimp only [balanceR!, balance!] - all_goals try simp only [*] - all_goals try dsimp only [dreduceDIte, dreduceIte] + induction k, v, l, r using balance!.fun_cases_unfolding all_goals try rfl - all_goals try contradiction + all_goals simp only [balanceR!, *, ↓reduceDIte, ↓reduceIte] all_goals try (exfalso; tree_tac; done) congr; tree_tac theorem balance_eq_balance! {k : α} {v : β k} {l r : Impl α β} {hlb hrb hlr} : balance k v l r hlb hrb hlr = balance! k v l r := by - fun_cases balance! k v l r - all_goals dsimp only [balance, balance!] - all_goals simp only [*] - all_goals dsimp only [dreduceDIte] - all_goals contradiction + induction k, v, l, r, hlb, hrb, hlr using balance.fun_cases_unfolding + all_goals try rfl + all_goals simp only [balance!, *, ↓reduceIte, ↓reduceDIte] theorem balance_eq_inner [Ord α] {sz k v} {l r : Impl α β} (hl : (inner sz k v l r).Balanced) {h} : @@ -755,22 +731,19 @@ theorem balance_eq_inner [Ord α] {sz k v} {l r : Impl α β} theorem balance!_desc {k : α} {v : β k} {l r : Impl α β} (hlb : l.Balanced) (hrb : r.Balanced) (hlr : BalanceLErasePrecond l.size r.size ∨ BalanceLErasePrecond r.size l.size) : (balance! k v l r).size = l.size + 1 + r.size ∧ (balance! k v l r).Balanced := by - rw [balance!_eq_balanceₘ hlb hrb hlr, balanceₘ] - fun_cases balanceₘ k v l r - · rw [if_pos ‹_›, bin, balanced_inner_iff] + rw [balance!_eq_balanceₘ hlb hrb hlr] + induction k, v, l, r using balanceₘ.fun_cases_unfolding + · rw [bin, balanced_inner_iff] exact ⟨rfl, hlb, hrb, Or.inl ‹_›, rfl⟩ - · rw [if_neg ‹_›, dif_pos ‹_›] - simp only [size_rotateL (.left ‹_›), size_bin, size_inner] + · simp only [size_rotateL (.left ‹_›), size_bin, size_inner] rw [← Balanced.eq ‹_›] refine ⟨rfl, ?_⟩ apply balanced_rotateL <;> assumption - · rw [if_neg ‹_›, dif_neg ‹_›, dif_pos ‹_›] - simp only [size_rotateR (.right ‹_›), size_bin, size_inner] + · simp only [size_rotateR (.right ‹_›), size_bin, size_inner] rw [← Balanced.eq ‹_›] refine ⟨rfl, ?_⟩ apply balanced_rotateR <;> assumption - · rw [if_neg ‹_›, dif_neg ‹_›, dif_neg ‹_›] - exact ⟨rfl, ✓⟩ + · exact ⟨rfl, ✓⟩ @[Std.Internal.tree_tac] theorem size_balance! {k : α} {v : β k} {l r : Impl α β} (hlb : l.Balanced) (hrb : r.Balanced) diff --git a/src/Std/Tactic/BVDecide/Bitblast/BVExpr/Circuit/Impl/Expr.lean b/src/Std/Tactic/BVDecide/Bitblast/BVExpr/Circuit/Impl/Expr.lean index 24eb0871174c..071b207193db 100644 --- a/src/Std/Tactic/BVDecide/Bitblast/BVExpr/Circuit/Impl/Expr.lean +++ b/src/Std/Tactic/BVDecide/Bitblast/BVExpr/Circuit/Impl/Expr.lean @@ -106,7 +106,7 @@ structure Return (aig : AIG BVBit) (w : Nat) where result : AIG.ExtendingRefVecEntry aig w cache : Cache result.val.aig -set_option maxHeartbeats 400000 in +set_option maxHeartbeats 2000000 in def bitblast (aig : AIG BVBit) (input : WithCache (BVExpr w) aig) : Return aig w := let ⟨expr, cache⟩ := input goCache aig expr cache diff --git a/src/Std/Tactic/BVDecide/Bitblast/BVExpr/Circuit/Lemmas/Expr.lean b/src/Std/Tactic/BVDecide/Bitblast/BVExpr/Circuit/Lemmas/Expr.lean index a5d05af99b13..e461327c5364 100644 --- a/src/Std/Tactic/BVDecide/Bitblast/BVExpr/Circuit/Lemmas/Expr.lean +++ b/src/Std/Tactic/BVDecide/Bitblast/BVExpr/Circuit/Lemmas/Expr.lean @@ -158,7 +158,7 @@ theorem goCache_denote_mem_prefix (aig : AIG BVBit) (expr : BVExpr w) (assign : · intros apply (goCache aig expr cache).result.property -set_option maxHeartbeats 400000 +set_option maxHeartbeats 2000000 mutual diff --git a/stage0/stdlib/Lean/Compiler/IR/Basic.c b/stage0/stdlib/Lean/Compiler/IR/Basic.c index 6468938b29e8..d23d94db0a80 100644 --- a/stage0/stdlib/Lean/Compiler/IR/Basic.c +++ b/stage0/stdlib/Lean/Compiler/IR/Basic.c @@ -81,7 +81,6 @@ static lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprVarId____x static lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprVarId____x40_Lean_Compiler_IR_Basic___hyg_34____closed__12; static lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprIRType____x40_Lean_Compiler_IR_Basic___hyg_321____closed__7; LEAN_EXPORT lean_object* l_Lean_IR_instAlphaEqvArrayArg; -LEAN_EXPORT lean_object* l_Lean_IR_AltCore_setBody(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instReprVarId; static lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprCtorInfo____x40_Lean_Compiler_IR_Basic___hyg_1499____closed__10; LEAN_EXPORT lean_object* l_Lean_IR_CtorInfo_isRef___boxed(lean_object*); @@ -98,7 +97,6 @@ LEAN_EXPORT lean_object* l_Lean_IR_instInhabitedFnBody; LEAN_EXPORT lean_object* l_Lean_IR_IRType_isIrrelevant___boxed(lean_object*); LEAN_EXPORT lean_object* lean_ir_mk_jdecl(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_ir_mk_case(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_AltCore_modifyBody(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprIRType____x40_Lean_Compiler_IR_Basic___hyg_321____closed__8; LEAN_EXPORT lean_object* l_Lean_IR_reshapeAux(lean_object*, lean_object*, lean_object*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); @@ -131,6 +129,7 @@ LEAN_EXPORT uint8_t l_Lean_IR_IRType_isUnion(lean_object*); static lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprParam____x40_Lean_Compiler_IR_Basic___hyg_2147____closed__2; LEAN_EXPORT lean_object* l_Lean_IR_FnBody_resetBody(lean_object*); lean_object* l_Lean_RBNode_appendTrees___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_Alt_mmodifyBody___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_IR_IRType_isIrrelevant(lean_object*); LEAN_EXPORT uint8_t l_Lean_IR_Index_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Decl_isExtern___boxed(lean_object*); @@ -208,15 +207,18 @@ LEAN_EXPORT lean_object* l_Lean_IR_LocalContext_getType___boxed(lean_object*, le LEAN_EXPORT uint8_t l_Lean_IR_CtorInfo_isRef(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_IR_modifyJPs___spec__1(lean_object*, size_t, size_t, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_Alt_modifyBody(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprParam____x40_Lean_Compiler_IR_Basic___hyg_2147____closed__1; LEAN_EXPORT lean_object* l_Lean_IR_instInhabitedArg; LEAN_EXPORT lean_object* l_Lean_IR_LocalContext_addLocal(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instInhabitedIndexSet; +static lean_object* l_Lean_IR_Alt_mmodifyBody___rarg___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprIRType____x40_Lean_Compiler_IR_Basic___hyg_321____boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_IR_FnBody_alphaEqv(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprIRType____x40_Lean_Compiler_IR_Basic___hyg_321____closed__2; lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_Alt_mmodifyBody(lean_object*); LEAN_EXPORT lean_object* lean_ir_mk_uproj_expr(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_IR_Decl_isExtern(lean_object*); static lean_object* l_Lean_IR_instInhabitedCtorInfo___closed__1; @@ -233,12 +235,10 @@ LEAN_EXPORT lean_object* l_Lean_IR_LocalContext_getJPBody___boxed(lean_object*, LEAN_EXPORT lean_object* l_repr___at___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprIRType____x40_Lean_Compiler_IR_Basic___hyg_321____spec__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_FnBody_alphaEqv___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_instReprVarId___closed__1; -LEAN_EXPORT lean_object* l_Lean_IR_Alt_ctor(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprIRType____x40_Lean_Compiler_IR_Basic___hyg_321____closed__58; LEAN_EXPORT uint8_t l_Lean_IR_Arg_alphaEqv(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprIRType____x40_Lean_Compiler_IR_Basic___hyg_321____closed__36; static lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprParam____x40_Lean_Compiler_IR_Basic___hyg_2147____closed__9; -LEAN_EXPORT lean_object* l_Lean_IR_AltCore_body___boxed(lean_object*); static lean_object* l_Lean_IR_instAlphaEqvArrayArg___closed__1; LEAN_EXPORT uint8_t l_Lean_IR_LocalContext_isParam(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_LocalContext_isJP___boxed(lean_object*, lean_object*); @@ -268,6 +268,7 @@ static lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprIRType____ LEAN_EXPORT lean_object* l_Lean_IR_instHashableVarId___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprCtorInfo____x40_Lean_Compiler_IR_Basic___hyg_1499____boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprIRType____x40_Lean_Compiler_IR_Basic___hyg_321____closed__34; +LEAN_EXPORT lean_object* l_Lean_IR_Alt_mmodifyBody___rarg___lambda__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instAlphaEqvArg; lean_object* lean_array_fget(lean_object*, lean_object*); static lean_object* l_Lean_IR_instToStringJoinPointId___closed__1; @@ -282,7 +283,6 @@ LEAN_EXPORT lean_object* l_Lean_IR_instBEqArg; LEAN_EXPORT lean_object* l_Lean_IR_Decl_updateBody_x21(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprCtorInfo____x40_Lean_Compiler_IR_Basic___hyg_1499____closed__13; static lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprIRType____x40_Lean_Compiler_IR_Basic___hyg_321____closed__11; -LEAN_EXPORT lean_object* l_Lean_IR_AltCore_mmodifyBody(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_LocalContext_isLocalVar___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Decl_params___boxed(lean_object*); static lean_object* l_Lean_IR_instInhabitedArg___closed__1; @@ -306,12 +306,13 @@ LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_IR_addVarRename___spec__2(l LEAN_EXPORT lean_object* l_panic___at_Lean_IR_reshapeAux___spec__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at_Lean_IR_IRType_beq___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_getUnboxOpName(lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_Alt_default(lean_object*); LEAN_EXPORT uint8_t l_Lean_IR_args_alphaEqv(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_Alt_setBody(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprIRType____x40_Lean_Compiler_IR_Basic___hyg_321____closed__1; LEAN_EXPORT lean_object* lean_ir_mk_proj_expr(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprVarId____x40_Lean_Compiler_IR_Basic___hyg_34____closed__11; static lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprIRType____x40_Lean_Compiler_IR_Basic___hyg_321____closed__18; +LEAN_EXPORT lean_object* l_Lean_IR_Alt_body(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_LocalContext_getValue(lean_object*, lean_object*); lean_object* l_Repr_addAppParen(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_IR_addParamsRename___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -322,7 +323,6 @@ lean_object* lean_panic_fn(lean_object*, lean_object*); static lean_object* l_Lean_IR_instReprJoinPointId___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprParam____x40_Lean_Compiler_IR_Basic___hyg_2147_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_IR_mmodifyJPs___spec__1___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_AltCore_body(lean_object*); static lean_object* l_Lean_IR_instReprIRType___closed__1; lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_IR_addVarRename___spec__1(lean_object*, lean_object*, lean_object*); @@ -360,6 +360,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_IRType_instBEq; LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprIRType____x40_Lean_Compiler_IR_Basic___hyg_321____spec__3(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprIRType____x40_Lean_Compiler_IR_Basic___hyg_321____closed__4; size_t lean_usize_add(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_IR_Alt_body___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_mkIndexSet(lean_object*); static lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprVarId____x40_Lean_Compiler_IR_Basic___hyg_34____closed__10; static lean_object* l_Lean_IR_mkIf___closed__2; @@ -381,7 +382,6 @@ static lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprIRType____ LEAN_EXPORT lean_object* lean_ir_mk_param(lean_object*, uint8_t, lean_object*); static lean_object* l_Array_Array_repr___at___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprIRType____x40_Lean_Compiler_IR_Basic___hyg_321____spec__1___closed__7; LEAN_EXPORT lean_object* l_Lean_IR_LocalContext_addJP(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_AltCore_mmodifyBody___rarg(lean_object*, lean_object*, lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprIRType____x40_Lean_Compiler_IR_Basic___hyg_321____closed__17; LEAN_EXPORT lean_object* l_Lean_IR_VarId_alphaEqv___boxed(lean_object*, lean_object*, lean_object*); @@ -417,12 +417,12 @@ static lean_object* l_Lean_IR_instInhabitedExpr___closed__2; static lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprIRType____x40_Lean_Compiler_IR_Basic___hyg_321____closed__64; static lean_object* l_Lean_IR_getUnboxOpName___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprCtorInfo____x40_Lean_Compiler_IR_Basic___hyg_1499_(lean_object*, lean_object*); -static lean_object* l_Lean_IR_AltCore_mmodifyBody___rarg___closed__1; static lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprIRType____x40_Lean_Compiler_IR_Basic___hyg_321____closed__56; static lean_object* l_Lean_IR_Decl_updateBody_x21___closed__4; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instReprCtorInfo; lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_Alt_mmodifyBody___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_IR_VarId_alphaEqv___spec__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_IR_IRType_isScalar(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_IR_mmodifyJPs___spec__1___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -4053,25 +4053,6 @@ x_2 = lean_box(13); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_IR_Alt_ctor(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_IR_Alt_default(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} static lean_object* _init_l_Lean_IR_instInhabitedAlt___closed__1() { _start: { @@ -4548,7 +4529,7 @@ lean_ctor_set(x_5, 1, x_2); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_IR_AltCore_body(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_IR_Alt_body(lean_object* x_1) { _start: { if (lean_obj_tag(x_1) == 0) @@ -4567,16 +4548,16 @@ return x_3; } } } -LEAN_EXPORT lean_object* l_Lean_IR_AltCore_body___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_IR_Alt_body___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_IR_AltCore_body(x_1); +x_2 = l_Lean_IR_Alt_body(x_1); lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_IR_AltCore_setBody(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_IR_Alt_setBody(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -4626,7 +4607,7 @@ return x_9; } } } -LEAN_EXPORT lean_object* l_Lean_IR_AltCore_modifyBody(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_IR_Alt_modifyBody(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -4682,15 +4663,34 @@ return x_15; } } } -static lean_object* _init_l_Lean_IR_AltCore_mmodifyBody___rarg___closed__1() { +LEAN_EXPORT lean_object* l_Lean_IR_Alt_mmodifyBody___rarg___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_Alt_mmodifyBody___rarg___lambda__2(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_IR_Alt_mmodifyBody___rarg___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_IR_Alt_default), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_IR_Alt_mmodifyBody___rarg___lambda__2), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_AltCore_mmodifyBody___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_IR_Alt_mmodifyBody___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -4710,7 +4710,7 @@ lean_dec(x_6); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); lean_dec(x_7); -x_9 = lean_alloc_closure((void*)(l_Lean_IR_Alt_ctor), 2, 1); +x_9 = lean_alloc_closure((void*)(l_Lean_IR_Alt_mmodifyBody___rarg___lambda__1), 2, 1); lean_closure_set(x_9, 0, x_4); x_10 = lean_apply_1(x_2, x_5); x_11 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_9, x_10); @@ -4732,17 +4732,17 @@ x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); lean_dec(x_14); x_16 = lean_apply_1(x_2, x_12); -x_17 = l_Lean_IR_AltCore_mmodifyBody___rarg___closed__1; +x_17 = l_Lean_IR_Alt_mmodifyBody___rarg___closed__1; x_18 = lean_apply_4(x_15, lean_box(0), lean_box(0), x_17, x_16); return x_18; } } } -LEAN_EXPORT lean_object* l_Lean_IR_AltCore_mmodifyBody(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_IR_Alt_mmodifyBody(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_IR_AltCore_mmodifyBody___rarg), 3, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_IR_Alt_mmodifyBody___rarg), 3, 0); return x_2; } } @@ -16347,8 +16347,8 @@ l_Lean_IR_instInhabitedAlt___closed__1 = _init_l_Lean_IR_instInhabitedAlt___clos lean_mark_persistent(l_Lean_IR_instInhabitedAlt___closed__1); l_Lean_IR_instInhabitedAlt = _init_l_Lean_IR_instInhabitedAlt(); lean_mark_persistent(l_Lean_IR_instInhabitedAlt); -l_Lean_IR_AltCore_mmodifyBody___rarg___closed__1 = _init_l_Lean_IR_AltCore_mmodifyBody___rarg___closed__1(); -lean_mark_persistent(l_Lean_IR_AltCore_mmodifyBody___rarg___closed__1); +l_Lean_IR_Alt_mmodifyBody___rarg___closed__1 = _init_l_Lean_IR_Alt_mmodifyBody___rarg___closed__1(); +lean_mark_persistent(l_Lean_IR_Alt_mmodifyBody___rarg___closed__1); l_Lean_IR_FnBody_flatten___closed__1 = _init_l_Lean_IR_FnBody_flatten___closed__1(); lean_mark_persistent(l_Lean_IR_FnBody_flatten___closed__1); l_Lean_IR_reshapeAux___closed__1 = _init_l_Lean_IR_reshapeAux___closed__1(); diff --git a/stage0/stdlib/Lean/Compiler/IR/Borrow.c b/stage0/stdlib/Lean/Compiler/IR/Borrow.c index 065fa2a89ce9..7f92831d0cf3 100644 --- a/stage0/stdlib/Lean/Compiler/IR/Borrow.c +++ b/stage0/stdlib/Lean/Compiler/IR/Borrow.c @@ -166,12 +166,12 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_IR_Borrow_ownArgsIf uint64_t l_Lean_Name_hash___override(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_IR_Borrow_updateParamMap___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Borrow_mkInitParamMap(lean_object*, lean_object*); +lean_object* l_Lean_IR_Alt_body(lean_object*); uint64_t lean_uint64_xor(uint64_t, uint64_t); extern lean_object* l_Id_instMonad; lean_object* lean_panic_fn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Borrow_collectDecls(lean_object*, lean_object*); static lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_IR_Borrow_ParamMap_fmt___spec__1___closed__2; -lean_object* l_Lean_IR_AltCore_body(lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); static lean_object* l_Lean_IR_Borrow_ApplyParamMap_visitFnBody___closed__1; lean_object* lean_nat_mul(lean_object*, lean_object*); @@ -1637,7 +1637,7 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; size_t x_13; size_t x_14; lean_dec(x_5); x_8 = lean_array_uget(x_2, x_3); -x_9 = l_Lean_IR_AltCore_body(x_8); +x_9 = l_Lean_IR_Alt_body(x_8); lean_dec(x_8); lean_inc(x_1); x_10 = l_Lean_IR_Borrow_InitParamMap_visitFnBody(x_1, x_9, x_6); @@ -5829,7 +5829,7 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; size_t x_13; size_t x_14; lean_dec(x_4); x_8 = lean_array_uget(x_1, x_2); -x_9 = l_Lean_IR_AltCore_body(x_8); +x_9 = l_Lean_IR_Alt_body(x_8); lean_dec(x_8); lean_inc(x_5); x_10 = l_Lean_IR_Borrow_collectFnBody(x_9, x_5, x_6); diff --git a/stage0/stdlib/Lean/Compiler/IR/Boxing.c b/stage0/stdlib/Lean/Compiler/IR/Boxing.c index edeac53ad089..722ca9cc0954 100644 --- a/stage0/stdlib/Lean/Compiler/IR/Boxing.c +++ b/stage0/stdlib/Lean/Compiler/IR/Boxing.c @@ -80,6 +80,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_Explicit static lean_object* l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux___closed__2; LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitVDeclExpr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_Alt_mmodifyBody___at_Lean_IR_ExplicitBoxing_visitFnBody___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_IR_findEnvDecl_x27(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_ExplicitBoxing_mkBoxedVersionAux___closed__2; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); @@ -102,7 +103,6 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_IR_ExplicitBoxing_getS lean_object* lean_name_append_index_after(lean_object*, lean_object*); uint8_t l_Lean_IR_FnBody_beq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_mkBoxedVersionAux(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_AltCore_mmodifyBody___at_Lean_IR_ExplicitBoxing_visitFnBody___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); lean_object* l_Lean_IR_Decl_updateBody_x21(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_mkBoxedVersionAux___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -4345,7 +4345,7 @@ return x_250; } } } -LEAN_EXPORT lean_object* l_Lean_IR_AltCore_mmodifyBody___at_Lean_IR_ExplicitBoxing_visitFnBody___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_IR_Alt_mmodifyBody___at_Lean_IR_ExplicitBoxing_visitFnBody___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { if (lean_obj_tag(x_2) == 0) @@ -4511,7 +4511,7 @@ x_9 = lean_unsigned_to_nat(0u); x_10 = lean_array_uset(x_3, x_2, x_9); x_11 = l_Array_mapMUnsafe_map___at_Lean_IR_ExplicitBoxing_visitFnBody___spec__2___closed__1; lean_inc(x_4); -x_12 = l_Lean_IR_AltCore_mmodifyBody___at_Lean_IR_ExplicitBoxing_visitFnBody___spec__1(x_11, x_8, x_4, x_5); +x_12 = l_Lean_IR_Alt_mmodifyBody___at_Lean_IR_ExplicitBoxing_visitFnBody___spec__1(x_11, x_8, x_4, x_5); x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); x_14 = lean_ctor_get(x_12, 1); diff --git a/stage0/stdlib/Lean/Compiler/IR/Checker.c b/stage0/stdlib/Lean/Compiler/IR/Checker.c index 1120e6f88b63..e9c5eb8b1e10 100644 --- a/stage0/stdlib/Lean/Compiler/IR/Checker.c +++ b/stage0/stdlib/Lean/Compiler/IR/Checker.c @@ -126,8 +126,8 @@ LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkEqTypes(lean_object*, lean_objec static lean_object* l_Lean_IR_Checker_checkExpr___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkExpr___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkFnBody(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_IR_Alt_body(lean_object*); static lean_object* l_Lean_IR_Checker_checkExpr___closed__1; -lean_object* l_Lean_IR_AltCore_body(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkScalarVar(lean_object*, lean_object*, lean_object*); lean_object* lean_nat_mul(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Checker_getDecl___lambda__1___boxed(lean_object*); @@ -4880,7 +4880,7 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_dec(x_4); x_8 = lean_array_uget(x_1, x_2); -x_9 = l_Lean_IR_AltCore_body(x_8); +x_9 = l_Lean_IR_Alt_body(x_8); lean_dec(x_8); lean_inc(x_5); x_10 = l_Lean_IR_Checker_checkFnBody(x_9, x_5, x_6); diff --git a/stage0/stdlib/Lean/Compiler/IR/EmitC.c b/stage0/stdlib/Lean/Compiler/IR/EmitC.c index beafc9bc3fe4..8c9773dee31d 100644 --- a/stage0/stdlib/Lean/Compiler/IR/EmitC.c +++ b/stage0/stdlib/Lean/Compiler/IR/EmitC.c @@ -471,6 +471,7 @@ static lean_object* l_Lean_IR_EmitC_emitDeclInit___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFnBody___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___lambda__2___closed__3; static lean_object* l_String_foldlAux___at_Lean_IR_EmitC_quoteString___spec__1___closed__7; +lean_object* l_Lean_IR_Alt_body(lean_object*); uint64_t lean_uint64_xor(uint64_t, uint64_t); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitDeclAux___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLn___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -480,7 +481,6 @@ static lean_object* l_Lean_IR_EmitC_emitDec___closed__2; static lean_object* l_Lean_IR_EmitC_toCType___closed__11; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitMarkPersistent(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___lambda__2___closed__17; -lean_object* l_Lean_IR_AltCore_body(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitDec(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitTag___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); @@ -5245,7 +5245,7 @@ lean_inc(x_11); lean_dec(x_9); x_12 = lean_unsigned_to_nat(1u); x_13 = lean_array_fget(x_1, x_12); -x_14 = l_Lean_IR_AltCore_body(x_13); +x_14 = l_Lean_IR_Alt_body(x_13); lean_dec(x_13); lean_ctor_set(x_7, 1, x_14); lean_ctor_set(x_7, 0, x_10); @@ -5269,7 +5269,7 @@ lean_inc(x_19); lean_dec(x_17); x_20 = lean_unsigned_to_nat(1u); x_21 = lean_array_fget(x_1, x_20); -x_22 = l_Lean_IR_AltCore_body(x_21); +x_22 = l_Lean_IR_Alt_body(x_21); lean_dec(x_21); x_23 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_23, 0, x_18); diff --git a/stage0/stdlib/Lean/Compiler/IR/EmitLLVM.c b/stage0/stdlib/Lean/Compiler/IR/EmitLLVM.c index f70256de88e4..c1823cb289ee 100644 --- a/stage0/stdlib/Lean/Compiler/IR/EmitLLVM.c +++ b/stage0/stdlib/Lean/Compiler/IR/EmitLLVM.c @@ -677,6 +677,7 @@ static lean_object* l_Lean_IR_EmitLLVM_emitCase___closed__2; LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitSet___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_llvm_build_store(size_t, size_t, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitDeclInit___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_IR_Alt_body(lean_object*); uint64_t lean_uint64_xor(uint64_t, uint64_t); static lean_object* l_Lean_IR_EmitLLVM_emitSSet___closed__4; LEAN_EXPORT lean_object* l_Lean_IR_getModuleFunctions_go(size_t, size_t, lean_object*, lean_object*); @@ -685,7 +686,6 @@ static lean_object* l_Lean_IR_EmitLLVM_emitMainFn___lambda__3___closed__2; static lean_object* l_Lean_IR_EmitLLVM_emitDeclInit___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_getDecl(size_t); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_buildIfThen_____lambda__1(size_t, size_t, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_IR_AltCore_body(lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); lean_object* l_Lean_IR_mkVarJPMaps(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_getModName___boxed(lean_object*, lean_object*); @@ -42907,7 +42907,7 @@ lean_dec(x_2); x_5 = l_Lean_IR_instInhabitedAlt; x_6 = l_Array_back_x21___rarg(x_5, x_1); x_7 = lean_array_pop(x_1); -x_8 = l_Lean_IR_AltCore_body(x_6); +x_8 = l_Lean_IR_Alt_body(x_6); lean_dec(x_6); x_9 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_9, 0, x_8); @@ -42927,7 +42927,7 @@ lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean x_14 = l_Lean_IR_instInhabitedAlt; x_15 = l_Array_back_x21___rarg(x_14, x_1); x_16 = lean_array_pop(x_1); -x_17 = l_Lean_IR_AltCore_body(x_15); +x_17 = l_Lean_IR_Alt_body(x_15); lean_dec(x_15); x_18 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_18, 0, x_17); diff --git a/stage0/stdlib/Lean/Compiler/IR/EmitUtil.c b/stage0/stdlib/Lean/Compiler/IR/EmitUtil.c index 37f2778ce110..22c34d07307a 100644 --- a/stage0/stdlib/Lean/Compiler/IR/EmitUtil.c +++ b/stage0/stdlib/Lean/Compiler/IR/EmitUtil.c @@ -61,8 +61,8 @@ uint8_t lean_nat_dec_eq(lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); lean_object* l_Lean_Environment_allImportedModuleNames(lean_object*); LEAN_EXPORT uint8_t l_Lean_IR_isTailCallTo(lean_object*, lean_object*); +lean_object* l_Lean_IR_Alt_body(lean_object*); uint64_t lean_uint64_xor(uint64_t, uint64_t); -lean_object* l_Lean_IR_AltCore_body(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_mkVarJPMaps(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_CollectUsedDecls_collectInitDecl(lean_object*, lean_object*, lean_object*); lean_object* lean_nat_mul(lean_object*, lean_object*); @@ -215,7 +215,7 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; size_t x_13; size_t x_14; lean_dec(x_4); x_8 = lean_array_uget(x_1, x_2); -x_9 = l_Lean_IR_AltCore_body(x_8); +x_9 = l_Lean_IR_Alt_body(x_8); lean_dec(x_8); x_10 = l_Lean_IR_CollectUsedDecls_collectFnBody(x_9, x_5, x_6); x_11 = lean_ctor_get(x_10, 0); @@ -2163,7 +2163,7 @@ if (x_5 == 0) { lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; x_6 = lean_array_uget(x_1, x_2); -x_7 = l_Lean_IR_AltCore_body(x_6); +x_7 = l_Lean_IR_Alt_body(x_6); lean_dec(x_6); x_8 = l_Lean_IR_CollectMaps_collectFnBody(x_7, x_4); x_9 = 1; diff --git a/stage0/stdlib/Lean/Compiler/IR/ExpandResetReuse.c b/stage0/stdlib/Lean/Compiler/IR/ExpandResetReuse.c index b2d03bdcbd53..a9bbbc24ebfa 100644 --- a/stage0/stdlib/Lean/Compiler/IR/ExpandResetReuse.c +++ b/stage0/stdlib/Lean/Compiler/IR/ExpandResetReuse.c @@ -105,8 +105,8 @@ LEAN_EXPORT uint8_t l_Lean_IR_ExpandResetReuse_isSelfSSet(lean_object*, lean_obj LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_IR_ExpandResetReuse_mkSlowPath___spec__1(lean_object*, size_t, size_t, lean_object*); lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_IR_ExpandResetReuse_eraseProjIncForAux___spec__1___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_IR_Alt_body(lean_object*); uint64_t lean_uint64_xor(uint64_t, uint64_t); -lean_object* l_Lean_IR_AltCore_body(lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExpandResetReuse_computeProjCounts___boxed(lean_object*); static lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_IR_ExpandResetReuse_eraseProjIncForAux___spec__1___closed__3; @@ -926,7 +926,7 @@ if (x_5 == 0) { lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; x_6 = lean_array_uget(x_1, x_2); -x_7 = l_Lean_IR_AltCore_body(x_6); +x_7 = l_Lean_IR_Alt_body(x_6); lean_dec(x_6); x_8 = l_Lean_IR_ExpandResetReuse_CollectProjMap_collectFnBody(x_7, x_4); x_9 = 1; @@ -1553,7 +1553,7 @@ if (x_5 == 0) { lean_object* x_6; lean_object* x_7; uint8_t x_8; x_6 = lean_array_uget(x_2, x_3); -x_7 = l_Lean_IR_AltCore_body(x_6); +x_7 = l_Lean_IR_Alt_body(x_6); lean_dec(x_6); x_8 = l_Lean_IR_ExpandResetReuse_consumed(x_1, x_7); if (x_8 == 0) diff --git a/stage0/stdlib/Lean/Compiler/IR/FreeVars.c b/stage0/stdlib/Lean/Compiler/IR/FreeVars.c index cd0e2f0523dc..c82240362a0f 100644 --- a/stage0/stdlib/Lean/Compiler/IR/FreeVars.c +++ b/stage0/stdlib/Lean/Compiler/IR/FreeVars.c @@ -94,8 +94,8 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_IR_HasIndex_visitParam LEAN_EXPORT uint8_t l_Lean_IR_HasIndex_visitArgs(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_FreeIndices_seq(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_findCore___at___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_FreeIndices_collectIndex___spec__1___boxed(lean_object*, lean_object*); +lean_object* l_Lean_IR_Alt_body(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_FreeIndices_withParams___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_IR_AltCore_body(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_FreeIndices_collectJP(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_MaxIndex_collectVar___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_FreeIndices_collectAlts___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -634,7 +634,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_MaxInd _start: { lean_object* x_4; lean_object* x_5; -x_4 = l_Lean_IR_AltCore_body(x_2); +x_4 = l_Lean_IR_Alt_body(x_2); x_5 = lean_apply_2(x_1, x_4, x_3); return x_5; } @@ -1751,7 +1751,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_FreeIn _start: { lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_IR_AltCore_body(x_2); +x_5 = l_Lean_IR_Alt_body(x_2); x_6 = lean_apply_3(x_1, x_5, x_3, x_4); return x_6; } @@ -2503,7 +2503,7 @@ if (x_5 == 0) { lean_object* x_6; lean_object* x_7; uint8_t x_8; x_6 = lean_array_uget(x_2, x_3); -x_7 = l_Lean_IR_AltCore_body(x_6); +x_7 = l_Lean_IR_Alt_body(x_6); lean_dec(x_6); x_8 = l_Lean_IR_HasIndex_visitFnBody(x_1, x_7); lean_dec(x_7); diff --git a/stage0/stdlib/Lean/Compiler/IR/LiveVars.c b/stage0/stdlib/Lean/Compiler/IR/LiveVars.c index 0928b8e0aafb..0958689d15d0 100644 --- a/stage0/stdlib/Lean/Compiler/IR/LiveVars.c +++ b/stage0/stdlib/Lean/Compiler/IR/LiveVars.c @@ -64,8 +64,8 @@ uint8_t l_Lean_IR_HasIndex_visitArgs(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_LiveVars_collectFnBody___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_IR_mkLiveVarSet___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at___private_Lean_Compiler_IR_LiveVars_0__Lean_IR_LiveVars_collectJP___spec__1(lean_object*, lean_object*); +lean_object* l_Lean_IR_Alt_body(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instInhabitedLiveVarSet; -lean_object* l_Lean_IR_AltCore_body(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_IsLive_visitJP(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_RBNode_isBlack___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_IsLive_visitExpr___boxed(lean_object*, lean_object*, lean_object*); @@ -215,7 +215,7 @@ if (x_6 == 0) { lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; x_7 = lean_array_uget(x_2, x_3); -x_8 = l_Lean_IR_AltCore_body(x_7); +x_8 = l_Lean_IR_Alt_body(x_7); lean_dec(x_7); x_9 = l_Lean_IR_IsLive_visitFnBody(x_1, x_8, x_5); x_10 = lean_ctor_get(x_9, 0); @@ -6385,7 +6385,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_LiveVars_collectFnBody___lambda__1(lean_objec _start: { lean_object* x_4; lean_object* x_5; -x_4 = l_Lean_IR_AltCore_body(x_2); +x_4 = l_Lean_IR_Alt_body(x_2); x_5 = l_Lean_IR_LiveVars_collectFnBody(x_4, x_1, x_3); return x_5; } diff --git a/stage0/stdlib/Lean/Compiler/IR/NormIds.c b/stage0/stdlib/Lean/Compiler/IR/NormIds.c index 865696d07ea1..cef93b4f5e69 100644 --- a/stage0/stdlib/Lean/Compiler/IR/NormIds.c +++ b/stage0/stdlib/Lean/Compiler/IR/NormIds.c @@ -21,6 +21,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_IR_NormalizeIds_normFn LEAN_EXPORT lean_object* l_Lean_IR_NormalizeIds_normFnBody(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_UniqueIds_checkFnBody(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_UniqueIds_checkDecl(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_Alt_mmodifyBody___at_Lean_IR_NormalizeIds_normFnBody___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); lean_object* l_Lean_RBNode_insert___at_Lean_IR_mkIndexSet___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_IR_NormalizeIds_normFnBody___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -67,9 +68,9 @@ LEAN_EXPORT lean_object* l_Lean_IR_NormalizeIds_withVar___rarg(lean_object*, lea LEAN_EXPORT lean_object* l_Lean_IR_UniqueIds_checkParams___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_IR_MapVars_mapArgs___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_NormalizeIds_withParams(lean_object*); +lean_object* l_Lean_IR_Alt_body(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_NormalizeIds_withJP___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_NormalizeIds_normArgs(lean_object*, lean_object*); -lean_object* l_Lean_IR_AltCore_body(lean_object*); lean_object* l_Lean_RBNode_insert___at_Lean_IR_addVarRename___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_IR_MapVars_mapFnBody___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_findCore___at_Lean_IR_UniqueIds_checkId___spec__1___boxed(lean_object*, lean_object*); @@ -97,7 +98,6 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_IR_NormalizeIds_normAr lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* l_Lean_RBNode_find___at_Lean_IR_VarId_alphaEqv___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_FnBody_replaceVar___lambda__1(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_AltCore_mmodifyBody___at_Lean_IR_NormalizeIds_normFnBody___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_IR_NormalizeIds_normFnBody___spec__5___closed__1; LEAN_EXPORT lean_object* l_Lean_RBNode_findCore___at_Lean_IR_UniqueIds_checkId___spec__1(lean_object* x_1, lean_object* x_2) { _start: @@ -376,7 +376,7 @@ if (x_5 == 0) { lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; x_6 = lean_array_uget(x_1, x_2); -x_7 = l_Lean_IR_AltCore_body(x_6); +x_7 = l_Lean_IR_Alt_body(x_6); lean_dec(x_6); x_8 = l_Lean_IR_UniqueIds_checkFnBody(x_7, x_4); x_9 = lean_ctor_get(x_8, 0); @@ -1982,7 +1982,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_IR_AltCore_mmodifyBody___at_Lean_IR_NormalizeIds_normFnBody___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_IR_Alt_mmodifyBody___at_Lean_IR_NormalizeIds_normFnBody___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { if (lean_obj_tag(x_2) == 0) @@ -2148,7 +2148,7 @@ x_9 = lean_unsigned_to_nat(0u); x_10 = lean_array_uset(x_3, x_2, x_9); x_11 = l_Array_mapMUnsafe_map___at_Lean_IR_NormalizeIds_normFnBody___spec__5___closed__1; lean_inc(x_4); -x_12 = l_Lean_IR_AltCore_mmodifyBody___at_Lean_IR_NormalizeIds_normFnBody___spec__4(x_11, x_8, x_4, x_5); +x_12 = l_Lean_IR_Alt_mmodifyBody___at_Lean_IR_NormalizeIds_normFnBody___spec__4(x_11, x_8, x_4, x_5); x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); x_14 = lean_ctor_get(x_12, 1); diff --git a/stage0/stdlib/Lean/Compiler/IR/PushProj.c b/stage0/stdlib/Lean/Compiler/IR/PushProj.c index 8d03a49740c7..a84763a7219c 100644 --- a/stage0/stdlib/Lean/Compiler/IR/PushProj.c +++ b/stage0/stdlib/Lean/Compiler/IR/PushProj.c @@ -54,7 +54,7 @@ lean_object* l_Lean_IR_Decl_updateBody_x21(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at_Lean_IR_pushProjs___spec__9___at_Lean_IR_pushProjs___spec__10___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_Array_mapMUnsafe_map___at_Lean_IR_FnBody_pushProj___spec__3___boxed(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_IR_AltCore_body(lean_object*); +lean_object* l_Lean_IR_Alt_body(lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); lean_object* l_Array_back_x21___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_IR_pushProjs___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2069,7 +2069,7 @@ lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_obj x_5 = lean_array_uget(x_3, x_2); x_6 = lean_unsigned_to_nat(0u); x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = l_Lean_IR_AltCore_body(x_5); +x_8 = l_Lean_IR_Alt_body(x_5); lean_dec(x_5); x_9 = l_Lean_IR_FnBody_freeIndices(x_8); x_10 = 1; diff --git a/stage0/stdlib/Lean/Compiler/IR/ResetReuse.c b/stage0/stdlib/Lean/Compiler/IR/ResetReuse.c index 0bf2f5706de5..a1f2cbb6a427 100644 --- a/stage0/stdlib/Lean/Compiler/IR/ResetReuse.c +++ b/stage0/stdlib/Lean/Compiler/IR/ResetReuse.c @@ -81,7 +81,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_Rese uint8_t lean_nat_dec_eq(lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); -lean_object* l_Lean_IR_AltCore_body(lean_object*); +lean_object* l_Lean_IR_Alt_body(lean_object*); lean_object* l_Lean_PersistentHashMap_mkCollisionNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_tryS(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -3154,7 +3154,7 @@ if (x_6 == 0) lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; lean_dec(x_4); x_7 = lean_array_uget(x_1, x_2); -x_8 = l_Lean_IR_AltCore_body(x_7); +x_8 = l_Lean_IR_Alt_body(x_7); lean_dec(x_7); x_9 = l_Lean_IR_ResetReuse_collectResets(x_8, x_5); x_10 = lean_ctor_get(x_9, 0); diff --git a/stage0/stdlib/Lean/Compiler/IR/SimpCase.c b/stage0/stdlib/Lean/Compiler/IR/SimpCase.c index c04dbf5ccaff..8ca14a87b06a 100644 --- a/stage0/stdlib/Lean/Compiler/IR/SimpCase.c +++ b/stage0/stdlib/Lean/Compiler/IR/SimpCase.c @@ -47,9 +47,9 @@ lean_object* l_Lean_IR_Decl_updateBody_x21(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_IR_ensureHasDefault___spec__1___boxed(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_IR_Alt_body(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_SimpCase_0__Lean_IR_filterUnreachable___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_SimpCase_0__Lean_IR_getOccsOf___boxed(lean_object*, lean_object*); -lean_object* l_Lean_IR_AltCore_body(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_SimpCase_0__Lean_IR_addDefault(lean_object*); lean_object* l_Array_back_x21___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Compiler_IR_SimpCase_0__Lean_IR_maxOccs___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -146,7 +146,7 @@ lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_obj x_6 = l_Lean_IR_instInhabitedAlt; x_7 = l_Array_back_x21___rarg(x_6, x_1); x_8 = lean_array_pop(x_1); -x_9 = l_Lean_IR_AltCore_body(x_7); +x_9 = l_Lean_IR_Alt_body(x_7); lean_dec(x_7); x_10 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_10, 0, x_9); @@ -190,7 +190,7 @@ else { lean_object* x_11; lean_object* x_12; uint8_t x_13; x_11 = lean_array_fget(x_1, x_6); -x_12 = l_Lean_IR_AltCore_body(x_11); +x_12 = l_Lean_IR_Alt_body(x_11); lean_dec(x_11); lean_inc(x_3); x_13 = l_Lean_IR_FnBody_beq(x_12, x_3); @@ -229,7 +229,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_SimpCase_0__Lean_IR_getOcc lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; x_3 = l_Lean_IR_instInhabitedAlt; x_4 = lean_array_get(x_3, x_1, x_2); -x_5 = l_Lean_IR_AltCore_body(x_4); +x_5 = l_Lean_IR_Alt_body(x_4); lean_dec(x_4); x_6 = lean_unsigned_to_nat(1u); x_7 = lean_nat_add(x_2, x_6); @@ -436,8 +436,8 @@ if (x_6 == 0) { lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; size_t x_11; size_t x_12; x_7 = lean_array_uget(x_2, x_3); -x_8 = l_Lean_IR_AltCore_body(x_7); -x_9 = l_Lean_IR_AltCore_body(x_1); +x_8 = l_Lean_IR_Alt_body(x_7); +x_9 = l_Lean_IR_Alt_body(x_1); x_10 = l_Lean_IR_FnBody_beq(x_8, x_9); x_11 = 1; x_12 = lean_usize_add(x_3, x_11); @@ -535,7 +535,7 @@ if (x_8 == 0) lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; x_9 = lean_unsigned_to_nat(0u); x_10 = lean_nat_dec_lt(x_9, x_2); -x_11 = l_Lean_IR_AltCore_body(x_5); +x_11 = l_Lean_IR_Alt_body(x_5); x_12 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_12, 0, x_11); if (x_10 == 0) @@ -616,7 +616,7 @@ if (x_5 == 0) { lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; size_t x_10; size_t x_11; x_6 = lean_array_uget(x_1, x_2); -x_7 = l_Lean_IR_AltCore_body(x_6); +x_7 = l_Lean_IR_Alt_body(x_6); x_8 = lean_box(13); x_9 = l_Lean_IR_FnBody_beq(x_7, x_8); x_10 = 1; @@ -736,7 +736,7 @@ lean_dec(x_2); lean_dec(x_1); x_13 = lean_array_fget(x_6, x_8); lean_dec(x_6); -x_14 = l_Lean_IR_AltCore_body(x_13); +x_14 = l_Lean_IR_Alt_body(x_13); lean_dec(x_13); return x_14; } diff --git a/stage0/stdlib/Lean/Compiler/IR/Sorry.c b/stage0/stdlib/Lean/Compiler/IR/Sorry.c index dbe8c4973a09..2e68722955ee 100644 --- a/stage0/stdlib/Lean/Compiler/IR/Sorry.c +++ b/stage0/stdlib/Lean/Compiler/IR/Sorry.c @@ -38,7 +38,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_Sorry_visitFndBody(lean_object*, lean_object* LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_IR_updateSorryDep___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Sorry_visitFndBody___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_IR_AltCore_body(lean_object*); +lean_object* l_Lean_IR_Alt_body(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*); @@ -619,7 +619,7 @@ if (x_8 == 0) lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_dec(x_4); x_9 = lean_array_uget(x_1, x_2); -x_10 = l_Lean_IR_AltCore_body(x_9); +x_10 = l_Lean_IR_Alt_body(x_9); lean_dec(x_9); x_11 = l_Lean_IR_Sorry_visitFndBody(x_10, x_5, x_6, x_7); x_12 = lean_ctor_get(x_11, 0); diff --git a/stage0/stdlib/Lean/Compiler/LCNF/AlphaEqv.c b/stage0/stdlib/Lean/Compiler/LCNF/AlphaEqv.c index 6e29d7378d1c..25d7d8350902 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/AlphaEqv.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/AlphaEqv.c @@ -60,7 +60,6 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AlphaEqv_withParams_go___at_Lean_C LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_AlphaEqv_withParams_go___at_Lean_Compiler_LCNF_AlphaEqv_eqv___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_AlphaEqv_eqvAlts___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AlphaEqv_eqvArg___boxed(lean_object*, lean_object*, lean_object*); -uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_264_(lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_AlphaEqv_eqvTypes___lambda__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_AlphaEqv_eqvFVar(lean_object*, lean_object*, lean_object*); @@ -72,6 +71,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AlphaEqv_withParams(lean_object*, lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_AlphaEqv_eqvType(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_qsort_sort___at_Lean_Compiler_LCNF_AlphaEqv_sortAlts___spec__1___lambda__1(lean_object*, lean_object*); +uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_201_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AlphaEqv_withFVar___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Compiler_LCNF_AlphaEqv_eqvFVar___spec__1(lean_object*, lean_object*); uint8_t l_Lean_Name_quickCmp(lean_object*, lean_object*); @@ -940,7 +940,7 @@ x_4 = lean_ctor_get(x_1, 0); x_5 = lean_ctor_get(x_2, 0); lean_inc(x_5); lean_dec(x_2); -x_6 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_264_(x_4, x_5); +x_6 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_201_(x_4, x_5); lean_dec(x_5); return x_6; } diff --git a/stage0/stdlib/Lean/Compiler/LCNF/AuxDeclCache.c b/stage0/stdlib/Lean/Compiler/LCNF/AuxDeclCache.c index a54ac69032c7..c843dc54d71d 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/AuxDeclCache.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/AuxDeclCache.c @@ -47,7 +47,6 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Compiler_L lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_AuxDeclCache___hyg_3____spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_EStateM_pure___rarg(lean_object*, lean_object*); -uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6719_(lean_object*, lean_object*); lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); extern lean_object* l_Lean_Compiler_LCNF_instBEqDecl; lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -94,6 +93,7 @@ lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); static lean_object* l_List_foldl___at_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_AuxDeclCache___hyg_3____spec__9___closed__3; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_insert___at_Lean_Compiler_LCNF_cacheAuxDecl___spec__2___boxed(lean_object*); static lean_object* l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at_Lean_Compiler_LCNF_cacheAuxDecl___spec__1___rarg___closed__1; +uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6660_(lean_object*, lean_object*); size_t lean_usize_land(size_t, size_t); static lean_object* l_Lean_Compiler_LCNF_CacheExtension_insert___at_Lean_Compiler_LCNF_cacheAuxDecl___spec__2___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_AuxDeclCache___hyg_3____spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { @@ -116,7 +116,7 @@ else lean_object* x_9; uint8_t x_10; x_9 = lean_array_fget(x_1, x_4); lean_inc(x_5); -x_10 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6719_(x_5, x_9); +x_10 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6660_(x_5, x_9); if (x_10 == 0) { lean_object* x_11; lean_object* x_12; @@ -188,7 +188,7 @@ 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_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6719_(x_3, x_12); +x_14 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6660_(x_3, x_12); if (x_14 == 0) { lean_object* x_15; @@ -249,7 +249,7 @@ 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_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6719_(x_3, x_27); +x_29 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6660_(x_3, x_27); if (x_29 == 0) { lean_object* x_30; @@ -395,7 +395,7 @@ else lean_object* x_17; uint8_t x_18; x_17 = lean_array_fget(x_5, x_2); lean_inc(x_3); -x_18 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6719_(x_3, x_17); +x_18 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6660_(x_3, x_17); if (x_18 == 0) { lean_object* x_19; lean_object* x_20; @@ -493,7 +493,7 @@ x_19 = lean_ctor_get(x_15, 0); x_20 = lean_ctor_get(x_15, 1); lean_inc(x_19); lean_inc(x_4); -x_21 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6719_(x_4, x_19); +x_21 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6660_(x_4, x_19); if (x_21 == 0) { lean_object* x_22; lean_object* x_23; lean_object* x_24; @@ -529,7 +529,7 @@ lean_inc(x_26); lean_dec(x_15); lean_inc(x_26); lean_inc(x_4); -x_28 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6719_(x_4, x_26); +x_28 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6660_(x_4, x_26); if (x_28 == 0) { lean_object* x_29; lean_object* x_30; lean_object* x_31; @@ -652,7 +652,7 @@ if (lean_is_exclusive(x_57)) { } lean_inc(x_60); lean_inc(x_4); -x_63 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6719_(x_4, x_60); +x_63 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6660_(x_4, x_60); if (x_63 == 0) { lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Basic.c b/stage0/stdlib/Lean/Compiler/LCNF/Basic.c index 00bb75b3fb71..e7b949bc188c 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Basic.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Basic.c @@ -13,18 +13,22 @@ #ifdef __cplusplus extern "C" { #endif +LEAN_EXPORT uint64_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_280_(lean_object*); lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_eqAlt___boxed(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_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6660____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_Cases_extractAlt_x21___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_instantiateValueLevelParams_instLetDecl(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_LetValue_updateProjImp(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instBEqFunDecl; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_instantiateValueLevelParams_instParams(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_848____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDecl_getArity___boxed(lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_LetValue_updateArgsImp___closed__1; static lean_object* l_Lean_Compiler_LCNF_instInhabitedParam___closed__2; static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__2; LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_Decl_noinlineAttr(lean_object*); LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_Decl_alwaysInlineAttr(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_mapCodeM___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp(lean_object*, lean_object*); extern lean_object* l_Lean_instFVarIdSetInhabited; LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateLetDeclCoreImp(lean_object*, lean_object*, lean_object*); @@ -32,19 +36,18 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_DeclValue_forCodeM___rarg(lean_obj LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Code_collectUsed___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_markRecDecls_visit___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqArg____x40_Lean_Compiler_LCNF_Basic___hyg_462_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_eqImp___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_getArity___rarg___boxed(lean_object*); +LEAN_EXPORT uint64_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashArg____x40_Lean_Compiler_LCNF_Basic___hyg_489_(lean_object*); LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_Code_isFun(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_collectUsed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_mapCodeM___rarg___lambda__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6660____boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_LetValue_updateProjImp___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Decl_instantiateParamsLevelParams(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateJmpImp(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___lambda__2(lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_eqImp(lean_object*, lean_object*); lean_object* l_Lean_Level_instantiateParams(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instInhabitedLetValue; +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_getParams___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Code_forM_go___spec__1(lean_object*); uint8_t l___private_Lean_Compiler_ExternAttr_0__Lean_beqExternAttrData____x40_Lean_Compiler_ExternAttr___hyg_382_(lean_object*, lean_object*); uint64_t lean_uint64_of_nat(lean_object*); @@ -53,21 +56,17 @@ LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Param_toArg___boxed(lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_LetValue_updateProjImp___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_attachCodeDecls_go(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instInhabitedFunDeclCore___rarg(lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateFunImp___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_collectUsedAtExpr(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltsImp___spec__1(lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDeclValue____x40_Lean_Compiler_LCNF_Basic___hyg_6307_(lean_object*, lean_object*); lean_object* l_Lean_Expr_lit___override(lean_object*); -static lean_object* l_Lean_Compiler_LCNF_instInhabitedCasesCore___closed__1; lean_object* l_Array_findIdx_x3f_loop___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateFunImp(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1088____boxed(lean_object*); static lean_object* l_Lean_Compiler_LCNF_instBEqDeclValue___closed__1; static lean_object* l_Lean_Compiler_LCNF_instHashableLetValue___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instInhabitedCode; static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateContImp___closed__2; -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instInhabitedFunDeclCore(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Cases_extractAlt_x21___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_DeclValue_mapCodeM___rarg(lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Arg_toLetValue___boxed(lean_object*); @@ -75,53 +74,53 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Code_ LEAN_EXPORT lean_object* l_panic___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_collectType___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltsImp(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateContImp___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1025____boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateContImp(lean_object*, lean_object*); extern lean_object* l_Lean_Compiler_LCNF_erasedExpr; lean_object* l_Lean_Expr_proj___override(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instBEqDeclValue; +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instInhabitedFunDecl; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Code_size_go___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Arg_toLetValue(lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqParam____x40_Lean_Compiler_LCNF_Basic___hyg_53_(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instInhabitedAltCore___rarg(lean_object*); -static lean_object* l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___closed__1; -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6719____spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Cases_extractAlt_x21(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateCasesImp(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashArg____x40_Lean_Compiler_LCNF_Basic___hyg_552____boxed(lean_object*); lean_object* l_Lean_RBNode_insert___at_Lean_NameSet_insert___spec__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_getCode___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Cases_getCtorNames___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_instInhabitedCases___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_collectArgs___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_instInhabitedAltCore___rarg___closed__1; uint64_t lean_string_hash(lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_LetValue_updateConstImp___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_size_go___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_instantiateValueLevelParams_instCode(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Decl_instantiateTypeLevelParams(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instInhabitedAlt; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_forM_go___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6719____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_mapCodeM(lean_object*); LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_eqAlt___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instInhabitedAltCore(lean_object*); static lean_object* l_Lean_Compiler_LCNF_instInhabitedLetDecl___closed__1; static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateJmpImp___closed__2; -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_forCodeM(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Arg_toLetValue___closed__1; uint8_t lean_string_dec_eq(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_LetValue_updateFVarImp___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_sizeLe_go___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Decl_getArity(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDeclValue____x40_Lean_Compiler_LCNF_Basic___hyg_6248____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_LCNF_Cases_extractAlt_x21___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instBEqArg; -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_size(lean_object*); static lean_object* l_Lean_Compiler_LCNF_instBEqFunDecl___closed__1; -LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_instantiateValueLevelParams_instArg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_DeclValue_mapCode(lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_848_(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_eqCases___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint64_t l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1025____spec__1(lean_object*, size_t, size_t, uint64_t); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Compiler_LCNF_markRecDecls_visit___spec__1(lean_object*, lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_785____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_instantiateValueLevelParams_instFunDecl(lean_object*, lean_object*, lean_object*); size_t lean_ptr_addr(lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateReturnImp___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_forM_go___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___lambda__2___boxed(lean_object*); uint64_t l___private_Lean_Expr_0__Lean_hashFVarId____x40_Lean_Expr___hyg_1730_(lean_object*); lean_object* lean_expr_instantiate_rev_range(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instInhabitedDecl; @@ -131,16 +130,15 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_markR LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_LetValue_updateFVarImp(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__1; static lean_object* l_Lean_Compiler_LCNF_instInhabitedParam___closed__3; +static lean_object* l_Lean_Compiler_LCNF_instInhabitedCases___closed__2; static lean_object* l_Lean_Compiler_LCNF_instInhabitedLetValue___closed__1; static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateUnreachImp___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_collectArgs___spec__1(lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetDecl____x40_Lean_Compiler_LCNF_Basic___hyg_1722_(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_forCodeM___boxed(lean_object*, lean_object*); uint8_t l_List_beq___at_Lean_Core_instantiateTypeLevelParams___spec__8(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instBEqLetValue; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instBEqLitValue; -LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_264____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Code_forM_go___spec__1___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_forCodeM___rarg(lean_object*, lean_object*); uint8_t lean_expr_eqv(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_instHashableLitValue___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_attachCodeDecls_go___boxed(lean_object*, lean_object*, lean_object*); @@ -148,24 +146,27 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Decl_isCasesOnParam_x3f_go(lean_ob LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_LetValue_updateArgsImp(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_markRecDecls___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_Code_isDecl(lean_object*); +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6660____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_markRecDecls_visit___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_sizeLe___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateJmpImp___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_markRecDecls_go(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_eqCases(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqArg____x40_Lean_Compiler_LCNF_Basic___hyg_462____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_785____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_getParams(lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_LetValue_updateConstImp___closed__1; static lean_object* l_Lean_Compiler_LCNF_instInhabitedCodeDecl___closed__1; LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_eqCases___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Param_toArg(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_collectParams___spec__1(lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_280____boxed(lean_object*); LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_Code_sizeLe(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_collectType___spec__1___lambda__1(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_LetValue_updateConstImp(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_eqAlt(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_instBEqDecl___closed__1; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_CasesCore_getCtorNames___spec__1(lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT uint64_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_343_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Cases_extractAlt_x21___lambda__2___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_mapCodeM___rarg___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_Decl_inlineAttr(lean_object*); uint8_t l_Lean_Compiler_hasSpecializeAttribute(lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); @@ -175,19 +176,19 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_DeclValue_isCodeAndM(lean_object*) static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateCasesImp___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Decl_alwaysInlineAttr___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_instantiateRangeArgs___spec__1___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_panic___at_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___spec__1___closed__2; static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateFVarImp___closed__1; uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_instantiateValueLevelParams_instExpr(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Decl_isTemplateLike(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateFunImp___closed__1; +LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqArg____x40_Lean_Compiler_LCNF_Basic___hyg_399_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_markRecDecls_visit___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_getCode(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_attachCodeDecls___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_Cases_extractAlt_x21___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_DeclValue_forCodeM(lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6719_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Code_forM_go___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_instBEqLetDecl___closed__1; +LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_785____spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_forM_go___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_instInhabitedParam___closed__1; static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateUnreachImp___closed__1; @@ -198,9 +199,13 @@ LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_L uint64_t l_Lean_Expr_hash(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_instantiateValueLevelParams_instLetValue(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_BinderInfo_isInstImplicit(uint8_t); +static lean_object* l_Lean_Compiler_LCNF_Cases_extractAlt_x21___closed__3; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instHashableArg; +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instInhabitedCases; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_LitValue_toExpr(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_collectParams___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_785____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6660____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_markRecDecls(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instBEqParam; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_DeclValue_size(lean_object*); @@ -216,28 +221,23 @@ static lean_object* l_Lean_Compiler_LCNF_instBEqParam___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instInhabitedArg; lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Code_sizeLe_go___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_getArity___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_instantiateValueLevelParams(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6719____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6719____boxed(lean_object*, lean_object*); lean_object* l_Pi_instInhabited___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_collectParams(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_getArity(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_LetValue_toExpr(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_LetValue_toExpr___spec__1(size_t, size_t, lean_object*); LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_hasLocalInst(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CasesCore_getCtorNames(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_isReturnOf___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instInhabitedDeclValue; lean_object* l_Array_eraseIdx___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__3; uint8_t l___private_Lean_Compiler_InlineAttrs_0__Lean_Compiler_beqInlineAttributeKind____x40_Lean_Compiler_InlineAttrs___hyg_18_(uint8_t, uint8_t); -LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_848____spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instHashableLitValue; +LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashArg____x40_Lean_Compiler_LCNF_Basic___hyg_489____boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_markRecDecls_visit(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateUnreachImp(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqArg____x40_Lean_Compiler_LCNF_Basic___hyg_399____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateParamCoreImp(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instInhabitedCasesCore(lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltImp___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Decl_getArity___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_attachCodeDecls(lean_object*, lean_object*); @@ -245,25 +245,27 @@ lean_object* l_Lean_Meta_isInstance(lean_object*, lean_object*, lean_object*, le LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateReturnImp(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_size_go(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_DeclValue_forCodeM___at_Lean_Compiler_LCNF_markRecDecls_go___spec__1(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1088____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_collectType(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltsImp___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Decl_isCasesOnParam_x3f_go___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_getCode(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instInhabitedCodeDecl; -LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_264_(lean_object*, lean_object*); lean_object* l_Array_findFinIdx_x3f_loop___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDecl_collectUsed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_collectLetValue(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_785____spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateFVarImp___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_instantiateRangeArgs___spec__1(size_t, size_t, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6660____spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_instInhabitedDeclValue___closed__1; static lean_object* l_Lean_Compiler_LCNF_instHashableArg___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Decl_noinlineAttr___boxed(lean_object*); extern lean_object* l_Lean_NameSet_empty; static lean_object* l_panic___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_collectType___spec__1___closed__2; +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_forCodeM(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateReturnImp___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_eqCases___boxed(lean_object*, lean_object*); -static lean_object* l_panic___at_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instantiateRangeArgs___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instantiateRevRangeArgs___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -271,29 +273,27 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Decl_isCasesOnParam_x3f(lean_objec LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CodeDecl_fvarId___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_LetValue_toExpr___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_eqFunDecl___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___lambda__1___boxed(lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_isDecl___boxed(lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_LetValue_updateFVarImp___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateLetImp(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Code_forM_go___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___closed__3; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_forM_go___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint64_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1088_(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1025____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_Cases_extractAlt_x21___lambda__2(lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_eqFunDecl(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_LetValue_updateProjImp___closed__2; -LEAN_EXPORT uint64_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashArg____x40_Lean_Compiler_LCNF_Basic___hyg_552_(lean_object*); -static lean_object* l_panic___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltImp___spec__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_getCode___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Compiler_LCNF_markRecDecls_visit___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_Cases_extractAlt_x21___closed__2; +static lean_object* l_Lean_Compiler_LCNF_instInhabitedAlt___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_size___boxed(lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateLetImp___closed__2; static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_LetValue_updateArgsImp___closed__2; uint8_t l_Lean_NameSet_contains(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_848____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t l_Lean_Name_hash___override(lean_object*); LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_eqImp___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Arg_toExpr(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_forCodeM___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__4; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CodeDecl_fvarId(lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); @@ -301,17 +301,18 @@ uint64_t l_List_foldl___at_Lean_Expr_const___override___spec__1(uint64_t, lean_o lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instInhabitedParam; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_sizeLe_inc___boxed(lean_object*, lean_object*); +static lean_object* l_panic___at_Lean_Compiler_LCNF_Cases_extractAlt_x21___spec__1___closed__1; static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateLetImp___closed__1; static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_collectType___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_DeclValue_isCodeAndM___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_848____spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_785_(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_201_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Cases_getCtorNames(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Code_sizeLe_go___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltImp___closed__1; lean_object* l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_instInhabitedDecl___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_markRecDecls_visit___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_mapCodeM___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_getParams(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Decl_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_forM_go(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateFunDeclCoreImp(lean_object*, lean_object*, lean_object*, lean_object*); @@ -319,32 +320,29 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instBEqLetDecl; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Decl_size___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_forM___rarg(lean_object*, lean_object*, lean_object*); lean_object* lean_array_mk(lean_object*); -LEAN_EXPORT uint64_t l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1088____spec__1(lean_object*, size_t, size_t, uint64_t); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CasesCore_getCtorNames___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Decl_isCasesOnParam_x3f_go___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_instInhabitedDecl___closed__1; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Cases_getCtorNames___spec__1(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltCodeImp(lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_343____boxed(lean_object*); LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_eqFunDecl___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltImp(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Param_toExpr(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_isFun___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_collectArgs(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_mapCodeM(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetDecl____x40_Lean_Compiler_LCNF_Basic___hyg_1659____boxed(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_Compiler_LCNF_instBEqLitValue___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instInhabitedLetDecl; LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqParam____x40_Lean_Compiler_LCNF_Basic___hyg_53____boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_collectUsed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_collectType___spec__1___lambda__1___boxed(lean_object*); -static lean_object* l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_DeclValue_mapCodeM___rarg___lambda__1(lean_object*, lean_object*); lean_object* l_Lean_Expr_instantiateLevelParamsNoCache(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Decl_inlineable___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_forM_go___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDeclValue____x40_Lean_Compiler_LCNF_Basic___hyg_6248_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_DeclValue_mapCodeM(lean_object*); LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_Code_isReturnOf(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_instBEqCode___closed__1; @@ -355,45 +353,44 @@ static lean_object* l_panic___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Comp LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_instantiateValueLevelParams_instLevel(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instantiateRevRangeArgs(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Decl_inlineAttr___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_forCodeM___rarg(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_201____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_BasicAux_0__List_mapMonoMImp___at_Lean_Compiler_LCNF_Code_instantiateValueLevelParams_instLetValue___spec__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Cases_getCtorNames___boxed(lean_object*); LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_eqImp___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_collectParams___boxed(lean_object*, lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instantiateRangeArgs(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDeclValue____x40_Lean_Compiler_LCNF_Basic___hyg_6307____boxed(lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltImp___closed__2; static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_collectType___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_sizeLe_go(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetDecl____x40_Lean_Compiler_LCNF_Basic___hyg_1722____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_instInhabitedLitValue___closed__1; static lean_object* l_Lean_Compiler_LCNF_instBEqArg___closed__1; lean_object* lean_nat_add(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_instInhabitedFunDecl___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_collectArgs___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDecl_getArity(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_sizeLe_inc(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Decl_inlineIfReduceAttr___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_CasesCore_getCtorNames___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Code_forM_go___spec__1___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_848____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetDecl____x40_Lean_Compiler_LCNF_Basic___hyg_1659_(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_instInhabitedCode___closed__1; +LEAN_EXPORT uint64_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1025_(lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_markRecDecls_go___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); lean_object* lean_expr_instantiate_range(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_getParams___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Decl_instantiateParamsLevelParams___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_forM(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateFVarImp(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltsImp___closed__1; lean_object* lean_mk_empty_array_with_capacity(lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6660_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Code_size_go___spec__1(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Decl_isTemplateLike___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6719____spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_hasLocalInst___boxed(lean_object*); -LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___spec__1(lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateCasesImp___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_DeclValue_size___boxed(lean_object*); LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_Decl_isCasesOnParam_x3f_go___lambda__1(lean_object*, lean_object*); @@ -551,36 +548,6 @@ x_3 = l_Lean_Expr_fvar___override(x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_instInhabitedAltCore___rarg___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(0u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instInhabitedAltCore___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 = l_Lean_Compiler_LCNF_instInhabitedAltCore___rarg___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_3); -lean_ctor_set(x_4, 2, x_1); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instInhabitedAltCore(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_instInhabitedAltCore___rarg), 1, 0); -return x_2; -} -} static lean_object* _init_l_Lean_Compiler_LCNF_instInhabitedLitValue___closed__1() { _start: { @@ -599,7 +566,7 @@ x_1 = l_Lean_Compiler_LCNF_instInhabitedLitValue___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_264_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_201_(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -638,11 +605,11 @@ return x_10; } } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_264____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_201____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_264_(x_1, x_2); +x_3 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_201_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -653,7 +620,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_instBEqLitValue___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_264____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_201____boxed), 2, 0); return x_1; } } @@ -665,7 +632,7 @@ x_1 = l_Lean_Compiler_LCNF_instBEqLitValue___closed__1; return x_1; } } -LEAN_EXPORT uint64_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_343_(lean_object* x_1) { +LEAN_EXPORT uint64_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_280_(lean_object* x_1) { _start: { if (lean_obj_tag(x_1) == 0) @@ -688,11 +655,11 @@ return x_9; } } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_343____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_280____boxed(lean_object* x_1) { _start: { uint64_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_343_(x_1); +x_2 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_280_(x_1); lean_dec(x_1); x_3 = lean_box_uint64(x_2); return x_3; @@ -702,7 +669,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_instHashableLitValue___closed__1( _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_343____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_280____boxed), 1, 0); return x_1; } } @@ -771,7 +738,7 @@ x_1 = lean_box(0); return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqArg____x40_Lean_Compiler_LCNF_Basic___hyg_462_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqArg____x40_Lean_Compiler_LCNF_Basic___hyg_399_(lean_object* x_1, lean_object* x_2) { _start: { switch (lean_obj_tag(x_1)) { @@ -827,11 +794,11 @@ return x_12; } } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqArg____x40_Lean_Compiler_LCNF_Basic___hyg_462____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqArg____x40_Lean_Compiler_LCNF_Basic___hyg_399____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqArg____x40_Lean_Compiler_LCNF_Basic___hyg_462_(x_1, x_2); +x_3 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqArg____x40_Lean_Compiler_LCNF_Basic___hyg_399_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -842,7 +809,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_instBEqArg___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqArg____x40_Lean_Compiler_LCNF_Basic___hyg_462____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqArg____x40_Lean_Compiler_LCNF_Basic___hyg_399____boxed), 2, 0); return x_1; } } @@ -854,7 +821,7 @@ x_1 = l_Lean_Compiler_LCNF_instBEqArg___closed__1; return x_1; } } -LEAN_EXPORT uint64_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashArg____x40_Lean_Compiler_LCNF_Basic___hyg_552_(lean_object* x_1) { +LEAN_EXPORT uint64_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashArg____x40_Lean_Compiler_LCNF_Basic___hyg_489_(lean_object* x_1) { _start: { switch (lean_obj_tag(x_1)) { @@ -885,11 +852,11 @@ return x_10; } } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashArg____x40_Lean_Compiler_LCNF_Basic___hyg_552____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashArg____x40_Lean_Compiler_LCNF_Basic___hyg_489____boxed(lean_object* x_1) { _start: { uint64_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashArg____x40_Lean_Compiler_LCNF_Basic___hyg_552_(x_1); +x_2 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashArg____x40_Lean_Compiler_LCNF_Basic___hyg_489_(x_1); lean_dec(x_1); x_3 = lean_box_uint64(x_2); return x_3; @@ -899,7 +866,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_instHashableArg___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashArg____x40_Lean_Compiler_LCNF_Basic___hyg_552____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashArg____x40_Lean_Compiler_LCNF_Basic___hyg_489____boxed), 1, 0); return x_1; } } @@ -1000,7 +967,7 @@ static lean_object* _init_l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_ lean_object* x_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_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__1; x_2 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__2; -x_3 = lean_unsigned_to_nat(69u); +x_3 = lean_unsigned_to_nat(64u); x_4 = lean_unsigned_to_nat(9u); x_5 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -1071,7 +1038,7 @@ static lean_object* _init_l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_ lean_object* x_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_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__1; x_2 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateFVarImp___closed__1; -x_3 = lean_unsigned_to_nat(76u); +x_3 = lean_unsigned_to_nat(71u); x_4 = lean_unsigned_to_nat(9u); x_5 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -1144,7 +1111,7 @@ x_1 = l_Lean_Compiler_LCNF_instInhabitedLetValue___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_848____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_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_785____spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { lean_object* x_8; uint8_t x_9; @@ -1158,7 +1125,7 @@ x_11 = lean_nat_sub(x_6, x_10); lean_dec(x_6); x_12 = lean_array_fget(x_4, x_11); x_13 = lean_array_fget(x_5, x_11); -x_14 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqArg____x40_Lean_Compiler_LCNF_Basic___hyg_462_(x_12, x_13); +x_14 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqArg____x40_Lean_Compiler_LCNF_Basic___hyg_399_(x_12, x_13); lean_dec(x_13); lean_dec(x_12); if (x_14 == 0) @@ -1185,7 +1152,7 @@ return x_17; } } } -LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_848____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_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_785____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) { _start: { lean_object* x_8; uint8_t x_9; @@ -1199,7 +1166,7 @@ x_11 = lean_nat_sub(x_6, x_10); lean_dec(x_6); x_12 = lean_array_fget(x_4, x_11); x_13 = lean_array_fget(x_5, x_11); -x_14 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqArg____x40_Lean_Compiler_LCNF_Basic___hyg_462_(x_12, x_13); +x_14 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqArg____x40_Lean_Compiler_LCNF_Basic___hyg_399_(x_12, x_13); lean_dec(x_13); lean_dec(x_12); if (x_14 == 0) @@ -1226,7 +1193,7 @@ return x_17; } } } -LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_848_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_785_(lean_object* x_1, lean_object* x_2) { _start: { switch (lean_obj_tag(x_1)) { @@ -1237,7 +1204,7 @@ 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_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_264_(x_3, x_4); +x_5 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_201_(x_3, x_4); return x_5; } else @@ -1350,7 +1317,7 @@ return x_34; else { uint8_t x_35; -x_35 = l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_848____spec__1(x_23, x_26, lean_box(0), x_23, x_26, x_31, lean_box(0)); +x_35 = l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_785____spec__1(x_23, x_26, lean_box(0), x_23, x_26, x_31, lean_box(0)); return x_35; } } @@ -1396,7 +1363,7 @@ return x_46; else { uint8_t x_47; -x_47 = l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_848____spec__2(x_38, x_40, lean_box(0), x_38, x_40, x_43, lean_box(0)); +x_47 = l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_785____spec__2(x_38, x_40, lean_box(0), x_38, x_40, x_43, lean_box(0)); return x_47; } } @@ -1411,11 +1378,11 @@ return x_48; } } } -LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_848____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_785____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { uint8_t x_8; lean_object* x_9; -x_8 = l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_848____spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_8 = l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_785____spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); @@ -1424,11 +1391,11 @@ x_9 = lean_box(x_8); return x_9; } } -LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_848____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_785____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { uint8_t x_8; lean_object* x_9; -x_8 = l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_848____spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_8 = l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_785____spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); @@ -1437,11 +1404,11 @@ x_9 = lean_box(x_8); return x_9; } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_848____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_785____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_848_(x_1, x_2); +x_3 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_785_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -1452,7 +1419,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_instBEqLetValue___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_848____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_785____boxed), 2, 0); return x_1; } } @@ -1464,7 +1431,7 @@ x_1 = l_Lean_Compiler_LCNF_instBEqLetValue___closed__1; return x_1; } } -LEAN_EXPORT uint64_t l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1088____spec__1(lean_object* x_1, size_t x_2, size_t x_3, uint64_t x_4) { +LEAN_EXPORT uint64_t l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1025____spec__1(lean_object* x_1, size_t x_2, size_t x_3, uint64_t x_4) { _start: { uint8_t x_5; @@ -1473,7 +1440,7 @@ if (x_5 == 0) { lean_object* x_6; uint64_t x_7; uint64_t x_8; size_t x_9; size_t x_10; x_6 = lean_array_uget(x_1, x_2); -x_7 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashArg____x40_Lean_Compiler_LCNF_Basic___hyg_552_(x_6); +x_7 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashArg____x40_Lean_Compiler_LCNF_Basic___hyg_489_(x_6); lean_dec(x_6); x_8 = lean_uint64_mix_hash(x_4, x_7); x_9 = 1; @@ -1488,7 +1455,7 @@ return x_4; } } } -LEAN_EXPORT uint64_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1088_(lean_object* x_1) { +LEAN_EXPORT uint64_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1025_(lean_object* x_1) { _start: { switch (lean_obj_tag(x_1)) { @@ -1497,7 +1464,7 @@ case 0: lean_object* x_2; uint64_t x_3; uint64_t x_4; uint64_t x_5; x_2 = lean_ctor_get(x_1, 0); x_3 = 0; -x_4 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_343_(x_2); +x_4 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_280_(x_2); x_5 = lean_uint64_mix_hash(x_3, x_4); return x_5; } @@ -1561,7 +1528,7 @@ size_t x_32; size_t x_33; uint64_t x_34; uint64_t x_35; x_32 = 0; x_33 = lean_usize_of_nat(x_26); lean_dec(x_26); -x_34 = l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1088____spec__1(x_19, x_32, x_33, x_23); +x_34 = l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1025____spec__1(x_19, x_32, x_33, x_23); x_35 = lean_uint64_mix_hash(x_25, x_34); return x_35; } @@ -1603,7 +1570,7 @@ size_t x_48; size_t x_49; uint64_t x_50; uint64_t x_51; x_48 = 0; x_49 = lean_usize_of_nat(x_42); lean_dec(x_42); -x_50 = l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1088____spec__1(x_37, x_48, x_49, x_41); +x_50 = l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1025____spec__1(x_37, x_48, x_49, x_41); x_51 = lean_uint64_mix_hash(x_40, x_50); return x_51; } @@ -1612,7 +1579,7 @@ return x_51; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1088____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1025____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { size_t x_5; size_t x_6; uint64_t x_7; uint64_t x_8; lean_object* x_9; @@ -1622,17 +1589,17 @@ x_6 = lean_unbox_usize(x_3); lean_dec(x_3); x_7 = lean_unbox_uint64(x_4); lean_dec(x_4); -x_8 = l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1088____spec__1(x_1, x_5, x_6, x_7); +x_8 = l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1025____spec__1(x_1, x_5, x_6, x_7); lean_dec(x_1); x_9 = lean_box_uint64(x_8); return x_9; } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1088____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1025____boxed(lean_object* x_1) { _start: { uint64_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1088_(x_1); +x_2 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1025_(x_1); lean_dec(x_1); x_3 = lean_box_uint64(x_2); return x_3; @@ -1642,7 +1609,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_instHashableLetValue___closed__1( _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1088____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1025____boxed), 1, 0); return x_1; } } @@ -1717,7 +1684,7 @@ static lean_object* _init_l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_ lean_object* x_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_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__1; x_2 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_LetValue_updateProjImp___closed__1; -x_3 = lean_unsigned_to_nat(97u); +x_3 = lean_unsigned_to_nat(92u); x_4 = lean_unsigned_to_nat(9u); x_5 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -1798,7 +1765,7 @@ static lean_object* _init_l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_ lean_object* x_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_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__1; x_2 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_LetValue_updateConstImp___closed__1; -x_3 = lean_unsigned_to_nat(104u); +x_3 = lean_unsigned_to_nat(99u); x_4 = lean_unsigned_to_nat(9u); x_5 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -1960,7 +1927,7 @@ static lean_object* _init_l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_ lean_object* x_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_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__1; x_2 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_LetValue_updateFVarImp___closed__1; -x_3 = lean_unsigned_to_nat(111u); +x_3 = lean_unsigned_to_nat(106u); x_4 = lean_unsigned_to_nat(9u); x_5 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -2071,7 +2038,7 @@ static lean_object* _init_l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_ lean_object* x_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_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__1; x_2 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_LetValue_updateArgsImp___closed__1; -x_3 = lean_unsigned_to_nat(119u); +x_3 = lean_unsigned_to_nat(114u); x_4 = lean_unsigned_to_nat(9u); x_5 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -2356,7 +2323,7 @@ x_1 = l_Lean_Compiler_LCNF_instInhabitedLetDecl___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetDecl____x40_Lean_Compiler_LCNF_Basic___hyg_1722_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetDecl____x40_Lean_Compiler_LCNF_Basic___hyg_1659_(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; uint8_t x_11; @@ -2398,18 +2365,18 @@ return x_16; else { uint8_t x_17; -x_17 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_848_(x_6, x_10); +x_17 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_785_(x_6, x_10); return x_17; } } } } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetDecl____x40_Lean_Compiler_LCNF_Basic___hyg_1722____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetDecl____x40_Lean_Compiler_LCNF_Basic___hyg_1659____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetDecl____x40_Lean_Compiler_LCNF_Basic___hyg_1722_(x_1, x_2); +x_3 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetDecl____x40_Lean_Compiler_LCNF_Basic___hyg_1659_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -2420,7 +2387,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_instBEqLetDecl___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetDecl____x40_Lean_Compiler_LCNF_Basic___hyg_1722____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetDecl____x40_Lean_Compiler_LCNF_Basic___hyg_1659____boxed), 2, 0); return x_1; } } @@ -2432,100 +2399,124 @@ x_1 = l_Lean_Compiler_LCNF_instBEqLetDecl___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instInhabitedFunDeclCore___rarg(lean_object* x_1) { +static lean_object* _init_l_Lean_Compiler_LCNF_instInhabitedCases___closed__1() { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_2 = lean_box(0); -x_3 = l_Lean_Compiler_LCNF_instInhabitedAltCore___rarg___closed__1; -x_4 = l_Lean_Compiler_LCNF_instInhabitedParam___closed__3; -x_5 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_5, 0, x_2); -lean_ctor_set(x_5, 1, x_2); -lean_ctor_set(x_5, 2, x_3); -lean_ctor_set(x_5, 3, x_4); -lean_ctor_set(x_5, 4, x_1); -return x_5; +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instInhabitedFunDeclCore(lean_object* x_1) { +static lean_object* _init_l_Lean_Compiler_LCNF_instInhabitedCases___closed__2() { _start: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_instInhabitedFunDeclCore___rarg), 1, 0); -return x_2; +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_Compiler_LCNF_instInhabitedParam___closed__3; +x_3 = l_Lean_Compiler_LCNF_instInhabitedCases___closed__1; +x_4 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_1); +lean_ctor_set(x_4, 3, x_3); +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_getArity___rarg(lean_object* x_1) { +static lean_object* _init_l_Lean_Compiler_LCNF_instInhabitedCases() { _start: { -lean_object* x_2; lean_object* x_3; -x_2 = lean_ctor_get(x_1, 2); -x_3 = lean_array_get_size(x_2); -return x_3; +lean_object* x_1; +x_1 = l_Lean_Compiler_LCNF_instInhabitedCases___closed__2; +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_getArity(lean_object* x_1) { +static lean_object* _init_l_Lean_Compiler_LCNF_instInhabitedCode___closed__1() { _start: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_FunDeclCore_getArity___rarg___boxed), 1, 0); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Compiler_LCNF_instInhabitedCases___closed__1; +x_3 = lean_alloc_ctor(3, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_getArity___rarg___boxed(lean_object* x_1) { +static lean_object* _init_l_Lean_Compiler_LCNF_instInhabitedCode() { _start: { -lean_object* x_2; -x_2 = l_Lean_Compiler_LCNF_FunDeclCore_getArity___rarg(x_1); -lean_dec(x_1); -return x_2; +lean_object* x_1; +x_1 = l_Lean_Compiler_LCNF_instInhabitedCode___closed__1; +return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_instInhabitedCasesCore___closed__1() { +static lean_object* _init_l_Lean_Compiler_LCNF_instInhabitedAlt___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_Compiler_LCNF_instInhabitedParam___closed__3; -x_3 = l_Lean_Compiler_LCNF_instInhabitedAltCore___rarg___closed__1; -x_4 = lean_alloc_ctor(0, 4, 0); +x_2 = l_Lean_Compiler_LCNF_instInhabitedCases___closed__1; +x_3 = l_Lean_Compiler_LCNF_instInhabitedCode___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_1); -lean_ctor_set(x_4, 3, x_3); +lean_ctor_set(x_4, 2, x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instInhabitedCasesCore(lean_object* x_1) { +static lean_object* _init_l_Lean_Compiler_LCNF_instInhabitedAlt() { _start: { -lean_object* x_2; -x_2 = l_Lean_Compiler_LCNF_instInhabitedCasesCore___closed__1; -return x_2; +lean_object* x_1; +x_1 = l_Lean_Compiler_LCNF_instInhabitedAlt___closed__1; +return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_instInhabitedCode___closed__1() { +static lean_object* _init_l_Lean_Compiler_LCNF_instInhabitedFunDecl___closed__1() { _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 = lean_box(0); -x_2 = l_Lean_Compiler_LCNF_instInhabitedAltCore___rarg___closed__1; -x_3 = lean_alloc_ctor(3, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; +x_2 = l_Lean_Compiler_LCNF_instInhabitedCases___closed__1; +x_3 = l_Lean_Compiler_LCNF_instInhabitedParam___closed__3; +x_4 = l_Lean_Compiler_LCNF_instInhabitedCode___closed__1; +x_5 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_1); +lean_ctor_set(x_5, 2, x_2); +lean_ctor_set(x_5, 3, x_3); +lean_ctor_set(x_5, 4, x_4); +return x_5; } } -static lean_object* _init_l_Lean_Compiler_LCNF_instInhabitedCode() { +static lean_object* _init_l_Lean_Compiler_LCNF_instInhabitedFunDecl() { _start: { lean_object* x_1; -x_1 = l_Lean_Compiler_LCNF_instInhabitedCode___closed__1; +x_1 = l_Lean_Compiler_LCNF_instInhabitedFunDecl___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_CasesCore_getCtorNames___spec__1(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDecl_getArity(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_ctor_get(x_1, 2); +x_3 = lean_array_get_size(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDecl_getArity___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Compiler_LCNF_FunDecl_getArity(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Cases_getCtorNames___spec__1(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -2561,7 +2552,7 @@ return x_4; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CasesCore_getCtorNames(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Cases_getCtorNames(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; uint8_t x_5; @@ -2594,13 +2585,13 @@ x_9 = 0; x_10 = lean_usize_of_nat(x_3); lean_dec(x_3); x_11 = l_Lean_NameSet_empty; -x_12 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_CasesCore_getCtorNames___spec__1(x_2, x_9, x_10, x_11); +x_12 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Cases_getCtorNames___spec__1(x_2, x_9, x_10, x_11); return x_12; } } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_CasesCore_getCtorNames___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Cases_getCtorNames___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { size_t x_5; size_t x_6; lean_object* x_7; @@ -2608,16 +2599,16 @@ x_5 = lean_unbox_usize(x_2); lean_dec(x_2); x_6 = lean_unbox_usize(x_3); lean_dec(x_3); -x_7 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_CasesCore_getCtorNames___spec__1(x_1, x_5, x_6, x_4); +x_7 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Cases_getCtorNames___spec__1(x_1, x_5, x_6, x_4); lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CasesCore_getCtorNames___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Cases_getCtorNames___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Compiler_LCNF_CasesCore_getCtorNames(x_1); +x_2 = l_Lean_Compiler_LCNF_Cases_getCtorNames(x_1); lean_dec(x_1); return x_2; } @@ -2763,7 +2754,7 @@ x_11 = lean_nat_sub(x_6, x_10); lean_dec(x_6); x_12 = lean_array_fget(x_4, x_11); x_13 = lean_array_fget(x_5, x_11); -x_14 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqArg____x40_Lean_Compiler_LCNF_Basic___hyg_462_(x_12, x_13); +x_14 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqArg____x40_Lean_Compiler_LCNF_Basic___hyg_399_(x_12, x_13); lean_dec(x_13); lean_dec(x_12); if (x_14 == 0) @@ -2809,7 +2800,7 @@ x_6 = lean_ctor_get(x_1, 0); x_7 = lean_ctor_get(x_1, 1); x_8 = lean_ctor_get(x_2, 0); x_9 = lean_ctor_get(x_2, 1); -x_10 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetDecl____x40_Lean_Compiler_LCNF_Basic___hyg_1722_(x_6, x_8); +x_10 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetDecl____x40_Lean_Compiler_LCNF_Basic___hyg_1659_(x_6, x_8); if (x_10 == 0) { uint8_t x_11; @@ -3481,7 +3472,7 @@ x_1 = l_Lean_Compiler_LCNF_instBEqFunDecl___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_getCode(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_getCode(lean_object* x_1) { _start: { if (lean_obj_tag(x_1) == 0) @@ -3500,16 +3491,16 @@ return x_3; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_getCode___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_getCode___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Compiler_LCNF_AltCore_getCode(x_1); +x_2 = l_Lean_Compiler_LCNF_Alt_getCode(x_1); lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_getParams(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_getParams(lean_object* x_1) { _start: { if (lean_obj_tag(x_1) == 0) @@ -3527,16 +3518,16 @@ return x_3; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_getParams___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_getParams___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Compiler_LCNF_AltCore_getParams(x_1); +x_2 = l_Lean_Compiler_LCNF_Alt_getParams(x_1); lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_forCodeM___rarg(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_forCodeM___rarg(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -3559,19 +3550,19 @@ return x_6; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_forCodeM(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_forCodeM(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_AltCore_forCodeM___rarg), 2, 0); +x_3 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Alt_forCodeM___rarg), 2, 0); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_forCodeM___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_forCodeM___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Compiler_LCNF_AltCore_forCodeM(x_1, x_2); +x_3 = l_Lean_Compiler_LCNF_Alt_forCodeM(x_1, x_2); lean_dec(x_2); return x_3; } @@ -3665,20 +3656,11 @@ return x_1; } } } -static lean_object* _init_l_panic___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltImp___spec__1___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Compiler_LCNF_instInhabitedCode; -x_2 = l_Lean_Compiler_LCNF_instInhabitedAltCore___rarg(x_1); -return x_2; -} -} LEAN_EXPORT lean_object* l_panic___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltImp___spec__1(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_panic___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltImp___spec__1___closed__1; +x_2 = l_Lean_Compiler_LCNF_instInhabitedAlt; x_3 = lean_panic_fn(x_2, x_1); return x_3; } @@ -3697,7 +3679,7 @@ static lean_object* _init_l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_ lean_object* x_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_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__1; x_2 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltImp___closed__1; -x_3 = lean_unsigned_to_nat(267u); +x_3 = lean_unsigned_to_nat(268u); x_4 = lean_unsigned_to_nat(9u); x_5 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -3828,7 +3810,7 @@ static lean_object* _init_l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_ lean_object* x_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_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__1; x_2 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltsImp___closed__1; -x_3 = lean_unsigned_to_nat(274u); +x_3 = lean_unsigned_to_nat(275u); x_4 = lean_unsigned_to_nat(9u); x_5 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -3961,7 +3943,7 @@ static lean_object* _init_l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_ lean_object* x_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_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__1; x_2 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateCasesImp___closed__1; -x_3 = lean_unsigned_to_nat(281u); +x_3 = lean_unsigned_to_nat(282u); x_4 = lean_unsigned_to_nat(9u); x_5 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -4232,7 +4214,7 @@ static lean_object* _init_l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_ lean_object* x_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_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__1; x_2 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateLetImp___closed__1; -x_3 = lean_unsigned_to_nat(288u); +x_3 = lean_unsigned_to_nat(289u); x_4 = lean_unsigned_to_nat(9u); x_5 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -4345,7 +4327,7 @@ static lean_object* _init_l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_ lean_object* x_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_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__1; x_2 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateContImp___closed__1; -x_3 = lean_unsigned_to_nat(297u); +x_3 = lean_unsigned_to_nat(298u); x_4 = lean_unsigned_to_nat(9u); x_5 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -4508,7 +4490,7 @@ static lean_object* _init_l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_ lean_object* x_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_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__1; x_2 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateFunImp___closed__1; -x_3 = lean_unsigned_to_nat(305u); +x_3 = lean_unsigned_to_nat(306u); x_4 = lean_unsigned_to_nat(9u); x_5 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -4700,7 +4682,7 @@ static lean_object* _init_l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_ lean_object* x_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_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__1; x_2 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateReturnImp___closed__1; -x_3 = lean_unsigned_to_nat(312u); +x_3 = lean_unsigned_to_nat(313u); x_4 = lean_unsigned_to_nat(9u); x_5 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -4769,7 +4751,7 @@ static lean_object* _init_l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_ lean_object* x_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_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__1; x_2 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateJmpImp___closed__1; -x_3 = lean_unsigned_to_nat(319u); +x_3 = lean_unsigned_to_nat(320u); x_4 = lean_unsigned_to_nat(9u); x_5 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -4880,7 +4862,7 @@ static lean_object* _init_l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_ lean_object* x_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_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__1; x_2 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateUnreachImp___closed__1; -x_3 = lean_unsigned_to_nat(326u); +x_3 = lean_unsigned_to_nat(327u); x_4 = lean_unsigned_to_nat(9u); x_5 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -5239,36 +5221,28 @@ return x_1; } } } -static lean_object* _init_l_panic___at_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___spec__1___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Compiler_LCNF_instInhabitedCasesCore(lean_box(0)); -return x_1; -} -} -static lean_object* _init_l_panic___at_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___spec__1___closed__2() { +static lean_object* _init_l_panic___at_Lean_Compiler_LCNF_Cases_extractAlt_x21___spec__1___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_panic___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltImp___spec__1___closed__1; -x_2 = l_panic___at_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___spec__1___closed__1; +x_1 = l_Lean_Compiler_LCNF_instInhabitedAlt; +x_2 = l_Lean_Compiler_LCNF_instInhabitedCases; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___spec__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_LCNF_Cases_extractAlt_x21___spec__1(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_panic___at_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___spec__1___closed__2; +x_2 = l_panic___at_Lean_Compiler_LCNF_Cases_extractAlt_x21___spec__1___closed__1; x_3 = lean_panic_fn(x_2, x_1); return x_3; } } -LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___lambda__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_Cases_extractAlt_x21___lambda__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -5286,7 +5260,7 @@ return x_5; } } } -LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___lambda__2(lean_object* x_1) { +LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_Cases_extractAlt_x21___lambda__2(lean_object* x_1) { _start: { if (lean_obj_tag(x_1) == 0) @@ -5303,40 +5277,40 @@ return x_3; } } } -static lean_object* _init_l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___closed__1() { +static lean_object* _init_l_Lean_Compiler_LCNF_Cases_extractAlt_x21___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___lambda__2___boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Cases_extractAlt_x21___lambda__2___boxed), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___closed__2() { +static lean_object* _init_l_Lean_Compiler_LCNF_Cases_extractAlt_x21___closed__2() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Compiler.LCNF.CasesCore.extractAlt!", 40, 40); +x_1 = lean_mk_string_unchecked("Lean.Compiler.LCNF.Cases.extractAlt!", 36, 36); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___closed__3() { +static lean_object* _init_l_Lean_Compiler_LCNF_Cases_extractAlt_x21___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 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__1; -x_2 = l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___closed__2; -x_3 = lean_unsigned_to_nat(376u); +x_2 = l_Lean_Compiler_LCNF_Cases_extractAlt_x21___closed__2; +x_3 = lean_unsigned_to_nat(377u); x_4 = lean_unsigned_to_nat(4u); x_5 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___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_Compiler_LCNF_CasesCore_extractAlt_x21(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Cases_extractAlt_x21(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; uint8_t x_4; -x_3 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___lambda__1___boxed), 2, 1); +x_3 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Cases_extractAlt_x21___lambda__1___boxed), 2, 1); lean_closure_set(x_3, 0, x_2); x_4 = !lean_is_exclusive(x_1); if (x_4 == 0) @@ -5351,7 +5325,7 @@ x_10 = l_Array_findFinIdx_x3f_loop___rarg(x_3, x_8, x_9); if (lean_obj_tag(x_10) == 0) { lean_object* x_11; lean_object* x_12; -x_11 = l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___closed__1; +x_11 = l_Lean_Compiler_LCNF_Cases_extractAlt_x21___closed__1; x_12 = l_Array_findFinIdx_x3f_loop___rarg(x_11, x_8, x_9); if (lean_obj_tag(x_12) == 0) { @@ -5361,8 +5335,8 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_13 = l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___closed__3; -x_14 = l_panic___at_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___spec__1(x_13); +x_13 = l_Lean_Compiler_LCNF_Cases_extractAlt_x21___closed__3; +x_14 = l_panic___at_Lean_Compiler_LCNF_Cases_extractAlt_x21___spec__1(x_13); return x_14; } else @@ -5412,7 +5386,7 @@ x_28 = l_Array_findFinIdx_x3f_loop___rarg(x_3, x_26, x_27); if (lean_obj_tag(x_28) == 0) { lean_object* x_29; lean_object* x_30; -x_29 = l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___closed__1; +x_29 = l_Lean_Compiler_LCNF_Cases_extractAlt_x21___closed__1; x_30 = l_Array_findFinIdx_x3f_loop___rarg(x_29, x_26, x_27); if (lean_obj_tag(x_30) == 0) { @@ -5421,8 +5395,8 @@ lean_dec(x_26); lean_dec(x_25); lean_dec(x_24); lean_dec(x_23); -x_31 = l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___closed__3; -x_32 = l_panic___at_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___spec__1(x_31); +x_31 = l_Lean_Compiler_LCNF_Cases_extractAlt_x21___closed__3; +x_32 = l_panic___at_Lean_Compiler_LCNF_Cases_extractAlt_x21___spec__1(x_31); return x_32; } else @@ -5465,28 +5439,28 @@ return x_42; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Cases_extractAlt_x21___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___lambda__1(x_1, x_2); +x_3 = l_Lean_Compiler_LCNF_Cases_extractAlt_x21___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_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___lambda__2___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Cases_extractAlt_x21___lambda__2___boxed(lean_object* x_1) { _start: { uint8_t x_2; lean_object* x_3; -x_2 = l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___lambda__2(x_1); +x_2 = l_Lean_Compiler_LCNF_Cases_extractAlt_x21___lambda__2(x_1); lean_dec(x_1); x_3 = lean_box(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_mapCodeM___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_mapCodeM___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; @@ -5501,26 +5475,26 @@ x_7 = lean_apply_2(x_5, lean_box(0), x_6); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_mapCodeM___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_mapCodeM___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; x_4 = lean_ctor_get(x_1, 1); lean_inc(x_4); -x_5 = l_Lean_Compiler_LCNF_AltCore_getCode(x_2); +x_5 = l_Lean_Compiler_LCNF_Alt_getCode(x_2); x_6 = lean_apply_1(x_3, x_5); -x_7 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_AltCore_mapCodeM___rarg___lambda__1), 3, 2); +x_7 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Alt_mapCodeM___rarg___lambda__1), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); x_8 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_6, x_7); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_mapCodeM(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_mapCodeM(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_AltCore_mapCodeM___rarg), 3, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Alt_mapCodeM___rarg), 3, 0); return x_2; } } @@ -5630,7 +5604,7 @@ if (x_5 == 0) { lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; size_t x_11; size_t x_12; x_6 = lean_array_uget(x_1, x_2); -x_7 = l_Lean_Compiler_LCNF_AltCore_getCode(x_6); +x_7 = l_Lean_Compiler_LCNF_Alt_getCode(x_6); lean_dec(x_6); x_8 = lean_unsigned_to_nat(1u); x_9 = lean_nat_add(x_4, x_8); @@ -5817,7 +5791,7 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_dec(x_5); x_8 = lean_array_uget(x_2, x_3); -x_9 = l_Lean_Compiler_LCNF_AltCore_getCode(x_8); +x_9 = l_Lean_Compiler_LCNF_Alt_getCode(x_8); lean_dec(x_8); x_10 = l_Lean_Compiler_LCNF_Code_sizeLe_go(x_1, x_9, x_6); lean_dec(x_9); @@ -6213,7 +6187,7 @@ lean_dec(x_6); x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); x_9 = lean_array_uget(x_3, x_4); -x_10 = l_Lean_Compiler_LCNF_AltCore_getCode(x_9); +x_10 = l_Lean_Compiler_LCNF_Alt_getCode(x_9); lean_dec(x_9); lean_inc(x_2); lean_inc(x_1); @@ -7394,7 +7368,7 @@ x_1 = l_Lean_Compiler_LCNF_instInhabitedDeclValue___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDeclValue____x40_Lean_Compiler_LCNF_Basic___hyg_6307_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDeclValue____x40_Lean_Compiler_LCNF_Basic___hyg_6248_(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -7433,11 +7407,11 @@ return x_10; } } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDeclValue____x40_Lean_Compiler_LCNF_Basic___hyg_6307____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDeclValue____x40_Lean_Compiler_LCNF_Basic___hyg_6248____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDeclValue____x40_Lean_Compiler_LCNF_Basic___hyg_6307_(x_1, x_2); +x_3 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDeclValue____x40_Lean_Compiler_LCNF_Basic___hyg_6248_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -7448,7 +7422,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_instBEqDeclValue___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDeclValue____x40_Lean_Compiler_LCNF_Basic___hyg_6307____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDeclValue____x40_Lean_Compiler_LCNF_Basic___hyg_6248____boxed), 2, 0); return x_1; } } @@ -7688,7 +7662,7 @@ x_1 = lean_box(0); x_2 = lean_box(0); x_3 = lean_box(0); x_4 = l_Lean_Compiler_LCNF_instInhabitedDecl___closed__1; -x_5 = l_Lean_Compiler_LCNF_instInhabitedAltCore___rarg___closed__1; +x_5 = l_Lean_Compiler_LCNF_instInhabitedCases___closed__1; x_6 = l_Lean_Compiler_LCNF_instInhabitedDeclValue___closed__1; x_7 = 0; x_8 = lean_alloc_ctor(0, 6, 2); @@ -7711,7 +7685,7 @@ x_1 = l_Lean_Compiler_LCNF_instInhabitedDecl___closed__2; return x_1; } } -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6719____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6660____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -7758,7 +7732,7 @@ return x_10; } } } -LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6719____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_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6660____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) { _start: { lean_object* x_8; uint8_t x_9; @@ -7799,7 +7773,7 @@ return x_17; } } } -LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6719_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6660_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; uint8_t x_17; lean_object* x_18; uint8_t x_19; uint8_t x_26; @@ -7913,7 +7887,7 @@ return x_35; else { uint8_t x_36; -x_36 = l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6719____spec__2(x_6, x_14, lean_box(0), x_6, x_14, x_32, lean_box(0)); +x_36 = l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6660____spec__2(x_6, x_14, lean_box(0), x_6, x_14, x_32, lean_box(0)); lean_dec(x_14); lean_dec(x_6); if (x_36 == 0) @@ -7929,7 +7903,7 @@ return x_37; else { uint8_t x_38; -x_38 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDeclValue____x40_Lean_Compiler_LCNF_Basic___hyg_6307_(x_7, x_15); +x_38 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDeclValue____x40_Lean_Compiler_LCNF_Basic___hyg_6248_(x_7, x_15); lean_dec(x_15); lean_dec(x_7); if (x_38 == 0) @@ -7999,7 +7973,7 @@ if (x_9 == 0) if (x_17 == 0) { uint8_t x_21; -x_21 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6719____spec__1(x_10, x_18); +x_21 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6660____spec__1(x_10, x_18); return x_21; } else @@ -8024,7 +7998,7 @@ return x_23; else { uint8_t x_24; -x_24 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6719____spec__1(x_10, x_18); +x_24 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6660____spec__1(x_10, x_18); return x_24; } } @@ -8032,20 +8006,20 @@ return x_24; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6719____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6660____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6719____spec__1(x_1, x_2); +x_3 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6660____spec__1(x_1, x_2); x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6719____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6660____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { uint8_t x_8; lean_object* x_9; -x_8 = l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6719____spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_8 = l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6660____spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); @@ -8054,11 +8028,11 @@ x_9 = lean_box(x_8); return x_9; } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6719____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6660____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6719_(x_1, x_2); +x_3 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6660_(x_1, x_2); x_4 = lean_box(x_3); return x_4; } @@ -8067,7 +8041,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_instBEqDecl___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6719____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6660____boxed), 2, 0); return x_1; } } @@ -8724,7 +8698,7 @@ static lean_object* _init_l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_ lean_object* x_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_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__1; x_2 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_collectType___closed__1; -x_3 = lean_unsigned_to_nat(658u); +x_3 = lean_unsigned_to_nat(659u); x_4 = lean_unsigned_to_nat(27u); x_5 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -9059,7 +9033,7 @@ lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_collectUsed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDecl_collectUsed(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; @@ -9232,7 +9206,7 @@ lean_inc(x_34); x_35 = lean_ctor_get(x_1, 1); lean_inc(x_35); lean_dec(x_1); -x_36 = l_Lean_Compiler_LCNF_FunDeclCore_collectUsed(x_34, x_2); +x_36 = l_Lean_Compiler_LCNF_FunDecl_collectUsed(x_34, x_2); x_1 = x_35; x_2 = x_36; goto _start; @@ -9308,7 +9282,7 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; size_t x_13; size_t x_14; lean_dec(x_5); x_8 = lean_array_uget(x_2, x_3); -x_9 = l_Lean_Compiler_LCNF_AltCore_getCode(x_8); +x_9 = l_Lean_Compiler_LCNF_Alt_getCode(x_8); lean_dec(x_8); x_10 = l_Lean_Compiler_LCNF_markRecDecls_visit(x_1, x_9, x_6); x_11 = lean_ctor_get(x_10, 0); @@ -9978,8 +9952,6 @@ l_Lean_Compiler_LCNF_instBEqParam___closed__1 = _init_l_Lean_Compiler_LCNF_instB lean_mark_persistent(l_Lean_Compiler_LCNF_instBEqParam___closed__1); l_Lean_Compiler_LCNF_instBEqParam = _init_l_Lean_Compiler_LCNF_instBEqParam(); lean_mark_persistent(l_Lean_Compiler_LCNF_instBEqParam); -l_Lean_Compiler_LCNF_instInhabitedAltCore___rarg___closed__1 = _init_l_Lean_Compiler_LCNF_instInhabitedAltCore___rarg___closed__1(); -lean_mark_persistent(l_Lean_Compiler_LCNF_instInhabitedAltCore___rarg___closed__1); l_Lean_Compiler_LCNF_instInhabitedLitValue___closed__1 = _init_l_Lean_Compiler_LCNF_instInhabitedLitValue___closed__1(); lean_mark_persistent(l_Lean_Compiler_LCNF_instInhabitedLitValue___closed__1); l_Lean_Compiler_LCNF_instInhabitedLitValue = _init_l_Lean_Compiler_LCNF_instInhabitedLitValue(); @@ -10052,12 +10024,24 @@ l_Lean_Compiler_LCNF_instBEqLetDecl___closed__1 = _init_l_Lean_Compiler_LCNF_ins lean_mark_persistent(l_Lean_Compiler_LCNF_instBEqLetDecl___closed__1); l_Lean_Compiler_LCNF_instBEqLetDecl = _init_l_Lean_Compiler_LCNF_instBEqLetDecl(); lean_mark_persistent(l_Lean_Compiler_LCNF_instBEqLetDecl); -l_Lean_Compiler_LCNF_instInhabitedCasesCore___closed__1 = _init_l_Lean_Compiler_LCNF_instInhabitedCasesCore___closed__1(); -lean_mark_persistent(l_Lean_Compiler_LCNF_instInhabitedCasesCore___closed__1); +l_Lean_Compiler_LCNF_instInhabitedCases___closed__1 = _init_l_Lean_Compiler_LCNF_instInhabitedCases___closed__1(); +lean_mark_persistent(l_Lean_Compiler_LCNF_instInhabitedCases___closed__1); +l_Lean_Compiler_LCNF_instInhabitedCases___closed__2 = _init_l_Lean_Compiler_LCNF_instInhabitedCases___closed__2(); +lean_mark_persistent(l_Lean_Compiler_LCNF_instInhabitedCases___closed__2); +l_Lean_Compiler_LCNF_instInhabitedCases = _init_l_Lean_Compiler_LCNF_instInhabitedCases(); +lean_mark_persistent(l_Lean_Compiler_LCNF_instInhabitedCases); l_Lean_Compiler_LCNF_instInhabitedCode___closed__1 = _init_l_Lean_Compiler_LCNF_instInhabitedCode___closed__1(); lean_mark_persistent(l_Lean_Compiler_LCNF_instInhabitedCode___closed__1); l_Lean_Compiler_LCNF_instInhabitedCode = _init_l_Lean_Compiler_LCNF_instInhabitedCode(); lean_mark_persistent(l_Lean_Compiler_LCNF_instInhabitedCode); +l_Lean_Compiler_LCNF_instInhabitedAlt___closed__1 = _init_l_Lean_Compiler_LCNF_instInhabitedAlt___closed__1(); +lean_mark_persistent(l_Lean_Compiler_LCNF_instInhabitedAlt___closed__1); +l_Lean_Compiler_LCNF_instInhabitedAlt = _init_l_Lean_Compiler_LCNF_instInhabitedAlt(); +lean_mark_persistent(l_Lean_Compiler_LCNF_instInhabitedAlt); +l_Lean_Compiler_LCNF_instInhabitedFunDecl___closed__1 = _init_l_Lean_Compiler_LCNF_instInhabitedFunDecl___closed__1(); +lean_mark_persistent(l_Lean_Compiler_LCNF_instInhabitedFunDecl___closed__1); +l_Lean_Compiler_LCNF_instInhabitedFunDecl = _init_l_Lean_Compiler_LCNF_instInhabitedFunDecl(); +lean_mark_persistent(l_Lean_Compiler_LCNF_instInhabitedFunDecl); l_Lean_Compiler_LCNF_instInhabitedCodeDecl___closed__1 = _init_l_Lean_Compiler_LCNF_instInhabitedCodeDecl___closed__1(); lean_mark_persistent(l_Lean_Compiler_LCNF_instInhabitedCodeDecl___closed__1); l_Lean_Compiler_LCNF_instInhabitedCodeDecl = _init_l_Lean_Compiler_LCNF_instInhabitedCodeDecl(); @@ -10070,8 +10054,6 @@ l_Lean_Compiler_LCNF_instBEqFunDecl___closed__1 = _init_l_Lean_Compiler_LCNF_ins lean_mark_persistent(l_Lean_Compiler_LCNF_instBEqFunDecl___closed__1); l_Lean_Compiler_LCNF_instBEqFunDecl = _init_l_Lean_Compiler_LCNF_instBEqFunDecl(); lean_mark_persistent(l_Lean_Compiler_LCNF_instBEqFunDecl); -l_panic___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltImp___spec__1___closed__1 = _init_l_panic___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltImp___spec__1___closed__1(); -lean_mark_persistent(l_panic___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltImp___spec__1___closed__1); l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltImp___closed__1 = _init_l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltImp___closed__1(); lean_mark_persistent(l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltImp___closed__1); l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltImp___closed__2 = _init_l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltImp___closed__2(); @@ -10108,16 +10090,14 @@ l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateUnreachImp___cl lean_mark_persistent(l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateUnreachImp___closed__1); l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateUnreachImp___closed__2 = _init_l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateUnreachImp___closed__2(); lean_mark_persistent(l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateUnreachImp___closed__2); -l_panic___at_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___spec__1___closed__1 = _init_l_panic___at_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___spec__1___closed__1(); -lean_mark_persistent(l_panic___at_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___spec__1___closed__1); -l_panic___at_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___spec__1___closed__2 = _init_l_panic___at_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___spec__1___closed__2(); -lean_mark_persistent(l_panic___at_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___spec__1___closed__2); -l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___closed__1 = _init_l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___closed__1(); -lean_mark_persistent(l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___closed__1); -l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___closed__2 = _init_l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___closed__2(); -lean_mark_persistent(l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___closed__2); -l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___closed__3 = _init_l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___closed__3(); -lean_mark_persistent(l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___closed__3); +l_panic___at_Lean_Compiler_LCNF_Cases_extractAlt_x21___spec__1___closed__1 = _init_l_panic___at_Lean_Compiler_LCNF_Cases_extractAlt_x21___spec__1___closed__1(); +lean_mark_persistent(l_panic___at_Lean_Compiler_LCNF_Cases_extractAlt_x21___spec__1___closed__1); +l_Lean_Compiler_LCNF_Cases_extractAlt_x21___closed__1 = _init_l_Lean_Compiler_LCNF_Cases_extractAlt_x21___closed__1(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Cases_extractAlt_x21___closed__1); +l_Lean_Compiler_LCNF_Cases_extractAlt_x21___closed__2 = _init_l_Lean_Compiler_LCNF_Cases_extractAlt_x21___closed__2(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Cases_extractAlt_x21___closed__2); +l_Lean_Compiler_LCNF_Cases_extractAlt_x21___closed__3 = _init_l_Lean_Compiler_LCNF_Cases_extractAlt_x21___closed__3(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Cases_extractAlt_x21___closed__3); l_Lean_Compiler_LCNF_instInhabitedDeclValue___closed__1 = _init_l_Lean_Compiler_LCNF_instInhabitedDeclValue___closed__1(); lean_mark_persistent(l_Lean_Compiler_LCNF_instInhabitedDeclValue___closed__1); l_Lean_Compiler_LCNF_instInhabitedDeclValue = _init_l_Lean_Compiler_LCNF_instInhabitedDeclValue(); diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Bind.c b/stage0/stdlib/Lean/Compiler/LCNF/Bind.c index 7c4365e038d3..5a772cc46aef 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Bind.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Bind.c @@ -14,10 +14,11 @@ extern "C" { #endif lean_object* l_Lean_Compiler_LCNF_mkAuxLetDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_FunDeclCore_isEtaExpandCandidate(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDecl_etaExpand(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_CompilerM_codeBind_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* l_Lean_Compiler_LCNF_Code_inferType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_FunDecl_isEtaExpandCandidate(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Decl_etaExpand(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CompilerM_codeBind_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_updateFunDeclImp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -49,7 +50,6 @@ static lean_object* l_Lean_throwError___at_Lean_Compiler_LCNF_CompilerM_codeBind LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_bind___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_eraseCode(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CompilerM_codeBind_go___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_isEtaExpandCandidate___boxed(lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_bind(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instMonadCodeBindReaderT___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -69,6 +69,7 @@ static lean_object* l_Lean_Compiler_LCNF_etaExpandCore___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instMonadCodeBindReaderT(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CompilerM_codeBind_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* l_Lean_Compiler_LCNF_eraseParam(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDecl_isEtaExpandCandidate___boxed(lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); lean_object* l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_LCNF_CompilerM_codeBind_go___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -89,7 +90,6 @@ uint8_t lean_usize_dec_lt(size_t, size_t); lean_object* l_Lean_Compiler_LCNF_mkAuxParam(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_isEtaExpandCandidateCore(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_etaExpand(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_Compiler_LCNF_Code_bind___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: @@ -1979,7 +1979,7 @@ x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_FunDeclCore_isEtaExpandCandidate(lean_object* x_1) { +LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_FunDecl_isEtaExpandCandidate(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; @@ -1993,11 +1993,11 @@ lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_isEtaExpandCandidate___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDecl_isEtaExpandCandidate___boxed(lean_object* x_1) { _start: { uint8_t x_2; lean_object* x_3; -x_2 = l_Lean_Compiler_LCNF_FunDeclCore_isEtaExpandCandidate(x_1); +x_2 = l_Lean_Compiler_LCNF_FunDecl_isEtaExpandCandidate(x_1); x_3 = lean_box(x_2); return x_3; } @@ -2397,7 +2397,7 @@ return x_23; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_etaExpand(lean_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_Compiler_LCNF_FunDecl_etaExpand(lean_object* x_1, 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; diff --git a/stage0/stdlib/Lean/Compiler/LCNF/CSE.c b/stage0/stdlib/Lean/Compiler/LCNF/CSE.c index ac973a8055c0..33be57499cf7 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/CSE.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/CSE.c @@ -71,6 +71,7 @@ lean_object* l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_norm static lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Compiler_LCNF_CSE_addEntry___spec__2___closed__3; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Compiler_LCNF_Code_cse_go___spec__2(lean_object*, lean_object*); lean_object* lean_nat_div(lean_object*, lean_object*); +lean_object* l_Lean_Compiler_LCNF_FunDecl_toExpr(lean_object*, lean_object*); lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_cse_goFunDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_CSE___hyg_1124____closed__13; @@ -157,7 +158,6 @@ LEAN_EXPORT lean_object* l___private_Init_Data_Array_BasicAux_0__mapMonoMImp___a lean_object* lean_nat_add(lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CSE_replaceFun(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Compiler_LCNF_FunDeclCore_toExpr(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_cse(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_CSE___hyg_1124____closed__18; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); @@ -4104,7 +4104,7 @@ lean_inc(x_147); lean_dec(x_145); x_148 = l_Lean_Compiler_LCNF_Code_cse_go___closed__1; lean_inc(x_146); -x_149 = l_Lean_Compiler_LCNF_FunDeclCore_toExpr(x_146, x_148); +x_149 = l_Lean_Compiler_LCNF_FunDecl_toExpr(x_146, x_148); x_150 = lean_st_ref_get(x_3, x_147); x_151 = lean_ctor_get(x_150, 0); lean_inc(x_151); diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Check.c b/stage0/stdlib/Lean/Compiler/LCNF/Check.c index c4b24cb0099a..d9fa6b6b73f3 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Check.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Check.c @@ -99,7 +99,6 @@ static lean_object* l_Lean_Compiler_LCNF_Check_run___rarg___closed__5; lean_object* l_Lean_Compiler_LCNF_getParam(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Check_checkFVar___closed__1; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Compiler_LCNF_Check_checkAppArgs___spec__2___boxed(lean_object**); -uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetDecl____x40_Lean_Compiler_LCNF_Basic___hyg_1722_(lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Check_checkAppArgs___closed__1; static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Compiler_LCNF_Check_checkAppArgs___spec__2___lambda__3___closed__8; @@ -164,7 +163,6 @@ static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_Check_c static lean_object* l_Lean_Compiler_LCNF_Check_addFVarId___closed__1; static lean_object* l_Lean_Compiler_LCNF_Check_checkFunDecl___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Check_withParams(lean_object*); -lean_object* l_Lean_Compiler_LCNF_FunDeclCore_getArity___rarg(lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Compiler_LCNF_checkDeadLocalDecls___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_LCNF_Check_checkCases___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Check_run___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -286,8 +284,10 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Check_withJp___rarg(lean_object*, static lean_object* l_Lean_Compiler_LCNF_Check_check___closed__4; lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Check_withFVarId___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_Compiler_LCNF_FunDecl_getArity(lean_object*); static lean_object* l_Lean_Compiler_LCNF_Check_checkLetDecl___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Check_checkParams___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetDecl____x40_Lean_Compiler_LCNF_Basic___hyg_1659_(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Check_checkParam___closed__4; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_checkDeadLocalDecls_visitParams___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); @@ -3295,7 +3295,7 @@ if (x_13 == 0) lean_object* x_14; lean_object* x_15; uint8_t x_16; x_14 = lean_ctor_get(x_12, 0); x_15 = lean_ctor_get(x_12, 1); -x_16 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetDecl____x40_Lean_Compiler_LCNF_Basic___hyg_1722_(x_1, x_14); +x_16 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetDecl____x40_Lean_Compiler_LCNF_Basic___hyg_1659_(x_1, x_14); lean_dec(x_14); if (x_16 == 0) { @@ -3333,7 +3333,7 @@ x_26 = lean_ctor_get(x_12, 1); lean_inc(x_26); lean_inc(x_25); lean_dec(x_12); -x_27 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetDecl____x40_Lean_Compiler_LCNF_Basic___hyg_1722_(x_1, x_25); +x_27 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetDecl____x40_Lean_Compiler_LCNF_Basic___hyg_1659_(x_1, x_25); lean_dec(x_25); if (x_27 == 0) { @@ -7793,7 +7793,7 @@ lean_inc(x_126); x_127 = lean_ctor_get(x_125, 1); lean_inc(x_127); lean_dec(x_125); -x_128 = l_Lean_Compiler_LCNF_FunDeclCore_getArity___rarg(x_126); +x_128 = l_Lean_Compiler_LCNF_FunDecl_getArity(x_126); x_129 = lean_array_get_size(x_122); x_130 = lean_nat_dec_eq(x_128, x_129); if (x_130 == 0) @@ -7966,7 +7966,7 @@ lean_inc(x_168); x_169 = lean_ctor_get(x_167, 1); lean_inc(x_169); lean_dec(x_167); -x_170 = l_Lean_Compiler_LCNF_FunDeclCore_getArity___rarg(x_168); +x_170 = l_Lean_Compiler_LCNF_FunDecl_getArity(x_168); x_171 = lean_array_get_size(x_164); x_172 = lean_nat_dec_eq(x_170, x_171); if (x_172 == 0) diff --git a/stage0/stdlib/Lean/Compiler/LCNF/CompilerM.c b/stage0/stdlib/Lean/Compiler/LCNF/CompilerM.c index e30e5a76cc7f..c6610f968652 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/CompilerM.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/CompilerM.c @@ -46,6 +46,7 @@ static lean_object* l_Lean_Compiler_LCNF_mkFreshJpName___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_normFunDecl___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_eraseCodeDecls___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_normParam___at_Lean_Compiler_LCNF_normCodeImp___spec__4(uint8_t); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDecl_updateValue___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_replaceFVars___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_normParams(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_withPhase___rarg(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -54,6 +55,7 @@ static lean_object* l_Lean_Compiler_LCNF_instMonadCompilerM___closed__3; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_normArgs___rarg___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*); lean_object* l_Lean_PersistentHashMap_find_x3f___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_normLetDecl___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDecl_update_x27___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_instInhabitedCacheExtension___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_normFVar(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_withNormFVarResult___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -183,7 +185,6 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_normCode___rarg___lambda__1___boxe LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_updateParamImp___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_eraseParams(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instMonadFVarSubstStateOfMonadLift(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_update_x27___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_mkReturnErased___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_normFunDecl___rarg(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Compiler_LCNF_addFVarSubst___spec__2(lean_object*); @@ -214,6 +215,7 @@ lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_eraseParam___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_EStateM_pure___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_LCNF_getLetDecl___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDecl_update_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instMonadFVarSubstOfMonadLift___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_normArgsImp___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_eraseParams___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -223,6 +225,7 @@ LEAN_EXPORT lean_object* l_Lean_mkFreshId___at_Lean_Compiler_LCNF_mkParam___spec lean_object* l_Array_mapMono(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_inBasePhase(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_instMonadCompilerM___closed__5; +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDecl_updateValue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_normArgsImp___lambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_List_foldl___at_Lean_Compiler_LCNF_CacheExtension_register___spec__1___rarg___closed__1; lean_object* l_panic___at_Lean_Expr_appFn_x21___spec__1(lean_object*); @@ -242,7 +245,6 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_mkFreshBinderName(lean_object*, le static lean_object* l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_normFVarImp___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_normParam___at_Lean_Compiler_LCNF_normFunDeclImp___spec__2___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_normCode___rarg(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_update_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instMonadCompilerM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_BasicAux_0__mapMonoMImp___at_Lean_Compiler_LCNF_normCodeImp___spec__5(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instInhabitedCacheExtension___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -407,7 +409,6 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_LCNF_getFunDecl___ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_find_x3f___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_withNormFVarResult(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_BasicAux_0__mapMonoMImp___at_Lean_Compiler_LCNF_normCodeImp___spec__7(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_updateValue(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* l_Lean_MessageData_ofName(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_normExprImp_goApp___boxed(lean_object*, lean_object*, lean_object*); @@ -426,7 +427,6 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_normParam___at_Lean_Compiler_LCNF_ static lean_object* l_Lean_Compiler_LCNF_instInhabitedCacheExtension___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_normExpr___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_instAddMessageContextCompilerM___closed__3; -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_updateValue___boxed(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_Compiler_LCNF_getType___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_eraseCodeDecls(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); @@ -8328,7 +8328,7 @@ lean_dec(x_5); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_update_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_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDecl_update_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_9; lean_object* x_10; @@ -8338,11 +8338,11 @@ x_10 = l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_updateFunD return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_update_x27___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDecl_update_x27___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; -x_9 = l_Lean_Compiler_LCNF_FunDeclCore_update_x27(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_Compiler_LCNF_FunDecl_update_x27(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); @@ -8350,7 +8350,7 @@ lean_dec(x_4); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_updateValue(lean_object* x_1, lean_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_Compiler_LCNF_FunDecl_updateValue(lean_object* x_1, lean_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; @@ -8362,11 +8362,11 @@ x_10 = l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_updateFunD return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_updateValue___boxed(lean_object* x_1, lean_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_Compiler_LCNF_FunDecl_updateValue___boxed(lean_object* x_1, lean_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_Compiler_LCNF_FunDeclCore_updateValue(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_8 = l_Lean_Compiler_LCNF_FunDecl_updateValue(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); diff --git a/stage0/stdlib/Lean/Compiler/LCNF/DeclHash.c b/stage0/stdlib/Lean/Compiler/LCNF/DeclHash.c index d4844b329c5d..ee3abff69687 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/DeclHash.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/DeclHash.c @@ -23,6 +23,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_DeclHash_0__Lean_Compile LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_hashAlts___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint64_t l_Lean_Compiler_LCNF_hashParams(lean_object*); uint64_t l___private_Lean_Compiler_InlineAttrs_0__Lean_Compiler_hashInlineAttributeKind____x40_Lean_Compiler_InlineAttrs___hyg_36_(uint8_t); +uint64_t l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1025____spec__1(lean_object*, size_t, size_t, uint64_t); static lean_object* l_Lean_Compiler_LCNF_instHashableDecl___closed__1; uint64_t l___private_Lean_Expr_0__Lean_hashFVarId____x40_Lean_Expr___hyg_1730_(lean_object*); size_t lean_usize_of_nat(lean_object*); @@ -40,19 +41,18 @@ LEAN_EXPORT uint64_t l___private_Lean_Compiler_LCNF_DeclHash_0__Lean_Compiler_LC LEAN_EXPORT uint64_t l_Lean_Compiler_LCNF_instHashableCode(lean_object*); static lean_object* l_Lean_Compiler_LCNF_instHashableDeclValue___closed__1; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -uint64_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1088_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_hashAlt___boxed(lean_object*); LEAN_EXPORT uint64_t l_Lean_Compiler_LCNF_instHashableParam(lean_object*); uint64_t l_Lean_Name_hash___override(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_DeclHash_0__Lean_Compiler_LCNF_hashDecl____x40_Lean_Compiler_LCNF_DeclHash___hyg_319____boxed(lean_object*); LEAN_EXPORT uint64_t l_Lean_Compiler_LCNF_hashAlts(lean_object*); LEAN_EXPORT uint64_t l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_hashAlts___spec__1(lean_object*, size_t, size_t, uint64_t); -uint64_t l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1088____spec__1(lean_object*, size_t, size_t, uint64_t); size_t lean_usize_add(size_t, size_t); lean_object* lean_array_uget(lean_object*, size_t); lean_object* lean_array_get_size(lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); uint64_t l___private_Lean_Compiler_ExternAttr_0__Lean_hashExternAttrData____x40_Lean_Compiler_ExternAttr___hyg_456_(lean_object*); +uint64_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1025_(lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Compiler_LCNF_DeclHash_0__Lean_Compiler_LCNF_hashDecl____x40_Lean_Compiler_LCNF_DeclHash___hyg_319____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT uint64_t l_Lean_Compiler_LCNF_instHashableParam(lean_object* x_1) { _start: @@ -240,7 +240,7 @@ x_6 = lean_ctor_get(x_2, 2); x_7 = l_Lean_Expr_hash(x_6); x_8 = lean_uint64_mix_hash(x_5, x_7); x_9 = lean_ctor_get(x_2, 3); -x_10 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1088_(x_9); +x_10 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1025_(x_9); x_11 = l_Lean_Compiler_LCNF_hashCode(x_3); x_12 = lean_uint64_mix_hash(x_10, x_11); x_13 = lean_uint64_mix_hash(x_8, x_12); @@ -280,7 +280,7 @@ size_t x_24; size_t x_25; uint64_t x_26; uint64_t x_27; x_24 = 0; x_25 = lean_usize_of_nat(x_18); lean_dec(x_18); -x_26 = l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1088____spec__1(x_15, x_24, x_25, x_17); +x_26 = l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1025____spec__1(x_15, x_24, x_25, x_17); x_27 = lean_uint64_mix_hash(x_16, x_26); return x_27; } diff --git a/stage0/stdlib/Lean/Compiler/LCNF/DependsOn.c b/stage0/stdlib/Lean/Compiler/LCNF/DependsOn.c index 67bf6076a4d3..cf134ff5ae5c 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/DependsOn.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/DependsOn.c @@ -26,12 +26,12 @@ LEAN_EXPORT uint8_t l_Lean_Expr_hasAnyFVar_visit___at___private_Lean_Compiler_LC lean_object* l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_LetDecl_dependsOn___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_FunDecl_dependsOn(lean_object*, lean_object*); -lean_object* l_Lean_Compiler_LCNF_AltCore_getCode(lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_DependsOn_0__Lean_Compiler_LCNF_argDepOn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Compiler_LCNF_DependsOn_0__Lean_Compiler_LCNF_letValueDepOn___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_hasAnyFVar_visit___at___private_Lean_Compiler_LCNF_DependsOn_0__Lean_Compiler_LCNF_typeDepOn___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Compiler_LCNF_DependsOn_0__Lean_Compiler_LCNF_depOn___spec__1(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CodeDecl_dependsOn___boxed(lean_object*, lean_object*); +lean_object* l_Lean_Compiler_LCNF_Alt_getCode(lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_DependsOn_0__Lean_Compiler_LCNF_fvarDepOn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_dependsOn___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_DependsOn_0__Lean_Compiler_LCNF_letValueDepOn(lean_object*, lean_object*); @@ -465,7 +465,7 @@ if (x_5 == 0) { lean_object* x_6; lean_object* x_7; uint8_t x_8; x_6 = lean_array_uget(x_1, x_2); -x_7 = l_Lean_Compiler_LCNF_AltCore_getCode(x_6); +x_7 = l_Lean_Compiler_LCNF_Alt_getCode(x_6); lean_dec(x_6); x_8 = l___private_Lean_Compiler_LCNF_DependsOn_0__Lean_Compiler_LCNF_depOn(x_7, x_4); lean_dec(x_7); diff --git a/stage0/stdlib/Lean/Compiler/LCNF/ElimDead.c b/stage0/stdlib/Lean/Compiler/LCNF/ElimDead.c index 63e1bfe107fd..6882b0a5bd28 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/ElimDead.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/ElimDead.c @@ -54,7 +54,6 @@ 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_panic___at_Lean_Compiler_LCNF_collectLocalDeclsType_go___spec__5(lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Compiler_LCNF_collectLocalDeclsType_go___spec__3(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Compiler_LCNF_AltCore_getCode(lean_object*); static lean_object* l_Lean_Compiler_LCNF_Code_elimDead___closed__3; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Compiler_LCNF_collectLocalDeclsType_go___spec__1___boxed(lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); @@ -65,6 +64,7 @@ static lean_object* l_Lean_Compiler_LCNF_collectLocalDeclsType_go___closed__3; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Compiler_LCNF_collectLocalDeclsType_go___spec__4(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_collectLocalDeclsArg(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_ElimDead_elimDead___closed__1; +lean_object* l_Lean_Compiler_LCNF_Alt_getCode(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_collectLocalDeclsType_go___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_collectLocalDeclsType(lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); @@ -1801,7 +1801,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ElimDead_elimDead___lambda__1(lean _start: { lean_object* x_8; lean_object* x_9; -x_8 = l_Lean_Compiler_LCNF_AltCore_getCode(x_1); +x_8 = l_Lean_Compiler_LCNF_Alt_getCode(x_1); x_9 = l_Lean_Compiler_LCNF_ElimDead_elimDead(x_8, x_2, x_3, x_4, x_5, x_6, x_7); if (lean_obj_tag(x_9) == 0) { diff --git a/stage0/stdlib/Lean/Compiler/LCNF/ElimDeadBranches.c b/stage0/stdlib/Lean/Compiler/LCNF/ElimDeadBranches.c index 732d0ebede64..d4c72583c00c 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/ElimDeadBranches.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/ElimDeadBranches.c @@ -240,7 +240,6 @@ LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Compiler_LC lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_UnreachableBranches_elimDead_go___spec__7___closed__4; uint8_t l_Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Compiler_LCNF_AltCore_getCode(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_UnreachableBranches_interpCode___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_ElimDeadBranches___hyg_6391____closed__10; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_UnreachableBranches_initFn____x40_Lean_Compiler_LCNF_ElimDeadBranches___hyg_1933____lambda__4___boxed(lean_object*, lean_object*); @@ -310,6 +309,7 @@ double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); static lean_object* l_Lean_Compiler_LCNF_UnreachableBranches_Value_merge___closed__1; static lean_object* l_Lean_Compiler_LCNF_UnreachableBranches_initFn____x40_Lean_Compiler_LCNF_ElimDeadBranches___hyg_1933____closed__12; static lean_object* l_Lean_Compiler_LCNF_Decl_elimDeadBranches___closed__1; +lean_object* l_Lean_Compiler_LCNF_Alt_getCode(lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_ElimDeadBranches_0__Lean_Compiler_LCNF_UnreachableBranches_reprValue____x40_Lean_Compiler_LCNF_ElimDeadBranches___hyg_44____closed__7; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_UnreachableBranches_updateFunDeclParamsAssignment(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_zipWithAux___at_Lean_Compiler_LCNF_UnreachableBranches_Value_merge___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -10539,7 +10539,7 @@ if (x_12 == 0) lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; lean_dec(x_4); x_13 = lean_array_uget(x_1, x_2); -x_14 = l_Lean_Compiler_LCNF_AltCore_getCode(x_13); +x_14 = l_Lean_Compiler_LCNF_Alt_getCode(x_13); lean_dec(x_13); x_15 = l___private_Lean_Compiler_LCNF_ElimDeadBranches_0__Lean_Compiler_LCNF_UnreachableBranches_resetNestedFunDeclParams(x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11); x_16 = lean_ctor_get(x_15, 0); @@ -15849,7 +15849,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_Unreacha _start: { lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_9 = l_Lean_Compiler_LCNF_AltCore_getCode(x_1); +x_9 = l_Lean_Compiler_LCNF_Alt_getCode(x_1); x_10 = l_Lean_Compiler_LCNF_eraseCode(x_9, x_4, x_5, x_6, x_7, x_8); lean_dec(x_9); x_11 = !lean_is_exclusive(x_10); diff --git a/stage0/stdlib/Lean/Compiler/LCNF/FVarUtil.c b/stage0/stdlib/Lean/Compiler/LCNF/FVarUtil.c index 64c282df3688..196e3eada625 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/FVarUtil.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/FVarUtil.c @@ -21,7 +21,6 @@ LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_LCNF_Expr_forFVarM___spec__2 lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_LetValue_updateProjImp(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_instTraverseFVarCode___closed__2; static lean_object* l_Lean_Compiler_LCNF_Code_mapFVarM___rarg___lambda__14___closed__1; -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_forCodeM___at_Lean_Compiler_LCNF_Code_forFVarM___spec__4___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_FunDecl_forFVarM___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instTraverseFVarCodeDecl___lambda__3(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Code_mapFVarM___rarg___lambda__14___closed__2; @@ -32,11 +31,11 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Code_ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Expr_forFVarM(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_LetValue_mapFVarM___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Code_forFVarM___spec__2(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_forCodeM___at_Lean_Compiler_LCNF_Code_forFVarM___spec__4___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_LCNF_Expr_mapFVarM___spec__2___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Code_forFVarM___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Expr_mapFVarM___rarg___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_forFVarM___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_mapCodeM___at_Lean_Compiler_LCNF_Code_mapFVarM___spec__4___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_allFVarM___at_Lean_Compiler_LCNF_allFVar___spec__1(lean_object*); static lean_object* l_Lean_Compiler_LCNF_Code_mapFVarM___rarg___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_LetDecl_mapFVarM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -81,7 +80,6 @@ static lean_object* l_Lean_Compiler_LCNF_anyFVarM_go___rarg___lambda__2___closed LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instTraverseFVarLetDecl; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instTraverseFVarCode___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_instTraverseFVarParam___closed__1; -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_forCodeM___at_Lean_Compiler_LCNF_Code_forFVarM___spec__4(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_LetValue_mapFVarM___spec__1___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Expr_mapFVarM___rarg___lambda__4___closed__1; static lean_object* l_Lean_Compiler_LCNF_Code_mapFVarM___rarg___lambda__8___closed__1; @@ -90,6 +88,7 @@ LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_LCNF_Expr_mapFVarM___spec__2 LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_LetValue_forFVarM___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instTraverseFVarLetDecl___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_instTraverseFVarCode___closed__1; +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_mapCodeM___at_Lean_Compiler_LCNF_Code_mapFVarM___spec__4(lean_object*); static lean_object* l_Lean_Compiler_LCNF_Expr_mapFVarM___rarg___closed__1; static lean_object* l_Lean_Compiler_LCNF_instTraverseFVarFunDecl___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instTraverseFVarAlt___lambda__3(lean_object*, lean_object*); @@ -169,7 +168,6 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instTraverseFVarFunDecl___lambda__ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_FunDecl_forFVarM___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_instTraverseFVarArg___closed__2; static lean_object* l_Lean_Compiler_LCNF_Code_mapFVarM___rarg___lambda__1___closed__3; -lean_object* l_Lean_Compiler_LCNF_AltCore_getCode(lean_object*); static lean_object* l_Lean_Compiler_LCNF_instTraverseFVarLetValue___closed__3; static lean_object* l_Lean_Compiler_LCNF_instTraverseFVarLetDecl___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_forFVarM___rarg___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -180,6 +178,7 @@ LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_LCNF_Expr_mapFVarM___spec__3 LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_mapFVarM___rarg___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Code_forFVarM___spec__2___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instTraverseFVarExpr___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_forCodeM___at_Lean_Compiler_LCNF_Code_forFVarM___spec__4___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instTraverseFVarAlt___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_instTraverseFVarFunDecl___closed__3; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instTraverseFVarAlt___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -212,7 +211,6 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_instTrav LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instTraverseFVarCode; static lean_object* l_Lean_Compiler_LCNF_instTraverseFVarExpr___closed__2; static lean_object* l_Lean_Compiler_LCNF_instTraverseFVarCodeDecl___closed__3; -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_mapCodeM___at_Lean_Compiler_LCNF_Code_mapFVarM___spec__4(lean_object*); lean_object* l_OptionT_instMonad___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_anyFVarM_go___rarg___lambda__1(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_LCNF_Expr_mapFVarM___spec__3___rarg(lean_object*, lean_object*); @@ -222,6 +220,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_anyFVarM___rarg(lean_object*, lean LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_allFVarM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_forFVarM___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_LCNF_Expr_forFVarM___spec__1___rarg(lean_object*, lean_object*); +lean_object* l_Lean_Compiler_LCNF_Alt_getCode(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instTraverseFVarAlt___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Expr_mapFVarM___rarg___lambda__6___closed__2; LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_LCNF_Expr_mapFVarM___spec__1___rarg(lean_object*, lean_object*); @@ -244,6 +243,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_LetDecl_forFVarM___rarg(lean_objec LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instTraverseFVarArg___lambda__2(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_Compiler_LCNF_instTraverseFVarAlt; +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_forCodeM___at_Lean_Compiler_LCNF_Code_forFVarM___spec__4___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Expr_mapFVarM___rarg___closed__3; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_Code_mapFVarM___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -276,7 +276,6 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_FunDecl_ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_FunDecl_mapFVarM___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instTraverseFVarCode___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_Code_mapFVarM___spec__1(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_mapCodeM___at_Lean_Compiler_LCNF_Code_mapFVarM___spec__4___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_forFVarM___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_instTraverseFVarAlt___spec__2___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_LetValue_mapFVarM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -290,7 +289,6 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instTraverseFVarExpr; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_mapFVarM___rarg___lambda__13(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Arg_forFVarM(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_Code_mapFVarM___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_mapCodeM___at_Lean_Compiler_LCNF_Code_mapFVarM___spec__4___rarg___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_instTraverseFVarAlt___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_anyFVarM___rarg___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_forFVarM___rarg___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -335,10 +333,12 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_anyFVarM___at_Lean_Compiler_LCNF_a lean_object* l_Lean_Expr_lam___override(lean_object*, lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Compiler_LCNF_instTraverseFVarFunDecl___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDecl_mapFVarM___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_mapCodeM___at_Lean_Compiler_LCNF_Code_mapFVarM___spec__4___rarg___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_forFVarM___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_Code_mapFVarM___spec__5___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_Code_mapFVarM___spec__2___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_LetValue_forFVarM___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_forCodeM___at_Lean_Compiler_LCNF_Code_forFVarM___spec__4(lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_allFVar___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Code_mapFVarM___rarg___lambda__10___closed__1; @@ -2713,7 +2713,7 @@ x_2 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_ return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_mapCodeM___at_Lean_Compiler_LCNF_Code_mapFVarM___spec__4___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_mapCodeM___at_Lean_Compiler_LCNF_Code_mapFVarM___spec__4___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; @@ -2728,26 +2728,26 @@ x_7 = lean_apply_2(x_5, lean_box(0), x_6); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_mapCodeM___at_Lean_Compiler_LCNF_Code_mapFVarM___spec__4___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_mapCodeM___at_Lean_Compiler_LCNF_Code_mapFVarM___spec__4___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; x_4 = lean_ctor_get(x_1, 1); lean_inc(x_4); -x_5 = l_Lean_Compiler_LCNF_AltCore_getCode(x_2); +x_5 = l_Lean_Compiler_LCNF_Alt_getCode(x_2); x_6 = lean_apply_1(x_3, x_5); -x_7 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_AltCore_mapCodeM___at_Lean_Compiler_LCNF_Code_mapFVarM___spec__4___rarg___lambda__1), 3, 2); +x_7 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Alt_mapCodeM___at_Lean_Compiler_LCNF_Code_mapFVarM___spec__4___rarg___lambda__1), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); x_8 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_6, x_7); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_mapCodeM___at_Lean_Compiler_LCNF_Code_mapFVarM___spec__4(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_mapCodeM___at_Lean_Compiler_LCNF_Code_mapFVarM___spec__4(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_AltCore_mapCodeM___at_Lean_Compiler_LCNF_Code_mapFVarM___spec__4___rarg), 3, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Alt_mapCodeM___at_Lean_Compiler_LCNF_Code_mapFVarM___spec__4___rarg), 3, 0); return x_2; } } @@ -2797,7 +2797,7 @@ lean_closure_set(x_15, 0, x_1); lean_closure_set(x_15, 1, x_2); lean_closure_set(x_15, 2, x_3); lean_inc(x_2); -x_16 = l_Lean_Compiler_LCNF_AltCore_mapCodeM___at_Lean_Compiler_LCNF_Code_mapFVarM___spec__4___rarg(x_2, x_11, x_15); +x_16 = l_Lean_Compiler_LCNF_Alt_mapCodeM___at_Lean_Compiler_LCNF_Code_mapFVarM___spec__4___rarg(x_2, x_11, x_15); x_17 = lean_box_usize(x_5); x_18 = lean_box_usize(x_4); x_19 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_Code_mapFVarM___spec__5___rarg___lambda__1___boxed), 7, 6); @@ -2842,7 +2842,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_Code_mapFVarM___rarg___lambda__1_ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Compiler_LCNF_Code_mapFVarM___rarg___lambda__1___closed__1; x_2 = l_Lean_Compiler_LCNF_Code_mapFVarM___rarg___lambda__1___closed__2; -x_3 = lean_unsigned_to_nat(288u); +x_3 = lean_unsigned_to_nat(289u); x_4 = lean_unsigned_to_nat(9u); x_5 = l_Lean_Compiler_LCNF_Expr_mapFVarM___rarg___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -2988,7 +2988,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_Code_mapFVarM___rarg___lambda__3_ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Compiler_LCNF_Code_mapFVarM___rarg___lambda__1___closed__1; x_2 = l_Lean_Compiler_LCNF_Code_mapFVarM___rarg___lambda__3___closed__1; -x_3 = lean_unsigned_to_nat(305u); +x_3 = lean_unsigned_to_nat(306u); x_4 = lean_unsigned_to_nat(9u); x_5 = l_Lean_Compiler_LCNF_Expr_mapFVarM___rarg___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -3297,7 +3297,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_Code_mapFVarM___rarg___lambda__8_ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Compiler_LCNF_Code_mapFVarM___rarg___lambda__1___closed__1; x_2 = l_Lean_Compiler_LCNF_Code_mapFVarM___rarg___lambda__8___closed__1; -x_3 = lean_unsigned_to_nat(319u); +x_3 = lean_unsigned_to_nat(320u); x_4 = lean_unsigned_to_nat(9u); x_5 = l_Lean_Compiler_LCNF_Expr_mapFVarM___rarg___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -3443,7 +3443,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_Code_mapFVarM___rarg___lambda__10 lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Compiler_LCNF_Code_mapFVarM___rarg___lambda__1___closed__1; x_2 = l_Lean_Compiler_LCNF_Code_mapFVarM___rarg___lambda__10___closed__1; -x_3 = lean_unsigned_to_nat(281u); +x_3 = lean_unsigned_to_nat(282u); x_4 = lean_unsigned_to_nat(9u); x_5 = l_Lean_Compiler_LCNF_Expr_mapFVarM___rarg___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -3781,7 +3781,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_Code_mapFVarM___rarg___lambda__13 lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Compiler_LCNF_Code_mapFVarM___rarg___lambda__1___closed__1; x_2 = l_Lean_Compiler_LCNF_Code_mapFVarM___rarg___lambda__13___closed__1; -x_3 = lean_unsigned_to_nat(312u); +x_3 = lean_unsigned_to_nat(313u); x_4 = lean_unsigned_to_nat(9u); x_5 = l_Lean_Compiler_LCNF_Expr_mapFVarM___rarg___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -3867,7 +3867,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_Code_mapFVarM___rarg___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; x_1 = l_Lean_Compiler_LCNF_Code_mapFVarM___rarg___lambda__1___closed__1; x_2 = l_Lean_Compiler_LCNF_Code_mapFVarM___rarg___lambda__14___closed__1; -x_3 = lean_unsigned_to_nat(326u); +x_3 = lean_unsigned_to_nat(327u); x_4 = lean_unsigned_to_nat(9u); x_5 = l_Lean_Compiler_LCNF_Expr_mapFVarM___rarg___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -4390,7 +4390,7 @@ x_2 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LC return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_forCodeM___at_Lean_Compiler_LCNF_Code_forFVarM___spec__4___rarg(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_forCodeM___at_Lean_Compiler_LCNF_Code_forFVarM___spec__4___rarg(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -4413,11 +4413,11 @@ return x_6; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_forCodeM___at_Lean_Compiler_LCNF_Code_forFVarM___spec__4(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_forCodeM___at_Lean_Compiler_LCNF_Code_forFVarM___spec__4(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_AltCore_forCodeM___at_Lean_Compiler_LCNF_Code_forFVarM___spec__4___rarg), 2, 0); +x_3 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Alt_forCodeM___at_Lean_Compiler_LCNF_Code_forFVarM___spec__4___rarg), 2, 0); return x_3; } } @@ -4448,7 +4448,7 @@ lean_inc(x_1); x_10 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Code_forFVarM___rarg), 3, 2); lean_closure_set(x_10, 0, x_1); lean_closure_set(x_10, 1, x_2); -x_11 = l_Lean_Compiler_LCNF_AltCore_forCodeM___at_Lean_Compiler_LCNF_Code_forFVarM___spec__4___rarg(x_9, x_10); +x_11 = l_Lean_Compiler_LCNF_Alt_forCodeM___at_Lean_Compiler_LCNF_Code_forFVarM___spec__4___rarg(x_9, x_10); x_12 = lean_box_usize(x_4); x_13 = lean_box_usize(x_5); x_14 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Code_forFVarM___spec__5___rarg___lambda__1___boxed), 6, 5); @@ -4978,11 +4978,11 @@ x_9 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Code_forFVarM___spec__3_ return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_forCodeM___at_Lean_Compiler_LCNF_Code_forFVarM___spec__4___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_forCodeM___at_Lean_Compiler_LCNF_Code_forFVarM___spec__4___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Compiler_LCNF_AltCore_forCodeM___at_Lean_Compiler_LCNF_Code_forFVarM___spec__4(x_1, x_2); +x_3 = l_Lean_Compiler_LCNF_Alt_forCodeM___at_Lean_Compiler_LCNF_Code_forFVarM___spec__4(x_1, x_2); lean_dec(x_2); return x_3; } diff --git a/stage0/stdlib/Lean/Compiler/LCNF/FixedParams.c b/stage0/stdlib/Lean/Compiler/LCNF/FixedParams.c index ed27d860c851..1e7a2e7d1a77 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/FixedParams.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/FixedParams.c @@ -65,7 +65,6 @@ static lean_object* l_Lean_Compiler_LCNF_FixedParams_instHashableAbsValue___clos LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Compiler_LCNF_FixedParams_evalApp___spec__3___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_FixedParams_evalApp___spec__10___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FixedParams_evalApp(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Compiler_LCNF_AltCore_getCode(lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Compiler_LCNF_FixedParams_evalApp___spec__8___at_Lean_Compiler_LCNF_FixedParams_evalApp___spec__9(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_FixedParams_evalApp___closed__1; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Compiler_LCNF_FixedParams_evalApp___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -81,6 +80,7 @@ LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Compiler_LCNF_FixedParams_ev LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Compiler_LCNF_FixedParams_evalApp___spec__13(uint8_t, lean_object*, size_t, size_t); uint8_t l_Lean_Expr_isErased(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FixedParams_evalFVar(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Compiler_LCNF_Alt_getCode(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Compiler_LCNF_FixedParams_inMutualBlock___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_mkFixedParamsMap___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); @@ -2690,7 +2690,7 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_dec(x_4); x_8 = lean_array_uget(x_1, x_2); -x_9 = l_Lean_Compiler_LCNF_AltCore_getCode(x_8); +x_9 = l_Lean_Compiler_LCNF_Alt_getCode(x_8); lean_dec(x_8); lean_inc(x_5); x_10 = l_Lean_Compiler_LCNF_FixedParams_evalCode(x_9, x_5, x_6); diff --git a/stage0/stdlib/Lean/Compiler/LCNF/FloatLetIn.c b/stage0/stdlib/Lean/Compiler/LCNF/FloatLetIn.c index a3409e3aaf30..5bd4f6066601 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/FloatLetIn.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/FloatLetIn.c @@ -106,6 +106,7 @@ static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_FloatL lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Compiler_LCNF_FloatLetIn_initialNewArms___spec__11(lean_object*, lean_object*, uint64_t, uint64_t, size_t, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_FloatLetIn___hyg_2271____closed__3; +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_forCodeM___at_Lean_Compiler_LCNF_FloatLetIn_initialDecisions_goAlt___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_floatLetIn___closed__2; static lean_object* l_Lean_Compiler_LCNF_floatLetIn___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FloatLetIn_dontFloat_goFVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -120,6 +121,7 @@ static lean_object* l___private_Lean_Compiler_LCNF_FloatLetIn_0__Lean_Compiler_L lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_FloatLetIn_floatLetIn_go___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_FloatLetIn___hyg_2271____closed__4; +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_forCodeM___at_Lean_Compiler_LCNF_FloatLetIn_dontFloat___spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_DeclValue_mapCodeM___at_Lean_Compiler_LCNF_FloatLetIn_floatLetIn___spec__1(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_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_FloatLetIn_floatLetIn_go___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -130,7 +132,6 @@ LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Com LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_FloatLetIn_0__Lean_Compiler_LCNF_FloatLetIn_hashDecision____x40_Lean_Compiler_LCNF_FloatLetIn___hyg_16____boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_FloatLetIn_initialDecisions_goAlt___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*); uint8_t l_Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Compiler_LCNF_AltCore_getCode(lean_object*); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_FloatLetIn___hyg_2271____closed__6; static lean_object* l_Lean_Compiler_LCNF_FloatLetIn_instInhabitedDecision___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FloatLetIn_floatLetIn(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -173,8 +174,8 @@ lean_object* l_Lean_Compiler_LCNF_attachCodeDecls(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_FloatLetIn_dontFloat___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Param_forFVarM___at_Lean_Compiler_LCNF_FloatLetIn_dontFloat___spec__10(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_Compiler_LCNF_Alt_getCode(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_FloatLetIn_dontFloat___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_forCodeM___at_Lean_Compiler_LCNF_FloatLetIn_dontFloat___spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_FloatLetIn_floatLetIn_go___spec__3___closed__2; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Compiler_LCNF_FloatLetIn_initialNewArms___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FloatLetIn_float___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -251,7 +252,6 @@ lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_FloatLetIn_0__Lean_Compiler_LCNF_FloatLetIn_reprDecision____x40_Lean_Compiler_LCNF_FloatLetIn___hyg_161____closed__3; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_forCodeM___at_Lean_Compiler_LCNF_FloatLetIn_initialDecisions_goAlt___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FloatLetIn_withNewScope___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static uint64_t l_Lean_Compiler_LCNF_FloatLetIn_initialNewArms___closed__3; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_FloatLetIn_initialDecisions_goCases___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*); @@ -3448,7 +3448,7 @@ return x_25; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_forCodeM___at_Lean_Compiler_LCNF_FloatLetIn_initialDecisions_goAlt___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_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_forCodeM___at_Lean_Compiler_LCNF_FloatLetIn_initialDecisions_goAlt___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) { _start: { if (lean_obj_tag(x_1) == 0) @@ -3490,7 +3490,7 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_16 = l_Lean_Compiler_LCNF_AltCore_forCodeM___at_Lean_Compiler_LCNF_FloatLetIn_initialDecisions_goAlt___spec__13(x_14, x_15, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_16 = l_Lean_Compiler_LCNF_Alt_forCodeM___at_Lean_Compiler_LCNF_FloatLetIn_initialDecisions_goAlt___spec__13(x_14, x_15, x_6, x_7, x_8, x_9, x_10, x_11, x_12); if (lean_obj_tag(x_16) == 0) { lean_object* x_17; lean_object* x_18; size_t x_19; size_t x_20; @@ -9461,7 +9461,7 @@ return x_25; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_forCodeM___at_Lean_Compiler_LCNF_FloatLetIn_dontFloat___spec__15(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_forCodeM___at_Lean_Compiler_LCNF_FloatLetIn_dontFloat___spec__15(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { if (lean_obj_tag(x_1) == 0) @@ -9503,7 +9503,7 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_16 = l_Lean_Compiler_LCNF_AltCore_forCodeM___at_Lean_Compiler_LCNF_FloatLetIn_dontFloat___spec__15(x_14, x_15, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_16 = l_Lean_Compiler_LCNF_Alt_forCodeM___at_Lean_Compiler_LCNF_FloatLetIn_dontFloat___spec__15(x_14, x_15, x_6, x_7, x_8, x_9, x_10, x_11, x_12); if (lean_obj_tag(x_16) == 0) { lean_object* x_17; lean_object* x_18; size_t x_19; size_t x_20; @@ -15040,7 +15040,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_FloatLet { lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; x_10 = lean_array_mk(x_1); -x_11 = l_Lean_Compiler_LCNF_AltCore_getCode(x_2); +x_11 = l_Lean_Compiler_LCNF_Alt_getCode(x_2); x_12 = l_Lean_Compiler_LCNF_attachCodeDecls(x_10, x_11); lean_dec(x_10); x_13 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_FloatLetIn_floatLetIn_go), 7, 1); diff --git a/stage0/stdlib/Lean/Compiler/LCNF/InferType.c b/stage0/stdlib/Lean/Compiler/LCNF/InferType.c index 12c1c0b36ee2..0632e5f84712 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/InferType.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/InferType.c @@ -48,7 +48,6 @@ lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*) LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_InferType_getBinderName___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_mkForallParams___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_InferType_inferProjType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_Compiler_LCNF_instInhabitedCode; static lean_object* l_Lean_Compiler_LCNF_InferType_mkForallParams___closed__3; extern uint8_t l_instInhabitedBool; LEAN_EXPORT lean_object* l_List_isEqv___at_Lean_Compiler_LCNF_eqvTypes___spec__1___boxed(lean_object*, lean_object*); @@ -62,7 +61,6 @@ lean_object* l_Lean_Expr_proj___override(lean_object*, lean_object*, lean_object LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_InferType_inferForallType_go___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_InferType_inferForallType_go___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_EXPORT lean_object* l_Lean_Compiler_LCNF_InferType_inferLitValueType___boxed(lean_object*); -lean_object* l_Lean_Compiler_LCNF_instInhabitedAltCore___rarg(lean_object*); static lean_object* l_Lean_Compiler_LCNF_InferType_mkForallParams___closed__1; LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at_Lean_Compiler_LCNF_mkCasesResultType___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_Expr_fvarId_x21(lean_object*); @@ -80,6 +78,7 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_LCNF_InferType_inf LEAN_EXPORT lean_object* l_Nat_foldRevM_loop___at_Lean_Compiler_LCNF_InferType_mkForallFVars___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_Compiler_LCNF_mkAuxJpDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Decl_instantiateTypeLevelParams(lean_object*, lean_object*); +extern lean_object* l_Lean_Compiler_LCNF_instInhabitedAlt; static lean_object* l_panic___at_Lean_Compiler_LCNF_InferType_inferType___spec__1___closed__3; lean_object* l_Lean_stringToMessageData(lean_object*); static lean_object* l_Lean_Compiler_LCNF_InferType_inferLitValueType___closed__3; @@ -105,6 +104,7 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_In LEAN_EXPORT lean_object* l_Lean_mkFreshId___at_Lean_Compiler_LCNF_InferType_withLocalDecl___spec__2___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_mkAuxJpDecl_x27___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_getLevel___closed__1; +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_inferType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_expr_instantiate_rev_range(lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_InferType_inferForallType_go___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*); @@ -134,14 +134,12 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_LCNF_InferType_inf lean_object* l_Lean_Expr_forallE___override(lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Nat_foldRevM_loop___at_Lean_Compiler_LCNF_InferType_mkForallFVars___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_st_ref_get(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_inferType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_InferType_inferAppType___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_hasLooseBVars(lean_object*); static lean_object* l_Lean_Compiler_LCNF_InferType_inferConstType___closed__1; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Compiler_LCNF_InferType_inferAppType___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_mkAuxJpDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Compiler_LCNF_AltCore_getCode(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_InferType_getLevel_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_InferType_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_levelZero; @@ -158,7 +156,6 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_InferType_withLocalDecl___rarg___b lean_object* l_Lean_Compiler_LCNF_getDecl_x3f(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_Compiler_LCNF_mkFreshBinderName(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_inferType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_mkAuxJpDecl_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Compiler_LCNF_isErasedCompatible_go___spec__1___closed__1; uint8_t l_Lean_Expr_isErased(lean_object*); @@ -171,6 +168,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_InferType_inferForallType_go___lam LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_InferType_inferForallType_go___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* l_Lean_LocalDecl_userName(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_InferType_inferProjType___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_Compiler_LCNF_Alt_getCode(lean_object*); static lean_object* l_Lean_Compiler_LCNF_InferType_mkForallParams___closed__2; uint8_t l_Lean_Expr_isAny(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Compiler_LCNF_InferType_inferAppTypeCore___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*); @@ -186,6 +184,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_InferType_inferArgType(lean_object LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_InferType_inferForallType_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 uint8_t l_List_isEqv___at_Lean_Compiler_LCNF_eqvTypes___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_eqvTypes___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_inferType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Compiler_LCNF_InferType_inferAppTypeCore___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_Compiler_LCNF_InferType_mkForallParams___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); @@ -260,7 +259,6 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_inferType___boxed(lean_object LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_eqvTypes(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_InferType_inferAppTypeCore___lambda__1___boxed(lean_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_range(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_mkCasesResultType___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_LCNF_mkCasesResultType___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instMonad___rarg(lean_object*); static lean_object* l_Lean_Compiler_LCNF_InferType_mkForallParams___closed__5; @@ -11834,20 +11832,20 @@ lean_dec(x_3); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_inferType(lean_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_Compiler_LCNF_Alt_inferType(lean_object* x_1, 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_Compiler_LCNF_AltCore_getCode(x_1); +x_7 = l_Lean_Compiler_LCNF_Alt_getCode(x_1); x_8 = l_Lean_Compiler_LCNF_Code_inferType(x_7, x_2, x_3, x_4, x_5, x_6); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_inferType___boxed(lean_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_Compiler_LCNF_Alt_inferType___boxed(lean_object* x_1, 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_Compiler_LCNF_AltCore_inferType(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Lean_Compiler_LCNF_Alt_inferType(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); @@ -12101,7 +12099,7 @@ else lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_1, 0); x_13 = lean_array_uget(x_12, x_3); -x_14 = l_Lean_Compiler_LCNF_AltCore_inferType(x_13, x_5, x_6, x_7, x_8, x_9); +x_14 = l_Lean_Compiler_LCNF_Alt_inferType(x_13, x_5, x_6, x_7, x_8, x_9); lean_dec(x_13); if (lean_obj_tag(x_14) == 0) { @@ -12280,23 +12278,14 @@ return x_45; } } } -static lean_object* _init_l_Lean_Compiler_LCNF_mkCasesResultType___lambda__1___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Compiler_LCNF_instInhabitedCode; -x_2 = l_Lean_Compiler_LCNF_instInhabitedAltCore___rarg(x_1); -return x_2; -} -} LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_mkCasesResultType___lambda__1(lean_object* x_1, lean_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_Lean_Compiler_LCNF_mkCasesResultType___lambda__1___closed__1; +x_8 = l_Lean_Compiler_LCNF_instInhabitedAlt; x_9 = lean_unsigned_to_nat(0u); x_10 = lean_array_get(x_8, x_1, x_9); -x_11 = l_Lean_Compiler_LCNF_AltCore_inferType(x_10, x_3, x_4, x_5, x_6, x_7); +x_11 = l_Lean_Compiler_LCNF_Alt_inferType(x_10, x_3, x_4, x_5, x_6, x_7); lean_dec(x_10); if (lean_obj_tag(x_11) == 0) { @@ -13306,8 +13295,6 @@ l_Lean_Compiler_LCNF_getLevel___closed__1 = _init_l_Lean_Compiler_LCNF_getLevel_ lean_mark_persistent(l_Lean_Compiler_LCNF_getLevel___closed__1); l_Lean_Compiler_LCNF_getLevel___closed__2 = _init_l_Lean_Compiler_LCNF_getLevel___closed__2(); lean_mark_persistent(l_Lean_Compiler_LCNF_getLevel___closed__2); -l_Lean_Compiler_LCNF_mkCasesResultType___lambda__1___closed__1 = _init_l_Lean_Compiler_LCNF_mkCasesResultType___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Compiler_LCNF_mkCasesResultType___lambda__1___closed__1); l_Lean_Compiler_LCNF_mkCasesResultType___closed__1 = _init_l_Lean_Compiler_LCNF_mkCasesResultType___closed__1(); lean_mark_persistent(l_Lean_Compiler_LCNF_mkCasesResultType___closed__1); l_Lean_Compiler_LCNF_mkCasesResultType___closed__2 = _init_l_Lean_Compiler_LCNF_mkCasesResultType___closed__2(); diff --git a/stage0/stdlib/Lean/Compiler/LCNF/JoinPoints.c b/stage0/stdlib/Lean/Compiler/LCNF/JoinPoints.c index 21b23602fc9d..3dc7f7b151f6 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/JoinPoints.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/JoinPoints.c @@ -37,6 +37,7 @@ static lean_object* l___auto____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5005____ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Arg_mapFVarM___at_Lean_Compiler_LCNF_JoinPointContextExtender_extend_go___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_JoinPointContextExtender_mergeJpContextIfNecessary___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_JoinPointFinder_instInhabitedCandidateInfo___closed__1; +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_mapCodeM___at_Lean_Compiler_LCNF_JoinPointFinder_replace_go___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goReduce___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goAnalyzeFunDecl___spec__2(lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5005____closed__22; @@ -116,6 +117,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_JoinPointContextExtender_extend_go static lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goReduce___spec__6___closed__1; static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_4936____closed__17; static lean_object* l_Lean_Compiler_LCNF_Decl_findJoinPoints___closed__3; +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_mapCodeM___at_Lean_Compiler_LCNF_JoinPointContextExtender_extend_go___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_JoinPointContextExtender_withNewFunScope___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goAnalyze___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_JoinPointFinder_find(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -178,6 +180,7 @@ static lean_object* l___auto____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5005____ LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_JoinPoints_0__Lean_Compiler_LCNF_JoinPointFinder_findCandidate_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_extendJoinPointContext___closed__3; lean_object* lean_nat_div(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_mapCodeM___at_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goReduce___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goAnalyzeFunDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_4936____closed__5; static lean_object* l___auto____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5005____closed__8; @@ -187,6 +190,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_DeclValue_forCodeM___at_Lean_Compi LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goAnalyze___spec__12___boxed__const__1; LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_JoinPoints_0__Lean_Compiler_LCNF_JoinPointFinder_removeCandidatesInArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); +lean_object* l_Lean_Compiler_LCNF_Alt_getParams(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_LCNF_JoinPoints_0__Lean_Compiler_LCNF_JoinPointFinder_eraseCandidate___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_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goAnalyzeFunDecl___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_JoinPointCommonArgs_isInJpScope___closed__1; @@ -211,7 +215,6 @@ lean_object* l_Lean_Expr_forallE___override(lean_object*, lean_object*, lean_obj LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goAnalyze___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_LCNF_JoinPointFinder_find_go___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_JoinPointContextExtender_extend_go___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_Lean_Compiler_LCNF_AltCore_forCodeM___at_Lean_Compiler_LCNF_JoinPointFinder_find_go___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_ScopeM_isInScope(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_ReaderT_pure___at_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goAnalyzeFunDecl___spec__1(lean_object*); @@ -231,7 +234,6 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_JoinPoin lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_JoinPointContextExtender_withNewJpScope___spec__1___boxed(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Compiler_LCNF_AltCore_getCode(lean_object*); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_4936____closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Decl_findJoinPoints___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goAnalyze___spec__12___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -274,7 +276,6 @@ static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPo lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_JoinPointFinder_find_go___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Compiler_LCNF_JoinPointFinder_replace___spec__3(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Compiler_LCNF_FunDeclCore_getArity___rarg(lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Compiler_LCNF_JoinPoints_0__Lean_Compiler_LCNF_JoinPointFinder_addDependency___spec__4(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_JoinPointFinder_find___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goAnalyzeFunDecl___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -296,7 +297,6 @@ static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPo lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Compiler_LCNF_addFVarSubst___spec__5(lean_object*, lean_object*, lean_object*); lean_object* l_OptionT_instMonad___rarg(lean_object*); lean_object* l_Array_append___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_mapCodeM___at_Lean_Compiler_LCNF_JoinPointFinder_replace_go___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ScopeM_withNewScope___at_Lean_Compiler_LCNF_JoinPointContextExtender_withNewFunScope___spec__1(lean_object*); extern lean_object* l_Lean_Compiler_LCNF_instMonadCompilerM; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goAnalyze___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*); @@ -306,9 +306,10 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_J LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_JoinPointContextExtender_withNewFunScope(lean_object*); double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ScopeM_withNewScope___at_Lean_Compiler_LCNF_JoinPointContextExtender_withNewJpScope___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Compiler_LCNF_Alt_getCode(lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goReduce___spec__4(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_mapCodeM___at_Lean_Compiler_LCNF_JoinPointContextExtender_extend_go___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_erase___at___private_Lean_Compiler_LCNF_JoinPoints_0__Lean_Compiler_LCNF_JoinPointFinder_eraseCandidate___spec__4(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_forCodeM___at_Lean_Compiler_LCNF_JoinPointFinder_find_go___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Compiler_LCNF_JoinPoints_0__Lean_Compiler_LCNF_JoinPointFinder_removeCandidatesInArg___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_JoinPointContextExtender_extend_go___spec__9___closed__1; static lean_object* l___auto____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5005____closed__21; @@ -348,7 +349,6 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goAnaly LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_JoinPoints_0__Lean_Compiler_LCNF_JoinPointFinder_modifyCandidates___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_JoinPointContextExtender_extend_go___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_Compiler_LCNF_JoinPointContextExtender_replaceFVar___spec__1(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_mapCodeM___at_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goReduce___spec__5(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_Compiler_LCNF_JoinPoints_0__Lean_Compiler_LCNF_JoinPointFinder_eraseCandidate___spec__3___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_JoinPoints_0__Lean_Compiler_LCNF_JoinPointFinder_modifyCandidates(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_erase___at___private_Lean_Compiler_LCNF_JoinPoints_0__Lean_Compiler_LCNF_JoinPointFinder_eraseCandidate___spec__4___boxed(lean_object*, lean_object*); @@ -379,7 +379,6 @@ lean_object* l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(lean_obje static lean_object* l___auto____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5005____closed__10; static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5019____closed__1; LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_LCNF_Decl_findJoinPoints___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Compiler_LCNF_AltCore_getParams(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_LCNF_JoinPoints_0__Lean_Compiler_LCNF_JoinPointFinder_removeCandidatesInLetValue___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_JoinPointContextExtender_withNewJpScope(lean_object*); static lean_object* l___auto____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5005____closed__23; @@ -438,6 +437,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_JoinP lean_object* l_Lean_Compiler_LCNF_mkAuxParam(lean_object*, 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_Compiler_LCNF_JoinPointContextExtender_extend_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Compiler_LCNF_FunDecl_getArity(lean_object*); lean_object* l_Lean_Expr_lam___override(lean_object*, lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Compiler_LCNF_extendJoinPointContext___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_JoinPointFinder_replace___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -3337,7 +3337,7 @@ x_13 = l_Lean_Compiler_LCNF_ScopeM_withBackTrackingScope___at_Lean_Compiler_LCNF return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_forCodeM___at_Lean_Compiler_LCNF_JoinPointFinder_find_go___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_forCodeM___at_Lean_Compiler_LCNF_JoinPointFinder_find_go___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { if (lean_obj_tag(x_1) == 0) @@ -3386,7 +3386,7 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_16 = l_Lean_Compiler_LCNF_AltCore_forCodeM___at_Lean_Compiler_LCNF_JoinPointFinder_find_go___spec__5(x_14, x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_16 = l_Lean_Compiler_LCNF_Alt_forCodeM___at_Lean_Compiler_LCNF_JoinPointFinder_find_go___spec__5(x_14, x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); if (lean_obj_tag(x_16) == 0) { lean_object* x_17; lean_object* x_18; size_t x_19; size_t x_20; @@ -4591,7 +4591,7 @@ lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; x_190 = lean_ctor_get(x_189, 1); lean_inc(x_190); lean_dec(x_189); -x_191 = l_Lean_Compiler_LCNF_FunDeclCore_getArity___rarg(x_183); +x_191 = l_Lean_Compiler_LCNF_FunDecl_getArity(x_183); lean_dec(x_183); lean_inc(x_185); x_192 = l___private_Lean_Compiler_LCNF_JoinPoints_0__Lean_Compiler_LCNF_JoinPointFinder_addCandidate(x_185, x_191, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_190); @@ -5247,11 +5247,11 @@ return x_9; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_mapCodeM___at_Lean_Compiler_LCNF_JoinPointFinder_replace_go___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_mapCodeM___at_Lean_Compiler_LCNF_JoinPointFinder_replace_go___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; lean_object* x_10; -x_9 = l_Lean_Compiler_LCNF_AltCore_getCode(x_1); +x_9 = l_Lean_Compiler_LCNF_Alt_getCode(x_1); x_10 = lean_apply_7(x_2, x_9, x_3, x_4, x_5, x_6, x_7, x_8); if (lean_obj_tag(x_10) == 0) { @@ -5343,7 +5343,7 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_16 = l_Lean_Compiler_LCNF_AltCore_mapCodeM___at_Lean_Compiler_LCNF_JoinPointFinder_replace_go___spec__3(x_12, x_15, x_4, x_5, x_6, x_7, x_8, x_9); +x_16 = l_Lean_Compiler_LCNF_Alt_mapCodeM___at_Lean_Compiler_LCNF_JoinPointFinder_replace_go___spec__3(x_12, x_15, x_4, x_5, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_16) == 0) { lean_object* x_17; lean_object* x_18; size_t x_19; size_t x_20; lean_object* x_21; @@ -12243,7 +12243,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_JoinPointContextExtender_withNewAl _start: { lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_11 = l_Lean_Compiler_LCNF_AltCore_getParams(x_1); +x_11 = l_Lean_Compiler_LCNF_Alt_getParams(x_1); x_12 = lean_array_size(x_11); x_13 = 0; x_14 = l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_JoinPointContextExtender_withNewJpScope___spec__1(x_12, x_13, x_11); @@ -13977,11 +13977,11 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_mapCodeM___at_Lean_Compiler_LCNF_JoinPointContextExtender_extend_go___spec__13(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_mapCodeM___at_Lean_Compiler_LCNF_JoinPointContextExtender_extend_go___spec__13(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; lean_object* x_12; -x_11 = l_Lean_Compiler_LCNF_AltCore_getCode(x_1); +x_11 = l_Lean_Compiler_LCNF_Alt_getCode(x_1); x_12 = lean_apply_9(x_2, 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) { @@ -14071,7 +14071,7 @@ x_15 = lean_unsigned_to_nat(0u); x_16 = lean_array_uset(x_3, x_2, x_15); x_17 = l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_JoinPointContextExtender_extend_go___spec__14___closed__1; lean_inc(x_14); -x_18 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_AltCore_mapCodeM___at_Lean_Compiler_LCNF_JoinPointContextExtender_extend_go___spec__13), 10, 2); +x_18 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Alt_mapCodeM___at_Lean_Compiler_LCNF_JoinPointContextExtender_extend_go___spec__13), 10, 2); lean_closure_set(x_18, 0, x_14); lean_closure_set(x_18, 1, x_17); lean_inc(x_10); @@ -17997,7 +17997,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_JoinP _start: { lean_object* x_11; lean_object* x_12; -x_11 = l_Lean_Compiler_LCNF_AltCore_getCode(x_1); +x_11 = l_Lean_Compiler_LCNF_Alt_getCode(x_1); x_12 = l_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goAnalyze(x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_12; } @@ -18021,7 +18021,7 @@ if (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; lean_dec(x_4); x_14 = lean_array_uget(x_1, x_2); -x_15 = l_Lean_Compiler_LCNF_AltCore_getParams(x_14); +x_15 = l_Lean_Compiler_LCNF_Alt_getParams(x_14); x_16 = lean_array_get_size(x_15); x_17 = lean_unsigned_to_nat(0u); x_18 = lean_nat_dec_lt(x_17, x_16); @@ -19151,11 +19151,11 @@ x_3 = lean_panic_fn(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_mapCodeM___at_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goReduce___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Alt_mapCodeM___at_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goReduce___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; lean_object* x_10; -x_9 = l_Lean_Compiler_LCNF_AltCore_getCode(x_1); +x_9 = l_Lean_Compiler_LCNF_Alt_getCode(x_1); x_10 = lean_apply_7(x_2, x_9, x_3, x_4, x_5, x_6, x_7, x_8); if (lean_obj_tag(x_10) == 0) { @@ -19247,7 +19247,7 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_16 = l_Lean_Compiler_LCNF_AltCore_mapCodeM___at_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goReduce___spec__5(x_12, x_15, x_4, x_5, x_6, x_7, x_8, x_9); +x_16 = l_Lean_Compiler_LCNF_Alt_mapCodeM___at_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goReduce___spec__5(x_12, x_15, x_4, x_5, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_16) == 0) { lean_object* x_17; lean_object* x_18; size_t x_19; size_t x_20; lean_object* x_21; @@ -19496,7 +19496,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goRedu lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goReduce___closed__3; x_2 = l_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goReduce___closed__4; -x_3 = lean_unsigned_to_nat(319u); +x_3 = lean_unsigned_to_nat(320u); x_4 = lean_unsigned_to_nat(9u); x_5 = l_Lean_Compiler_LCNF_Expr_forFVarM___at___private_Lean_Compiler_LCNF_JoinPoints_0__Lean_Compiler_LCNF_JoinPointFinder_removeCandidatesInArg___spec__2___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); diff --git a/stage0/stdlib/Lean/Compiler/LCNF/MonoTypes.c b/stage0/stdlib/Lean/Compiler/LCNF/MonoTypes.c index 899625fe5f2f..5456e1b3e2c6 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/MonoTypes.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/MonoTypes.c @@ -3975,6 +3975,7 @@ if (x_73 == 0) lean_object* x_74; uint8_t x_75; x_74 = l_Lean_Compiler_LCNF_toMonoType_visitApp___closed__2; x_75 = lean_string_dec_eq(x_69, x_74); +x_75 = 0; // bootstrap: disable lean_dec(x_69); if (x_75 == 0) { diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Passes.c b/stage0/stdlib/Lean/Compiler/LCNF/Passes.c index 640bc1835d58..122332defea2 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Passes.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Passes.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Compiler.LCNF.Passes -// Imports: Lean.Compiler.LCNF.PassManager Lean.Compiler.LCNF.PullLetDecls Lean.Compiler.LCNF.CSE Lean.Compiler.LCNF.Simp Lean.Compiler.LCNF.PullFunDecls Lean.Compiler.LCNF.ReduceJpArity Lean.Compiler.LCNF.JoinPoints Lean.Compiler.LCNF.Specialize Lean.Compiler.LCNF.PhaseExt Lean.Compiler.LCNF.ToMono Lean.Compiler.LCNF.LambdaLifting Lean.Compiler.LCNF.FloatLetIn Lean.Compiler.LCNF.ReduceArity Lean.Compiler.LCNF.ElimDeadBranches +// Imports: Lean.Compiler.LCNF.PassManager Lean.Compiler.LCNF.PullLetDecls Lean.Compiler.LCNF.CSE Lean.Compiler.LCNF.Simp Lean.Compiler.LCNF.PullFunDecls Lean.Compiler.LCNF.ReduceJpArity Lean.Compiler.LCNF.JoinPoints Lean.Compiler.LCNF.Specialize Lean.Compiler.LCNF.PhaseExt Lean.Compiler.LCNF.ToMono Lean.Compiler.LCNF.LambdaLifting Lean.Compiler.LCNF.FloatLetIn Lean.Compiler.LCNF.ReduceArity Lean.Compiler.LCNF.ElimDeadBranches Lean.Compiler.LCNF.StructProjCases #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -16,167 +16,165 @@ extern "C" { lean_object* l_Lean_Compiler_LCNF_cse(uint8_t, uint8_t, lean_object*); static lean_object* l_Lean_Compiler_LCNF_init___closed__2; static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__26; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__22; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__14; +static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__49; +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__14; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_init___elambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__22; static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__22; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_955____closed__1; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__3; +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_saveBase___closed__4; static lean_object* l_Lean_Compiler_LCNF_saveBase___closed__1; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__12; lean_object* l_Lean_ConstantInfo_type(lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__6; lean_object* l_Lean_Compiler_LCNF_PassInstaller_runFromDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__33; static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__27; -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_init___elambda__1___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_Compiler_LCNF_builtinPassManager___closed__32; static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__34; static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__4; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_saveBase___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__10; -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____lambda__4___boxed(lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_955____closed__2; static lean_object* l_Lean_Compiler_LCNF_init___closed__4; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__3___closed__2; uint8_t lean_usize_dec_eq(size_t, size_t); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__7; lean_object* l_Lean_Compiler_LCNF_Decl_saveBase(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Compiler_LCNF_lambdaLifting; -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448_(lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__5; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__21; lean_object* l_Lean_Attribute_Builtin_ensureNoArgs(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__30; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_getPassManager___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__5; static lean_object* l_Lean_Compiler_LCNF_saveMono___closed__1; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__14; static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__2; lean_object* l_Lean_stringToMessageData(lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__1; static lean_object* l_Lean_Compiler_LCNF_saveMono___closed__4; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__7; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__4; uint8_t lean_string_dec_eq(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__19; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_957____closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_getPassManager(lean_object*); static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__6; static lean_object* l_Lean_Compiler_LCNF_saveBase___closed__2; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_runImportedDecls___spec__2(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807_(lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__8; static lean_object* l_Lean_Compiler_LCNF_addPass___closed__6; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__2; static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__18; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__4; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_saveMono___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Decl_saveMono(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__10; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__7; extern lean_object* l_Lean_Compiler_LCNF_commonJoinPointArgs; size_t lean_usize_of_nat(lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_957____closed__3; static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__15; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_saveBase; +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__48; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__16; lean_object* lean_st_ref_take(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____lambda__4(lean_object*); extern lean_object* l_Lean_Compiler_LCNF_reduceArity; lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__24; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_init___elambda__1___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_addPass___closed__5; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__17; lean_object* l_Lean_PersistentEnvExtension_addEntry___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__8; lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); lean_object* l_Lean_PersistentEnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*, uint8_t); lean_object* l_Lean_ImportM_runCoreM___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__12; extern lean_object* l_Lean_Compiler_LCNF_eagerLambdaLifting; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_runImportedDecls___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_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__31; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__17; static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__9; static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__47; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__13; static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__38; lean_object* lean_st_ref_get(lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_registerTagAttribute___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_reduceJpArity(uint8_t); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__21; -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__9; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_trace___elambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__44; static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__37; static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__21; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__12; static lean_object* l_Lean_Compiler_LCNF_addPass___closed__7; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_simp(lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__19; lean_object* l_EStateM_pure___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_addPass___closed__8; +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__1; static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__39; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__2___closed__1; static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__13; -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Compiler_LCNF_findJoinPoints; +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____lambda__3(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_saveMono___closed__2; static lean_object* l_Lean_Compiler_LCNF_getPassManager___rarg___closed__2; -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Compiler_LCNF_pullInstances; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_trace(uint8_t); extern lean_object* l_Lean_Compiler_LCNF_elimDeadBranches; static lean_object* l_Lean_Compiler_LCNF_addPass___closed__1; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__7; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__6; +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____lambda__4(lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__1; static lean_object* l_Lean_Compiler_LCNF_saveBase___closed__3; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__2; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_runImportedDecls___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__3; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__4; static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__7; -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____lambda__2(lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__5; lean_object* l_Lean_Compiler_LCNF_floatLetIn(uint8_t, lean_object*); static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__3; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_955____closed__3; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_passManagerExt; static lean_object* l_Lean_Compiler_LCNF_getPassManager___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_getPassManager___rarg___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____lambda__3(lean_object*); extern lean_object* l_Lean_Compiler_LCNF_toMono; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__5; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__20; lean_object* l_Lean_MessageData_ofExpr(lean_object*); static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__23; -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__3; static lean_object* l_Lean_Compiler_LCNF_addPass___closed__2; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__18; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__3___closed__1; static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__40; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__1; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__16; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_saveMono___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Compiler_LCNF_pullFunDecls; static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__35; lean_object* l_Lean_Compiler_LCNF_normalizeFVarIds(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__18; static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__8; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_runImportedDecls(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_saveMono___closed__3; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__8; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_init; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__36; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__3___closed__1; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__9; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_957____closed__4; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_957____closed__1; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__9; static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__5; static lean_object* l_Lean_Compiler_LCNF_trace___closed__3; static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__17; static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__20; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__19; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__13; static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__43; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__20; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__2___closed__1; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__3___closed__2; static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__11; lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); lean_object* l_Lean_registerBuiltinAttribute(lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__11; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__9; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__2___closed__2; lean_object* l_Lean_getConstInfo___at___private_Lean_Compiler_LCNF_Util_0__Lean_Compiler_LCNF_getCasesOnInductiveVal_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_reverse___rarg(lean_object*); uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_162_(uint8_t, uint8_t); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_955_(lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__10; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__6; lean_object* lean_array_mk(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_getPassManager___boxed(lean_object*); size_t lean_usize_add(size_t, size_t); @@ -186,40 +184,44 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_ru lean_object* l_Lean_Compiler_LCNF_Pass_mkPerDeclaration(lean_object*, lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_builtinPassManager; lean_object* lean_array_uget(lean_object*, size_t); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__10; size_t lean_array_size(lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__10; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__2; static lean_object* l_Lean_Compiler_LCNF_init___closed__1; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_addPass(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__16; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__2; +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809_(lean_object*); lean_object* l_Lean_Compiler_LCNF_extendJoinPointContext(lean_object*, uint8_t, lean_object*); static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__1; static lean_object* l_Lean_Compiler_LCNF_trace___closed__2; static lean_object* l_Lean_Compiler_LCNF_addPass___closed__4; static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__41; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_trace___elambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450_(lean_object*); lean_object* lean_array_get_size(lean_object*); extern lean_object* l_Lean_Compiler_LCNF_instInhabitedPassManager; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__8; static lean_object* l_Lean_Compiler_LCNF_addPass___closed__3; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__4; static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__28; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__11; uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_saveMono; uint8_t lean_usize_dec_lt(size_t, size_t); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__15; static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__29; +extern lean_object* l_Lean_Compiler_LCNF_structProjCases; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_init___elambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__6; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__3; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_955____closed__4; +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_957_(lean_object*); lean_object* l_Lean_throwError___at___private_Lean_ReducibilityAttrs_0__Lean_validate___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_runImportedDecls___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____lambda__2(lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_Attribute_Builtin_ensureNoArgs___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__25; static lean_object* l_Lean_Compiler_LCNF_trace___closed__1; -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__15; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__2___closed__2; +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____lambda__4___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_trace___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_saveBase___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Compiler_LCNF_specialize; @@ -941,7 +943,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinPassManager___closed__21() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__17; +x_1 = l_Lean_Compiler_LCNF_structProjCases; x_2 = l_Lean_Compiler_LCNF_builtinPassManager___closed__20; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -953,7 +955,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinPassManager___closed__22() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__16; +x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__17; x_2 = l_Lean_Compiler_LCNF_builtinPassManager___closed__21; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -965,7 +967,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinPassManager___closed__23() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_lambdaLifting; +x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__16; x_2 = l_Lean_Compiler_LCNF_builtinPassManager___closed__22; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -977,7 +979,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinPassManager___closed__24() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_elimDeadBranches; +x_1 = l_Lean_Compiler_LCNF_lambdaLifting; x_2 = l_Lean_Compiler_LCNF_builtinPassManager___closed__23; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -989,7 +991,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinPassManager___closed__25() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__15; +x_1 = l_Lean_Compiler_LCNF_elimDeadBranches; x_2 = l_Lean_Compiler_LCNF_builtinPassManager___closed__24; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1001,7 +1003,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinPassManager___closed__26() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__14; +x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__15; x_2 = l_Lean_Compiler_LCNF_builtinPassManager___closed__25; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1013,7 +1015,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinPassManager___closed__27() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_commonJoinPointArgs; +x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__14; x_2 = l_Lean_Compiler_LCNF_builtinPassManager___closed__26; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1025,7 +1027,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinPassManager___closed__28() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_reduceArity; +x_1 = l_Lean_Compiler_LCNF_commonJoinPointArgs; x_2 = l_Lean_Compiler_LCNF_builtinPassManager___closed__27; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1037,7 +1039,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinPassManager___closed__29() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__13; +x_1 = l_Lean_Compiler_LCNF_reduceArity; x_2 = l_Lean_Compiler_LCNF_builtinPassManager___closed__28; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1049,7 +1051,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinPassManager___closed__30() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__12; +x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__13; x_2 = l_Lean_Compiler_LCNF_builtinPassManager___closed__29; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1061,7 +1063,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinPassManager___closed__31() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__11; +x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__12; x_2 = l_Lean_Compiler_LCNF_builtinPassManager___closed__30; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1073,7 +1075,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinPassManager___closed__32() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__10; +x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__11; x_2 = l_Lean_Compiler_LCNF_builtinPassManager___closed__31; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1085,7 +1087,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinPassManager___closed__33() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_toMono; +x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__10; x_2 = l_Lean_Compiler_LCNF_builtinPassManager___closed__32; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1097,7 +1099,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinPassManager___closed__34() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_saveBase; +x_1 = l_Lean_Compiler_LCNF_toMono; x_2 = l_Lean_Compiler_LCNF_builtinPassManager___closed__33; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1109,7 +1111,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinPassManager___closed__35() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__9; +x_1 = l_Lean_Compiler_LCNF_saveBase; x_2 = l_Lean_Compiler_LCNF_builtinPassManager___closed__34; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1121,7 +1123,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinPassManager___closed__36() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__8; +x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__9; x_2 = l_Lean_Compiler_LCNF_builtinPassManager___closed__35; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1133,7 +1135,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinPassManager___closed__37() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_specialize; +x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__8; x_2 = l_Lean_Compiler_LCNF_builtinPassManager___closed__36; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1145,7 +1147,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinPassManager___closed__38() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_eagerLambdaLifting; +x_1 = l_Lean_Compiler_LCNF_specialize; x_2 = l_Lean_Compiler_LCNF_builtinPassManager___closed__37; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1157,7 +1159,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinPassManager___closed__39() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__7; +x_1 = l_Lean_Compiler_LCNF_eagerLambdaLifting; x_2 = l_Lean_Compiler_LCNF_builtinPassManager___closed__38; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1169,7 +1171,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinPassManager___closed__40() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__5; +x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__7; x_2 = l_Lean_Compiler_LCNF_builtinPassManager___closed__39; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1181,7 +1183,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinPassManager___closed__41() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_pullFunDecls; +x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__5; x_2 = l_Lean_Compiler_LCNF_builtinPassManager___closed__40; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1193,7 +1195,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinPassManager___closed__42() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_findJoinPoints; +x_1 = l_Lean_Compiler_LCNF_pullFunDecls; x_2 = l_Lean_Compiler_LCNF_builtinPassManager___closed__41; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1205,7 +1207,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinPassManager___closed__43() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__4; +x_1 = l_Lean_Compiler_LCNF_findJoinPoints; x_2 = l_Lean_Compiler_LCNF_builtinPassManager___closed__42; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1217,7 +1219,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinPassManager___closed__44() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__3; +x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__4; x_2 = l_Lean_Compiler_LCNF_builtinPassManager___closed__43; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1229,7 +1231,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinPassManager___closed__45() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__1; +x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__3; x_2 = l_Lean_Compiler_LCNF_builtinPassManager___closed__44; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1241,7 +1243,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinPassManager___closed__46() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_pullInstances; +x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__1; x_2 = l_Lean_Compiler_LCNF_builtinPassManager___closed__45; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1253,7 +1255,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinPassManager___closed__47() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_init; +x_1 = l_Lean_Compiler_LCNF_pullInstances; x_2 = l_Lean_Compiler_LCNF_builtinPassManager___closed__46; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1264,8 +1266,20 @@ return x_3; static lean_object* _init_l_Lean_Compiler_LCNF_builtinPassManager___closed__48() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Compiler_LCNF_init; +x_2 = l_Lean_Compiler_LCNF_builtinPassManager___closed__47; +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_Compiler_LCNF_builtinPassManager___closed__49() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__47; +x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__48; x_2 = lean_array_mk(x_1); return x_2; } @@ -1274,7 +1288,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinPassManager() { _start: { lean_object* x_1; -x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__48; +x_1 = l_Lean_Compiler_LCNF_builtinPassManager___closed__49; return x_1; } } @@ -1504,7 +1518,7 @@ lean_dec(x_1); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____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; @@ -1567,7 +1581,7 @@ return x_17; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____lambda__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____lambda__2(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; @@ -1638,7 +1652,7 @@ return x_16; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____lambda__3(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____lambda__3(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; @@ -1650,7 +1664,7 @@ x_4 = lean_array_mk(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____lambda__4(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____lambda__4(lean_object* x_1) { _start: { lean_object* x_2; @@ -1658,7 +1672,7 @@ x_2 = lean_box(0); return x_2; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__1() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__1() { _start: { lean_object* x_1; @@ -1666,7 +1680,7 @@ x_1 = lean_mk_string_unchecked("Lean", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__2() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__2() { _start: { lean_object* x_1; @@ -1674,7 +1688,7 @@ x_1 = lean_mk_string_unchecked("Compiler", 8, 8); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__3() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__3() { _start: { lean_object* x_1; @@ -1682,7 +1696,7 @@ x_1 = lean_mk_string_unchecked("LCNF", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__4() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__4() { _start: { lean_object* x_1; @@ -1690,19 +1704,19 @@ x_1 = lean_mk_string_unchecked("passManagerExt", 14, 14); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__5() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____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_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__1; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__2; -x_3 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__3; -x_4 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__4; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__1; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__2; +x_3 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__3; +x_4 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____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_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__6() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -1714,53 +1728,53 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__7() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__6; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__6; x_2 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__8() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__8() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____lambda__2), 2, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____lambda__2), 2, 0); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__9() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__9() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____lambda__3), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____lambda__3), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__10() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__10() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____lambda__4___boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____lambda__4___boxed), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450_(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; uint8_t x_10; lean_object* x_11; lean_object* x_12; x_2 = lean_box(0); -x_3 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____lambda__1), 4, 1); +x_3 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____lambda__1), 4, 1); lean_closure_set(x_3, 0, x_2); x_4 = lean_box(0); -x_5 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__5; -x_6 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__7; -x_7 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__8; -x_8 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__9; -x_9 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__10; +x_5 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__5; +x_6 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__7; +x_7 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__8; +x_8 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__9; +x_9 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__10; x_10 = 2; x_11 = lean_alloc_ctor(0, 8, 1); lean_ctor_set(x_11, 0, x_5); @@ -1776,11 +1790,11 @@ x_12 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg(x_11, x_1); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____lambda__4___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____lambda__4___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____lambda__4(x_1); +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____lambda__4(x_1); lean_dec(x_1); return x_2; } @@ -2014,7 +2028,7 @@ lean_dec(x_11); x_17 = lean_ctor_get(x_12, 1); lean_inc(x_17); lean_dec(x_12); -x_18 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__1; +x_18 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__1; x_19 = lean_string_dec_eq(x_17, x_18); lean_dec(x_17); if (x_19 == 0) @@ -2041,7 +2055,7 @@ return x_25; else { lean_object* x_26; uint8_t x_27; -x_26 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__2; +x_26 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__2; x_27 = lean_string_dec_eq(x_16, x_26); lean_dec(x_16); if (x_27 == 0) @@ -2067,7 +2081,7 @@ return x_33; else { lean_object* x_34; uint8_t x_35; -x_34 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__3; +x_34 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__3; x_35 = lean_string_dec_eq(x_15, x_34); lean_dec(x_15); if (x_35 == 0) @@ -2441,7 +2455,7 @@ return x_140; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____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; @@ -2496,7 +2510,7 @@ return x_16; } } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__2___closed__1() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__2___closed__1() { _start: { lean_object* x_1; @@ -2504,16 +2518,16 @@ x_1 = lean_mk_string_unchecked("invalid attribute 'cpass', must be global", 41, return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__2___closed__2() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__2___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__2___closed__1; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__2___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__2(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__2(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; @@ -2531,7 +2545,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_dec(x_1); -x_11 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__2___closed__2; +x_11 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__2___closed__2; x_12 = l_Lean_throwError___at_Lean_registerTagAttribute___spec__1(x_11, x_4, x_5, x_8); lean_dec(x_5); lean_dec(x_4); @@ -2558,7 +2572,7 @@ else { lean_object* x_17; lean_object* x_18; x_17 = lean_box(0); -x_18 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__1(x_1, x_17, x_4, x_5, x_8); +x_18 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__1(x_1, x_17, x_4, x_5, x_8); return x_18; } } @@ -2589,7 +2603,7 @@ return x_22; } } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__3___closed__1() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__3___closed__1() { _start: { lean_object* x_1; @@ -2597,55 +2611,55 @@ x_1 = lean_mk_string_unchecked("attribute cannot be erased", 26, 26); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__3___closed__2() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__3___closed__1; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__3___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__3___closed__2; +x_5 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__3___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_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__1() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__1; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__2() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__1; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__2; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__1; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__3() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__2; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__3; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__2; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__4() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__4() { _start: { lean_object* x_1; @@ -2653,17 +2667,17 @@ x_1 = lean_mk_string_unchecked("initFn", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__5() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__3; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__4; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__3; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__6() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__6() { _start: { lean_object* x_1; @@ -2671,47 +2685,47 @@ x_1 = lean_mk_string_unchecked("_@", 2, 2); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__7() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__5; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__6; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__5; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__8() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__7; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__1; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__7; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__9() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__8; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__2; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__8; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__10() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__9; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__3; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__9; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__11() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__11() { _start: { lean_object* x_1; @@ -2719,17 +2733,17 @@ x_1 = lean_mk_string_unchecked("Passes", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__12() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__10; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__11; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__10; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__11; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__13() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__13() { _start: { lean_object* x_1; @@ -2737,27 +2751,27 @@ x_1 = lean_mk_string_unchecked("_hyg", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__14() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__12; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__13; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__12; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__13; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__15() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__14; -x_2 = lean_unsigned_to_nat(807u); +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__14; +x_2 = lean_unsigned_to_nat(809u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__16() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__16() { _start: { lean_object* x_1; @@ -2765,17 +2779,17 @@ x_1 = lean_mk_string_unchecked("cpass", 5, 5); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__17() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__16; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__16; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__18() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__18() { _start: { lean_object* x_1; @@ -2783,13 +2797,13 @@ x_1 = lean_mk_string_unchecked("compiler passes for the code generator", 38, 38) return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__19() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__15; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__17; -x_3 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__18; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__15; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__17; +x_3 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__18; x_4 = 1; x_5 = lean_alloc_ctor(0, 3, 1); lean_ctor_set(x_5, 0, x_1); @@ -2799,29 +2813,29 @@ lean_ctor_set_uint8(x_5, sizeof(void*)*3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__20() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__20() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__2___boxed), 6, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__2___boxed), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__21() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__21() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__3___boxed), 4, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__3___boxed), 4, 0); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__22() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__19; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__20; -x_3 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__21; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__19; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__20; +x_3 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__21; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2829,92 +2843,92 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__22; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__22; x_3 = l_Lean_registerBuiltinAttribute(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; -x_6 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__1(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__1(x_1, x_2, x_3, x_4, x_5); lean_dec(x_2); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____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_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { uint8_t x_7; lean_object* x_8; x_7 = lean_unbox(x_3); lean_dec(x_3); -x_8 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__2(x_1, x_2, x_7, x_4, x_5, x_6); +x_8 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__2(x_1, x_2, x_7, x_4, x_5, x_6); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__3(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__3(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); return x_5; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_955____closed__1() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_957____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__2; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__2; x_2 = l_Lean_Compiler_LCNF_saveBase___closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_955____closed__2() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_957____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__14; -x_2 = lean_unsigned_to_nat(955u); +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__14; +x_2 = lean_unsigned_to_nat(957u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_955____closed__3() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_957____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__2; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__2; x_2 = l_Lean_Compiler_LCNF_saveMono___closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_955____closed__4() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_957____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__2; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__2; x_2 = l_Lean_Compiler_LCNF_trace___closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_955_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_957_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_955____closed__1; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_957____closed__1; x_3 = 1; -x_4 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_955____closed__2; +x_4 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_957____closed__2; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); if (lean_obj_tag(x_5) == 0) { @@ -2922,7 +2936,7 @@ lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = lean_ctor_get(x_5, 1); lean_inc(x_6); lean_dec(x_5); -x_7 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_955____closed__3; +x_7 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_957____closed__3; x_8 = l_Lean_registerTraceClass(x_7, x_3, x_4, x_6); if (lean_obj_tag(x_8) == 0) { @@ -2930,7 +2944,7 @@ lean_object* x_9; lean_object* x_10; lean_object* x_11; x_9 = lean_ctor_get(x_8, 1); lean_inc(x_9); lean_dec(x_8); -x_10 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_955____closed__4; +x_10 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_957____closed__4; x_11 = l_Lean_registerTraceClass(x_10, x_3, x_4, x_9); return x_11; } @@ -2995,6 +3009,7 @@ lean_object* initialize_Lean_Compiler_LCNF_LambdaLifting(uint8_t builtin, lean_o lean_object* initialize_Lean_Compiler_LCNF_FloatLetIn(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Compiler_LCNF_ReduceArity(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Compiler_LCNF_ElimDeadBranches(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Compiler_LCNF_StructProjCases(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Compiler_LCNF_Passes(uint8_t builtin, lean_object* w) { lean_object * res; @@ -3042,6 +3057,9 @@ lean_dec_ref(res); res = initialize_Lean_Compiler_LCNF_ElimDeadBranches(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Compiler_LCNF_StructProjCases(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); l_Lean_Compiler_LCNF_init___closed__1 = _init_l_Lean_Compiler_LCNF_init___closed__1(); lean_mark_persistent(l_Lean_Compiler_LCNF_init___closed__1); l_Lean_Compiler_LCNF_init___closed__2 = _init_l_Lean_Compiler_LCNF_init___closed__2(); @@ -3174,29 +3192,31 @@ l_Lean_Compiler_LCNF_builtinPassManager___closed__47 = _init_l_Lean_Compiler_LCN lean_mark_persistent(l_Lean_Compiler_LCNF_builtinPassManager___closed__47); l_Lean_Compiler_LCNF_builtinPassManager___closed__48 = _init_l_Lean_Compiler_LCNF_builtinPassManager___closed__48(); lean_mark_persistent(l_Lean_Compiler_LCNF_builtinPassManager___closed__48); +l_Lean_Compiler_LCNF_builtinPassManager___closed__49 = _init_l_Lean_Compiler_LCNF_builtinPassManager___closed__49(); +lean_mark_persistent(l_Lean_Compiler_LCNF_builtinPassManager___closed__49); l_Lean_Compiler_LCNF_builtinPassManager = _init_l_Lean_Compiler_LCNF_builtinPassManager(); lean_mark_persistent(l_Lean_Compiler_LCNF_builtinPassManager); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__1 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__1(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__1); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__2 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__2(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__2); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__3 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__3(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__3); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__4 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__4(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__4); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__5 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__5(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__5); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__6 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__6(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__6); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__7 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__7(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__7); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__8 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__8(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__8); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__9 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__9(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__9); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__10 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__10(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448____closed__10); -if (builtin) {res = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_448_(lean_io_mk_world()); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__1 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__1(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__1); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__2 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__2(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__2); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__3 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__3(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__3); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__4 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__4(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__4); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__5 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__5(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__5); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__6 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__6(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__6); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__7 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__7(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__7); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__8 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__8(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__8); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__9 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__9(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__9); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__10 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__10(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450____closed__10); +if (builtin) {res = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_450_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Compiler_LCNF_passManagerExt = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Compiler_LCNF_passManagerExt); @@ -3221,70 +3241,70 @@ l_Lean_Compiler_LCNF_addPass___closed__7 = _init_l_Lean_Compiler_LCNF_addPass___ lean_mark_persistent(l_Lean_Compiler_LCNF_addPass___closed__7); l_Lean_Compiler_LCNF_addPass___closed__8 = _init_l_Lean_Compiler_LCNF_addPass___closed__8(); lean_mark_persistent(l_Lean_Compiler_LCNF_addPass___closed__8); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__2___closed__1 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__2___closed__1); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__2___closed__2 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__2___closed__2(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__2___closed__2); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__3___closed__1 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__3___closed__1); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__3___closed__2 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____lambda__3___closed__2); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__1 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__1(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__1); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__2 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__2(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__2); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__3 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__3(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__3); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__4 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__4(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__4); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__5 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__5(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__5); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__6 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__6(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__6); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__7 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__7(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__7); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__8 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__8(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__8); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__9 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__9(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__9); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__10 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__10(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__10); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__11 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__11(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__11); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__12 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__12(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__12); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__13 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__13(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__13); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__14 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__14(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__14); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__15 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__15(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__15); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__16 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__16(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__16); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__17 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__17(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__17); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__18 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__18(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__18); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__19 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__19(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__19); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__20 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__20(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__20); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__21 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__21(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__21); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__22 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__22(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807____closed__22); -if (builtin) {res = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_807_(lean_io_mk_world()); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__2___closed__1 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__2___closed__1); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__2___closed__2 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__2___closed__2); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__3___closed__1 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__3___closed__1); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__3___closed__2 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____lambda__3___closed__2); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__1 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__1(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__1); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__2 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__2(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__2); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__3 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__3(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__3); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__4 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__4(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__4); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__5 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__5(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__5); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__6 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__6(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__6); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__7 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__7(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__7); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__8 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__8(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__8); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__9 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__9(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__9); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__10 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__10(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__10); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__11 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__11(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__11); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__12 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__12(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__12); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__13 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__13(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__13); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__14 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__14(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__14); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__15 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__15(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__15); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__16 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__16(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__16); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__17 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__17(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__17); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__18 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__18(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__18); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__19 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__19(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__19); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__20 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__20(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__20); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__21 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__21(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__21); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__22 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__22(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809____closed__22); +if (builtin) {res = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_809_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_955____closed__1 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_955____closed__1(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_955____closed__1); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_955____closed__2 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_955____closed__2(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_955____closed__2); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_955____closed__3 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_955____closed__3(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_955____closed__3); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_955____closed__4 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_955____closed__4(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_955____closed__4); -if (builtin) {res = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_955_(lean_io_mk_world()); +}l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_957____closed__1 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_957____closed__1(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_957____closed__1); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_957____closed__2 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_957____closed__2(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_957____closed__2); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_957____closed__3 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_957____closed__3(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_957____closed__3); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_957____closed__4 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_957____closed__4(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_957____closed__4); +if (builtin) {res = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_957_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Probing.c b/stage0/stdlib/Lean/Compiler/LCNF/Probing.c index 194bd988b7d0..d02a6029eee1 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Probing.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Probing.c @@ -41,9 +41,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Probe LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_filterByJp___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_uint64_to_usize(uint64_t); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Probe_filterByFun___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_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__19; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Probe_filterByLet___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__13; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_filterByUnreach___boxed(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_Compiler_LCNF_Probe_countUnique(lean_object*, lean_object*); @@ -59,6 +57,7 @@ lean_object* lean_mk_array(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Compiler_LCNF_Probe_filterByCases_go___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_tail(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_filterByFun___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__12; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_runOnDeclsNamed___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_filterByFunDecl_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_getDeclCore_x3f(lean_object*, lean_object*, lean_object*); @@ -75,19 +74,20 @@ lean_object* l_Nat_nextPowerOfTwo_go(lean_object*, lean_object*, lean_object*); lean_object* l_List_toString___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Probe_toPass___rarg___closed__1; lean_object* l_Lean_stringToMessageData(lean_object*); -static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__3; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_Probe_runOnDeclsNamed___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_Lean_Compiler_LCNF_Probe_filter___spec__1(lean_object*); +static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__4; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_Probe_toString___spec__1___rarg(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_LCtx_toLocalContext(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Compiler_LCNF_Probe_filterByFun_go___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Compiler_LCNF_Probe_countUniqueSorted___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__15; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_head___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_sum___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_filterByFun(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_filterByCases___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Probe_runOnModule___rarg___closed__3; +static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__6; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_Probe_runOnDeclsNamed___spec__2___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Probe_getLetValues_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_Compiler_LCNF_Probe_filterByReturn_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -97,10 +97,8 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_count(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_countUnique___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_LCNF_Probe_toPass___elambda__1___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Probe_filterByJmp___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__5; lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_sum(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__9; LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Compiler_LCNF_Probe_sorted___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_filterByCases_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_filterByJmp___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -144,28 +142,27 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_countUniqueSorted___boxed(le LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Probe_filterByJp___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_filterByFunDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Compiler_LCNF_AltCore_getCode(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_LCNF_Probe_runOnModule___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Probe_countUniqueSorted___rarg___closed__1; lean_object* l_Std_DHashMap_Internal_AssocList_replace___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_filterByFunDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__19; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_Probe_countUnique___spec__1___rarg(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_Compiler_LCNF_Probe_filterByJp_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addTrace___at_Lean_Compiler_LCNF_Probe_toPass___elambda__1___spec__2___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Probe_filter___spec__1___rarg(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_Compiler_LCNF_Probe_filterByUnreach(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_toPass___elambda__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__7; LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Compiler_LCNF_Probe_sortedBySize___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__18; LEAN_EXPORT uint8_t l_Array_qsort_sort___at_Lean_Compiler_LCNF_Probe_sortedBySize___spec__2___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_count___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_getLetValues_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Probe_getLetValues_go___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__15; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Probe_getJps_start___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_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_map___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__6; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_filterByFun_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_LCNF_Probe_runOnDeclsNamed___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_LCNF_Probe_toPass___elambda__1___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -175,7 +172,6 @@ LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Compiler_LCNF_Probe_sorted static lean_object* l_Lean_Compiler_LCNF_Probe_runOnModule___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_filterByLet(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_getLetValues_start___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__10; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_runGlobally(lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Compiler_LCNF_Probe_countUniqueSorted___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_getJps___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -184,7 +180,9 @@ lean_object* l_Array_append___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_toPass___rarg(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_map(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_DeclValue_forCodeM___at_Lean_Compiler_LCNF_Probe_getLetValues_start___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__3; double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); +lean_object* l_Lean_Compiler_LCNF_Alt_getCode(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_toPass___elambda__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_getJps_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_toPass___elambda__1___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -197,9 +195,8 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_getLetValues_start(lean_obje LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_countUniqueSorted___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_Compiler_LCNF_Probe_toString___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_getJps_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__12; +static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__5; static lean_object* l_Lean_Compiler_LCNF_Probe_toPass___elambda__1___rarg___closed__3; -static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Probe_filterByUnreach___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_Probe_countUnique___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_sorted___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -207,21 +204,23 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_toString(lean_object*); lean_object* l_Lean_Compiler_LCNF_CompilerM_run___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_instInhabited(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_qsort_sort___at_Lean_Compiler_LCNF_Probe_countUniqueSorted___spec__1___rarg___lambda__1(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__2; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Probe_getJps_go___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_filterByJp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_countUnique___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_filter___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__11; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); lean_object* l_Lean_Environment_getModuleIdx_x3f(lean_object*, lean_object*); lean_object* l_Lean_Environment_allImportedModuleNames(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Probe_countUnique___spec__3(lean_object*); +static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__10; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_countUnique___rarg___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_Compiler_LCNF_Probe_filterByFun___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__4; +static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__17; static lean_object* l_Lean_Compiler_LCNF_Probe_runOnDeclsNamed___rarg___closed__1; -static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__16; uint64_t lean_uint64_xor(uint64_t, uint64_t); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_tail___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_Compiler_LCNF_Probe_getJps_go___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -238,22 +237,25 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_sorted___boxed(lean_object*, LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Compiler_LCNF_Probe_countUnique___spec__2(lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); lean_object* l_Array_ofSubarray___rarg(lean_object*); -static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__14; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_sortedBySize(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Compiler_LCNF_Probe_sortedBySize___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Compiler_LCNF_Probe_sorted___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_toPass___elambda__1(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Compiler_LCNF_Probe_filterByJmp_go___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__11; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_declNames___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__13; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_getJps_start(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Decl_size(lean_object*); static lean_object* l_Lean_Compiler_LCNF_Probe_toPass___elambda__1___rarg___closed__4; +static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__18; +static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__14; size_t lean_usize_sub(size_t, size_t); lean_object* lean_array_mk(lean_object*); static lean_object* l_Lean_Compiler_LCNF_Probe_filter___rarg___closed__1; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_Probe_runOnDeclsNamed___spec__2___closed__2; +static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__16; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_filterByLet___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__1; static lean_object* l_Lean_addTrace___at_Lean_Compiler_LCNF_Probe_toPass___elambda__1___spec__2___closed__1; size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Probe_filterByCases___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -263,8 +265,8 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Probe LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Compiler_LCNF_Probe_filterByJmp_go___spec__1___boxed(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___rarg(lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); -static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__8; size_t lean_array_size(lean_object*); +static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__9; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Probe_filterByFunDecl___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_Compiler_LCNF_Probe_toPass(lean_object*); static double l_Lean_addTrace___at_Lean_Compiler_LCNF_Probe_toPass___elambda__1___spec__2___closed__4; @@ -274,12 +276,11 @@ LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Compiler_LCNF_Probe_countU LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Compiler_LCNF_Probe_filterByCases_go___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Compiler_LCNF_Probe_filterByUnreach_go___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Compiler_LCNF_Probe_filterByReturn_go___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__17; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_countUniqueSorted___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_Compiler_LCNF_Probe_filterByJmp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_DeclValue_forCodeM___at_Lean_Compiler_LCNF_Probe_getJps_start___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__7; +static lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__8; extern lean_object* l_Lean_Name_instBEq; lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_toPass___rarg___boxed(lean_object*, lean_object*, lean_object*); @@ -304,7 +305,6 @@ lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Compiler_LCNF_Probe_filterByJp_go___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Name_lt(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972_(lean_object*); lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_countUniqueSorted___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); @@ -2119,7 +2119,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; size_t x_17; size_t x_18; lean_dec(x_4); x_12 = lean_array_uget(x_1, x_2); -x_13 = l_Lean_Compiler_LCNF_AltCore_getCode(x_12); +x_13 = l_Lean_Compiler_LCNF_Alt_getCode(x_12); lean_dec(x_12); x_14 = l_Lean_Compiler_LCNF_Probe_getLetValues_go(x_13, x_5, x_6, x_7, x_8, x_9, x_10); x_15 = lean_ctor_get(x_14, 0); @@ -2602,7 +2602,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; size_t x_17; size_t x_18; lean_dec(x_4); x_12 = lean_array_uget(x_1, x_2); -x_13 = l_Lean_Compiler_LCNF_AltCore_getCode(x_12); +x_13 = l_Lean_Compiler_LCNF_Alt_getCode(x_12); lean_dec(x_12); x_14 = l_Lean_Compiler_LCNF_Probe_getJps_go(x_13, x_5, x_6, x_7, x_8, x_9, x_10); x_15 = lean_ctor_get(x_14, 0); @@ -3088,7 +3088,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; x_11 = lean_array_uget(x_2, x_3); -x_12 = l_Lean_Compiler_LCNF_AltCore_getCode(x_11); +x_12 = l_Lean_Compiler_LCNF_Alt_getCode(x_11); lean_dec(x_11); lean_inc(x_8); lean_inc(x_7); @@ -3723,7 +3723,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; x_11 = lean_array_uget(x_2, x_3); -x_12 = l_Lean_Compiler_LCNF_AltCore_getCode(x_11); +x_12 = l_Lean_Compiler_LCNF_Alt_getCode(x_11); lean_dec(x_11); lean_inc(x_8); lean_inc(x_7); @@ -4328,7 +4328,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; x_11 = lean_array_uget(x_2, x_3); -x_12 = l_Lean_Compiler_LCNF_AltCore_getCode(x_11); +x_12 = l_Lean_Compiler_LCNF_Alt_getCode(x_11); lean_dec(x_11); lean_inc(x_8); lean_inc(x_7); @@ -5010,7 +5010,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; x_11 = lean_array_uget(x_2, x_3); -x_12 = l_Lean_Compiler_LCNF_AltCore_getCode(x_11); +x_12 = l_Lean_Compiler_LCNF_Alt_getCode(x_11); lean_dec(x_11); lean_inc(x_8); lean_inc(x_7); @@ -5615,7 +5615,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; x_11 = lean_array_uget(x_2, x_3); -x_12 = l_Lean_Compiler_LCNF_AltCore_getCode(x_11); +x_12 = l_Lean_Compiler_LCNF_Alt_getCode(x_11); lean_dec(x_11); lean_inc(x_8); lean_inc(x_7); @@ -6262,7 +6262,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; x_11 = lean_array_uget(x_2, x_3); -x_12 = l_Lean_Compiler_LCNF_AltCore_getCode(x_11); +x_12 = l_Lean_Compiler_LCNF_Alt_getCode(x_11); lean_dec(x_11); lean_inc(x_8); lean_inc(x_7); @@ -6757,7 +6757,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; x_11 = lean_array_uget(x_2, x_3); -x_12 = l_Lean_Compiler_LCNF_AltCore_getCode(x_11); +x_12 = l_Lean_Compiler_LCNF_Alt_getCode(x_11); lean_dec(x_11); lean_inc(x_8); lean_inc(x_7); @@ -7269,7 +7269,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; x_11 = lean_array_uget(x_2, x_3); -x_12 = l_Lean_Compiler_LCNF_AltCore_getCode(x_11); +x_12 = l_Lean_Compiler_LCNF_Alt_getCode(x_11); lean_dec(x_11); lean_inc(x_8); lean_inc(x_7); @@ -9608,7 +9608,7 @@ x_5 = l_Lean_Compiler_LCNF_Probe_toPass___rarg(x_1, x_2, x_4); return x_5; } } -static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__1() { +static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__1() { _start: { lean_object* x_1; @@ -9616,27 +9616,27 @@ x_1 = lean_mk_string_unchecked("Lean", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__2() { +static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__1; +x_2 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__3() { +static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__2; +x_1 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__2; x_2 = l_Lean_Compiler_LCNF_Probe_toPass___elambda__1___rarg___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__4() { +static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__4() { _start: { lean_object* x_1; @@ -9644,17 +9644,17 @@ x_1 = lean_mk_string_unchecked("LCNF", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__5() { +static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__3; -x_2 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__4; +x_1 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__3; +x_2 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__6() { +static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__6() { _start: { lean_object* x_1; @@ -9662,17 +9662,17 @@ x_1 = lean_mk_string_unchecked("Probe", 5, 5); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__7() { +static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__5; -x_2 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__6; +x_1 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__5; +x_2 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__8() { +static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__8() { _start: { lean_object* x_1; @@ -9680,17 +9680,17 @@ x_1 = lean_mk_string_unchecked("initFn", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__9() { +static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__7; -x_2 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__8; +x_1 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__7; +x_2 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__8; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__10() { +static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__10() { _start: { lean_object* x_1; @@ -9698,47 +9698,47 @@ x_1 = lean_mk_string_unchecked("_@", 2, 2); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__11() { +static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__9; -x_2 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__10; +x_1 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__9; +x_2 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__12() { +static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__11; -x_2 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__1; +x_1 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__11; +x_2 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__13() { +static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__12; +x_1 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__12; x_2 = l_Lean_Compiler_LCNF_Probe_toPass___elambda__1___rarg___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__14() { +static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__13; -x_2 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__4; +x_1 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__13; +x_2 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__15() { +static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__15() { _start: { lean_object* x_1; @@ -9746,17 +9746,17 @@ x_1 = lean_mk_string_unchecked("Probing", 7, 7); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__16() { +static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__14; -x_2 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__15; +x_1 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__14; +x_2 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__15; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__17() { +static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__17() { _start: { lean_object* x_1; @@ -9764,33 +9764,33 @@ x_1 = lean_mk_string_unchecked("_hyg", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__18() { +static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__16; -x_2 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__17; +x_1 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__16; +x_2 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__17; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__19() { +static lean_object* _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__18; -x_2 = lean_unsigned_to_nat(2972u); +x_1 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__18; +x_2 = lean_unsigned_to_nat(2964u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_Compiler_LCNF_Probe_toPass___elambda__1___rarg___closed__3; x_3 = 1; -x_4 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__19; +x_4 = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__19; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -9873,45 +9873,45 @@ l_Lean_Compiler_LCNF_Probe_toPass___elambda__1___rarg___closed__4 = _init_l_Lean lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_toPass___elambda__1___rarg___closed__4); l_Lean_Compiler_LCNF_Probe_toPass___rarg___closed__1 = _init_l_Lean_Compiler_LCNF_Probe_toPass___rarg___closed__1(); lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_toPass___rarg___closed__1); -l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__1 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__1(); -lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__1); -l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__2 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__2(); -lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__2); -l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__3 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__3(); -lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__3); -l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__4 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__4(); -lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__4); -l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__5 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__5(); -lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__5); -l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__6 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__6(); -lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__6); -l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__7 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__7(); -lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__7); -l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__8 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__8(); -lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__8); -l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__9 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__9(); -lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__9); -l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__10 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__10(); -lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__10); -l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__11 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__11(); -lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__11); -l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__12 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__12(); -lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__12); -l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__13 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__13(); -lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__13); -l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__14 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__14(); -lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__14); -l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__15 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__15(); -lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__15); -l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__16 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__16(); -lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__16); -l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__17 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__17(); -lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__17); -l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__18 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__18(); -lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__18); -l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__19 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__19(); -lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972____closed__19); -if (builtin) {res = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2972_(lean_io_mk_world()); +l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__1 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__1(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__1); +l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__2 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__2(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__2); +l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__3 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__3(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__3); +l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__4 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__4(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__4); +l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__5 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__5(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__5); +l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__6 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__6(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__6); +l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__7 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__7(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__7); +l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__8 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__8(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__8); +l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__9 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__9(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__9); +l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__10 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__10(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__10); +l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__11 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__11(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__11); +l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__12 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__12(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__12); +l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__13 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__13(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__13); +l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__14 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__14(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__14); +l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__15 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__15(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__15); +l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__16 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__16(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__16); +l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__17 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__17(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__17); +l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__18 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__18(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__18); +l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__19 = _init_l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__19(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964____closed__19); +if (builtin) {res = l_Lean_Compiler_LCNF_Probe_initFn____x40_Lean_Compiler_LCNF_Probing___hyg_2964_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Compiler/LCNF/PullFunDecls.c b/stage0/stdlib/Lean/Compiler/LCNF/PullFunDecls.c index 208e62e8d661..2cd38d02e96a 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/PullFunDecls.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/PullFunDecls.c @@ -82,6 +82,7 @@ static lean_object* l_Lean_Compiler_LCNF_Decl_pullFunDecls___closed__1; lean_object* l_Array_append___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_PullFunDecls_attach___spec__1(size_t, size_t, lean_object*); static lean_object* l_Lean_Compiler_LCNF_pullFunDecls___closed__3; +lean_object* l_Lean_Compiler_LCNF_FunDecl_collectUsed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_PullFunDecls_ToPull_attach___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_pullFunDecls; lean_object* lean_array_fget(lean_object*, lean_object*); @@ -117,7 +118,6 @@ static lean_object* l_Lean_Compiler_LCNF_PullFunDecls_instInhabitedToPull___clos LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_PullFunDecls_findFVarDirectDeps___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_PullFunDecls_attachJps(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Compiler_LCNF_FunDeclCore_collectUsed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_PullFunDecls_attachJps___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_PullFunDecls_pull___closed__1; static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PullFunDecls___hyg_1780____closed__16; @@ -1501,7 +1501,7 @@ x_19 = lean_ctor_get(x_17, 0); x_20 = lean_ctor_get(x_17, 1); x_21 = lean_box(0); lean_inc(x_15); -x_22 = l_Lean_Compiler_LCNF_FunDeclCore_collectUsed(x_15, x_21); +x_22 = l_Lean_Compiler_LCNF_FunDecl_collectUsed(x_15, x_21); x_23 = lean_alloc_ctor(0, 2, 1); lean_ctor_set(x_23, 0, x_15); lean_ctor_set(x_23, 1, x_22); @@ -1544,7 +1544,7 @@ lean_inc(x_32); lean_dec(x_17); x_34 = lean_box(0); lean_inc(x_15); -x_35 = l_Lean_Compiler_LCNF_FunDeclCore_collectUsed(x_15, x_34); +x_35 = l_Lean_Compiler_LCNF_FunDecl_collectUsed(x_15, x_34); x_36 = lean_alloc_ctor(0, 2, 1); lean_ctor_set(x_36, 0, x_15); lean_ctor_set(x_36, 1, x_35); diff --git a/stage0/stdlib/Lean/Compiler/LCNF/PullLetDecls.c b/stage0/stdlib/Lean/Compiler/LCNF/PullLetDecls.c index c5e428711d4b..9003bdd174b6 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/PullLetDecls.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/PullLetDecls.c @@ -1664,7 +1664,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_PullLetDecls_pullDecls___lambda__ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Compiler_LCNF_PullLetDecls_pullDecls___lambda__1___closed__1; x_2 = l_Lean_Compiler_LCNF_PullLetDecls_pullDecls___lambda__1___closed__2; -x_3 = lean_unsigned_to_nat(305u); +x_3 = lean_unsigned_to_nat(306u); x_4 = lean_unsigned_to_nat(9u); x_5 = l_Lean_Compiler_LCNF_PullLetDecls_pullDecls___lambda__1___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -2818,6 +2818,7 @@ lean_inc(x_91); lean_dec(x_87); x_92 = l_Lean_Compiler_LCNF_PullLetDecls_pullDecls___closed__2; x_93 = lean_name_eq(x_88, x_92); +x_93 = 0; // bootstrap: disable if (x_93 == 0) { lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; diff --git a/stage0/stdlib/Lean/Compiler/LCNF/ReduceArity.c b/stage0/stdlib/Lean/Compiler/LCNF/ReduceArity.c index 49f2eba6c1ba..8df2e3b51873 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/ReduceArity.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/ReduceArity.c @@ -97,7 +97,6 @@ lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_ReduceArity___hyg_2393____closed__16; lean_object* l_Lean_Name_append(lean_object*, lean_object*); uint8_t l_Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Compiler_LCNF_AltCore_getCode(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Decl_reduceArity___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at_Lean_Compiler_LCNF_FindUsed_visitLetValue___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_Compiler_LCNF_ReduceArity_reduce___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -113,6 +112,7 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_De lean_object* l_Array_append___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_LCNF_Decl_reduceArity___spec__13(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_Compiler_LCNF_Alt_getCode(lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_LCNF_Decl_reduceArity___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FindUsed_visitFVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_reduceArity___closed__2; @@ -1046,7 +1046,7 @@ if (x_12 == 0) lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; lean_dec(x_4); x_13 = lean_array_uget(x_1, x_2); -x_14 = l_Lean_Compiler_LCNF_AltCore_getCode(x_13); +x_14 = l_Lean_Compiler_LCNF_Alt_getCode(x_13); lean_dec(x_13); lean_inc(x_5); x_15 = l_Lean_Compiler_LCNF_FindUsed_visit(x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11); @@ -2264,7 +2264,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ReduceArity_reduce___lambda__2(lea _start: { lean_object* x_8; lean_object* x_9; -x_8 = l_Lean_Compiler_LCNF_AltCore_getCode(x_1); +x_8 = l_Lean_Compiler_LCNF_Alt_getCode(x_1); x_9 = l_Lean_Compiler_LCNF_ReduceArity_reduce(x_8, x_2, x_3, x_4, x_5, x_6, x_7); if (lean_obj_tag(x_9) == 0) { diff --git a/stage0/stdlib/Lean/Compiler/LCNF/ReduceJpArity.c b/stage0/stdlib/Lean/Compiler/LCNF/ReduceJpArity.c index 69b37f1c6a54..661befcbc8fb 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/ReduceJpArity.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/ReduceJpArity.c @@ -38,7 +38,6 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_reduceJpArity(uint8_t); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_ReduceJpArity___hyg_879____closed__16; static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_ReduceJpArity___hyg_879____closed__6; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); -lean_object* l_Lean_Compiler_LCNF_AltCore_getCode(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_ReduceJpArity_reduce___spec__4(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_ReduceJpArity___hyg_879____closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ReduceJpArity_reduce___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -53,6 +52,7 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_Re LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_ReduceJpArity___hyg_879_(lean_object*); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_ReduceJpArity___hyg_879____closed__9; lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_collectType(lean_object*, lean_object*); +lean_object* l_Lean_Compiler_LCNF_Alt_getCode(lean_object*); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_ReduceJpArity___hyg_879____closed__13; static lean_object* l_Lean_Compiler_LCNF_ReduceJpArity_reduce___closed__2; static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_ReduceJpArity___hyg_879____closed__14; @@ -791,7 +791,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ReduceJpArity_reduce___lambda__1(l _start: { lean_object* x_8; lean_object* x_9; -x_8 = l_Lean_Compiler_LCNF_AltCore_getCode(x_1); +x_8 = l_Lean_Compiler_LCNF_Alt_getCode(x_1); x_9 = l_Lean_Compiler_LCNF_ReduceJpArity_reduce(x_8, x_2, x_3, x_4, x_5, x_6, x_7); if (lean_obj_tag(x_9) == 0) { diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Renaming.c b/stage0/stdlib/Lean/Compiler/LCNF/Renaming.c index 46d77fba702a..36c9a63dd209 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Renaming.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Renaming.c @@ -32,6 +32,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_LetDecl_applyRenaming(lean_object* lean_object* l_Lean_Compiler_LCNF_LCtx_addLetDecl(lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_LCtx_addFunDecl(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_applyRenaming___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDecl_applyRenaming(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_applyRenaming___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_LCtx_addParam(lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); @@ -40,7 +41,6 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Decl_applyRenaming(lean_object*, l LEAN_EXPORT lean_object* l___private_Init_Data_Array_BasicAux_0__mapMonoMImp_go___at_Lean_Compiler_LCNF_Code_applyRenaming___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltCodeImp(lean_object*, lean_object*); lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltImp(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_applyRenaming(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); @@ -413,7 +413,7 @@ lean_dec(x_2); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_applyRenaming(lean_object* x_1, lean_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_Compiler_LCNF_FunDecl_applyRenaming(lean_object* x_1, lean_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; @@ -1362,7 +1362,7 @@ lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); lean_inc(x_49); -x_51 = l_Lean_Compiler_LCNF_FunDeclCore_applyRenaming(x_49, x_2, x_3, x_4, x_5, x_6, x_7); +x_51 = l_Lean_Compiler_LCNF_FunDecl_applyRenaming(x_49, x_2, x_3, x_4, x_5, x_6, x_7); if (lean_obj_tag(x_51) == 0) { lean_object* x_52; lean_object* x_53; lean_object* x_54; @@ -1606,7 +1606,7 @@ lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); lean_inc(x_94); -x_96 = l_Lean_Compiler_LCNF_FunDeclCore_applyRenaming(x_94, x_2, x_3, x_4, x_5, x_6, x_7); +x_96 = l_Lean_Compiler_LCNF_FunDecl_applyRenaming(x_94, x_2, x_3, x_4, x_5, x_6, x_7); if (lean_obj_tag(x_96) == 0) { lean_object* x_97; lean_object* x_98; lean_object* x_99; diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Simp/DefaultAlt.c b/stage0/stdlib/Lean/Compiler/LCNF/Simp/DefaultAlt.c index e75f7217ef98..fb5cd4d096a0 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Simp/DefaultAlt.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Simp/DefaultAlt.c @@ -13,18 +13,16 @@ #ifdef __cplusplus extern "C" { #endif -static lean_object* l___private_Lean_Compiler_LCNF_Simp_DefaultAlt_0__Lean_Compiler_LCNF_Simp_getMaxOccs_getNumOccsOf___closed__1; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Compiler_LCNF_Simp_DefaultAlt_0__Lean_Compiler_LCNF_Simp_getMaxOccs___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_addDefaultAlt(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Compiler_LCNF_Simp_addDefaultAlt___spec__1___closed__1; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_Simp_addDefaultAlt___spec__2___closed__3; lean_object* lean_array_push(lean_object*, lean_object*); -extern lean_object* l_Lean_Compiler_LCNF_instInhabitedCode; uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Compiler_LCNF_Simp_addDefaultAlt___spec__3___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_instInhabitedReaderT___rarg___boxed(lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Compiler_LCNF_Simp_addDefaultAlt___spec__1___closed__2; -lean_object* l_Lean_Compiler_LCNF_instInhabitedAltCore___rarg(lean_object*); +extern lean_object* l_Lean_Compiler_LCNF_instInhabitedAlt; LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Simp_DefaultAlt_0__Lean_Compiler_LCNF_Simp_getMaxOccs___boxed(lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Compiler_LCNF_Simp_addDefaultAlt___spec__3(lean_object*, size_t, size_t); lean_object* l_Lean_Compiler_LCNF_Simp_markSimplified___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -39,12 +37,12 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_Si static lean_object* l_Lean_Compiler_LCNF_Simp_addDefaultAlt___closed__1; static lean_object* l_panic___at_Lean_Compiler_LCNF_Simp_addDefaultAlt___spec__1___closed__3; lean_object* l_Lean_Compiler_LCNF_eraseCode(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Compiler_LCNF_AltCore_getCode(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Compiler_LCNF_Simp_DefaultAlt_0__Lean_Compiler_LCNF_Simp_getMaxOccs___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_Simp_addDefaultAlt___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Compiler_LCNF_instMonadCompilerM; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_Simp_addDefaultAlt___spec__2___closed__4; +lean_object* l_Lean_Compiler_LCNF_Alt_getCode(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Simp_DefaultAlt_0__Lean_Compiler_LCNF_Simp_getMaxOccs(lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); @@ -85,7 +83,7 @@ else { lean_object* x_11; lean_object* x_12; uint8_t x_13; x_11 = lean_array_fget(x_1, x_6); -x_12 = l_Lean_Compiler_LCNF_AltCore_getCode(x_11); +x_12 = l_Lean_Compiler_LCNF_Alt_getCode(x_11); lean_dec(x_11); lean_inc(x_3); x_13 = l_Lean_Compiler_LCNF_Code_alphaEqv(x_12, x_3); @@ -118,22 +116,13 @@ goto _start; } } } -static lean_object* _init_l___private_Lean_Compiler_LCNF_Simp_DefaultAlt_0__Lean_Compiler_LCNF_Simp_getMaxOccs_getNumOccsOf___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Compiler_LCNF_instInhabitedCode; -x_2 = l_Lean_Compiler_LCNF_instInhabitedAltCore___rarg(x_1); -return x_2; -} -} LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Simp_DefaultAlt_0__Lean_Compiler_LCNF_Simp_getMaxOccs_getNumOccsOf(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_3 = l___private_Lean_Compiler_LCNF_Simp_DefaultAlt_0__Lean_Compiler_LCNF_Simp_getMaxOccs_getNumOccsOf___closed__1; +x_3 = l_Lean_Compiler_LCNF_instInhabitedAlt; x_4 = lean_array_get(x_3, x_1, x_2); -x_5 = l_Lean_Compiler_LCNF_AltCore_getCode(x_4); +x_5 = l_Lean_Compiler_LCNF_Alt_getCode(x_4); lean_dec(x_4); x_6 = lean_unsigned_to_nat(1u); x_7 = lean_nat_add(x_2, x_6); @@ -272,7 +261,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Simp_DefaultAlt_0__Lean_ _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; uint8_t x_11; -x_2 = l___private_Lean_Compiler_LCNF_Simp_DefaultAlt_0__Lean_Compiler_LCNF_Simp_getMaxOccs_getNumOccsOf___closed__1; +x_2 = l_Lean_Compiler_LCNF_instInhabitedAlt; x_3 = lean_unsigned_to_nat(0u); x_4 = lean_array_get(x_2, x_1, x_3); x_5 = l___private_Lean_Compiler_LCNF_Simp_DefaultAlt_0__Lean_Compiler_LCNF_Simp_getMaxOccs_getNumOccsOf(x_1, x_3); @@ -463,8 +452,8 @@ if (x_26 == 0) lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; x_27 = lean_ctor_get(x_7, 0); x_28 = lean_ctor_get(x_7, 1); -x_29 = l_Lean_Compiler_LCNF_AltCore_getCode(x_18); -x_30 = l_Lean_Compiler_LCNF_AltCore_getCode(x_2); +x_29 = l_Lean_Compiler_LCNF_Alt_getCode(x_18); +x_30 = l_Lean_Compiler_LCNF_Alt_getCode(x_2); x_31 = l_Lean_Compiler_LCNF_Code_alphaEqv(x_29, x_30); if (x_31 == 0) { @@ -672,8 +661,8 @@ x_70 = lean_ctor_get(x_7, 1); lean_inc(x_70); lean_inc(x_69); lean_dec(x_7); -x_71 = l_Lean_Compiler_LCNF_AltCore_getCode(x_18); -x_72 = l_Lean_Compiler_LCNF_AltCore_getCode(x_2); +x_71 = l_Lean_Compiler_LCNF_Alt_getCode(x_18); +x_72 = l_Lean_Compiler_LCNF_Alt_getCode(x_2); x_73 = l_Lean_Compiler_LCNF_Code_alphaEqv(x_71, x_72); if (x_73 == 0) { @@ -1001,7 +990,7 @@ x_25 = lean_ctor_get(x_23, 0); x_26 = lean_ctor_get(x_25, 0); lean_inc(x_26); lean_dec(x_25); -x_27 = l_Lean_Compiler_LCNF_AltCore_getCode(x_13); +x_27 = l_Lean_Compiler_LCNF_Alt_getCode(x_13); lean_dec(x_13); x_28 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_28, 0, x_27); @@ -1020,7 +1009,7 @@ lean_dec(x_23); x_32 = lean_ctor_get(x_30, 0); lean_inc(x_32); lean_dec(x_30); -x_33 = l_Lean_Compiler_LCNF_AltCore_getCode(x_13); +x_33 = l_Lean_Compiler_LCNF_Alt_getCode(x_13); lean_dec(x_13); x_34 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_34, 0, x_33); @@ -1112,7 +1101,7 @@ if (lean_is_exclusive(x_51)) { x_55 = lean_ctor_get(x_52, 0); lean_inc(x_55); lean_dec(x_52); -x_56 = l_Lean_Compiler_LCNF_AltCore_getCode(x_41); +x_56 = l_Lean_Compiler_LCNF_Alt_getCode(x_41); lean_dec(x_41); x_57 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_57, 0, x_56); @@ -1229,8 +1218,6 @@ _G_initialized = true; res = initialize_Lean_Compiler_LCNF_Simp_SimpM(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___private_Lean_Compiler_LCNF_Simp_DefaultAlt_0__Lean_Compiler_LCNF_Simp_getMaxOccs_getNumOccsOf___closed__1 = _init_l___private_Lean_Compiler_LCNF_Simp_DefaultAlt_0__Lean_Compiler_LCNF_Simp_getMaxOccs_getNumOccsOf___closed__1(); -lean_mark_persistent(l___private_Lean_Compiler_LCNF_Simp_DefaultAlt_0__Lean_Compiler_LCNF_Simp_getMaxOccs_getNumOccsOf___closed__1); l_panic___at_Lean_Compiler_LCNF_Simp_addDefaultAlt___spec__1___closed__1 = _init_l_panic___at_Lean_Compiler_LCNF_Simp_addDefaultAlt___spec__1___closed__1(); lean_mark_persistent(l_panic___at_Lean_Compiler_LCNF_Simp_addDefaultAlt___spec__1___closed__1); l_panic___at_Lean_Compiler_LCNF_Simp_addDefaultAlt___spec__1___closed__2 = _init_l_panic___at_Lean_Compiler_LCNF_Simp_addDefaultAlt___spec__1___closed__2(); diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Simp/FunDeclInfo.c b/stage0/stdlib/Lean/Compiler/LCNF/Simp/FunDeclInfo.c index 6d31bce0280c..95ebae87429f 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Simp/FunDeclInfo.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Simp/FunDeclInfo.c @@ -60,7 +60,6 @@ static lean_object* l___private_Lean_Compiler_LCNF_Simp_FunDeclInfo_0__Lean_Comp LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Compiler_LCNF_Simp_FunDeclInfoMap_add___spec__5(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_object* l_Lean_Compiler_LCNF_AltCore_getCode(lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Simp_FunDeclInfo_0__Lean_Compiler_LCNF_Simp_reprFunDeclInfo____x40_Lean_Compiler_LCNF_Simp_FunDeclInfo___hyg_10____closed__14; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Compiler_LCNF_Simp_FunDeclInfoMap_add___spec__6___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Compiler_LCNF_Simp_FunDeclInfoMap_format___spec__1___lambda__1___boxed(lean_object*); @@ -75,6 +74,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_FunDeclInfoMap_update_addArgO static lean_object* l___private_Lean_Compiler_LCNF_Simp_FunDeclInfo_0__Lean_Compiler_LCNF_Simp_reprFunDeclInfo____x40_Lean_Compiler_LCNF_Simp_FunDeclInfo___hyg_10____closed__16; static lean_object* l___private_Lean_Compiler_LCNF_Simp_FunDeclInfo_0__Lean_Compiler_LCNF_Simp_reprFunDeclInfo____x40_Lean_Compiler_LCNF_Simp_FunDeclInfo___hyg_10____closed__9; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldrM___at_Lean_Compiler_LCNF_Simp_FunDeclInfoMap_format___spec__2(lean_object*, lean_object*); +lean_object* l_Lean_Compiler_LCNF_Alt_getCode(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Simp_FunDeclInfoMap_update_addLetValueOccs___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Simp_FunDeclInfo_0__Lean_Compiler_LCNF_Simp_reprFunDeclInfo____x40_Lean_Compiler_LCNF_Simp_FunDeclInfo___hyg_10____closed__19; lean_object* lean_array_fget(lean_object*, lean_object*); @@ -2865,7 +2865,7 @@ if (x_12 == 0) lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_dec(x_5); x_13 = lean_array_uget(x_2, x_3); -x_14 = l_Lean_Compiler_LCNF_AltCore_getCode(x_13); +x_14 = l_Lean_Compiler_LCNF_Alt_getCode(x_13); lean_dec(x_13); x_15 = l_Lean_Compiler_LCNF_Simp_FunDeclInfoMap_update_go(x_1, x_14, x_6, x_7, x_8, x_9, x_10, x_11); if (lean_obj_tag(x_15) == 0) diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Simp/InlineCandidate.c b/stage0/stdlib/Lean/Compiler/LCNF/Simp/InlineCandidate.c index 1bbf8443ddb7..b6e4de1d5b46 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Simp/InlineCandidate.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Simp/InlineCandidate.c @@ -292,6 +292,7 @@ lean_object* x_16; lean_object* x_17; uint8_t x_18; x_16 = lean_ctor_get(x_2, 0); x_17 = l_Lean_Compiler_LCNF_Simp_inlineCandidate_x3f___lambda__5___closed__2; x_18 = lean_name_eq(x_16, x_17); +x_18 = 0; // bootstrap: disable if (x_18 == 0) { uint8_t x_19; lean_object* x_20; lean_object* x_21; diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Simp/JpCases.c b/stage0/stdlib/Lean/Compiler/LCNF/Simp/JpCases.c index fb75519b0a9a..f6a5cc5dd68b 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Simp/JpCases.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Simp/JpCases.c @@ -112,6 +112,7 @@ static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_J LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simpJpCases_x3f_visitJp_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_mkJpAlt_go___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_Compiler_LCNF_instInhabitedCases; lean_object* l_Lean_RBNode_insert___at_Lean_FVarIdMap_insert___spec__1___rarg(lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_collectJpCasesInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -124,14 +125,12 @@ extern lean_object* l_Lean_Compiler_LCNF_instInhabitedArg; lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_LCNF_Simp_simpJpCases_x3f___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_mkJpAlt_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Compiler_LCNF_CasesCore_getCtorNames(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_collectJpCasesInfo___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_mkJpAlt_go___closed__1; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_mkJpAlt_go___spec__3(lean_object*, lean_object*, size_t, size_t); lean_object* l_Array_append___rarg(lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofExpr(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simpJpCases_x3f_visitJp_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Compiler_LCNF_instInhabitedCasesCore(lean_object*); lean_object* l_Lean_Compiler_LCNF_attachCodeDecls(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBTree_toList___at_Lean_Compiler_LCNF_Simp_simpJpCases_x3f___spec__3(lean_object*); double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); @@ -167,12 +166,12 @@ extern lean_object* l_Lean_Compiler_LCNF_instInhabitedParam; static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3101____closed__19; static lean_object* l_panic___at___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_extractJpCases_go___spec__1___closed__1; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_mkJpAlt_go___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Compiler_LCNF_Cases_getCtorNames(lean_object*); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3101____closed__5; lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); lean_object* l_Array_ofSubarray___rarg(lean_object*); lean_object* l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_JpCasesInfoMap_isCandidate___boxed(lean_object*); -static lean_object* l_panic___at___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_extractJpCases_go___spec__1___closed__3; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_mkJpAlt_go___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3101____closed__8; lean_object* lean_array_mk(lean_object*); @@ -1489,17 +1488,9 @@ return x_2; static lean_object* _init_l_panic___at___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_extractJpCases_go___spec__1___closed__2() { _start: { -lean_object* x_1; -x_1 = l_Lean_Compiler_LCNF_instInhabitedCasesCore(lean_box(0)); -return x_1; -} -} -static lean_object* _init_l_panic___at___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_extractJpCases_go___spec__1___closed__3() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_panic___at___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_extractJpCases_go___spec__1___closed__1; -x_2 = l_panic___at___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_extractJpCases_go___spec__1___closed__2; +x_2 = l_Lean_Compiler_LCNF_instInhabitedCases; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -1510,7 +1501,7 @@ LEAN_EXPORT lean_object* l_panic___at___private_Lean_Compiler_LCNF_Simp_JpCases_ _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_panic___at___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_extractJpCases_go___spec__1___closed__3; +x_2 = l_panic___at___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_extractJpCases_go___spec__1___closed__2; x_3 = lean_panic_fn(x_2, x_1); return x_3; } @@ -6640,7 +6631,7 @@ lean_inc(x_149); x_150 = lean_ctor_get(x_148, 1); lean_inc(x_150); lean_dec(x_148); -x_151 = l_Lean_Compiler_LCNF_CasesCore_getCtorNames(x_4); +x_151 = l_Lean_Compiler_LCNF_Cases_getCtorNames(x_4); x_152 = lean_ctor_get(x_2, 1); lean_inc(x_152); x_153 = l_Lean_RBNode_any___at_Lean_Compiler_LCNF_Simp_simpJpCases_x3f_visitJp_x3f___spec__1(x_151, x_152); @@ -6985,7 +6976,7 @@ lean_inc(x_230); x_231 = lean_ctor_get(x_229, 1); lean_inc(x_231); lean_dec(x_229); -x_232 = l_Lean_Compiler_LCNF_CasesCore_getCtorNames(x_4); +x_232 = l_Lean_Compiler_LCNF_Cases_getCtorNames(x_4); x_233 = lean_ctor_get(x_2, 1); lean_inc(x_233); x_234 = l_Lean_RBNode_any___at_Lean_Compiler_LCNF_Simp_simpJpCases_x3f_visitJp_x3f___spec__1(x_232, x_233); @@ -9191,8 +9182,6 @@ l_panic___at___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Sim lean_mark_persistent(l_panic___at___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_extractJpCases_go___spec__1___closed__1); l_panic___at___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_extractJpCases_go___spec__1___closed__2 = _init_l_panic___at___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_extractJpCases_go___spec__1___closed__2(); lean_mark_persistent(l_panic___at___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_extractJpCases_go___spec__1___closed__2); -l_panic___at___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_extractJpCases_go___spec__1___closed__3 = _init_l_panic___at___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_extractJpCases_go___spec__1___closed__3(); -lean_mark_persistent(l_panic___at___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_extractJpCases_go___spec__1___closed__3); l___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_extractJpCases_go___closed__1 = _init_l___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_extractJpCases_go___closed__1(); lean_mark_persistent(l___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_extractJpCases_go___closed__1); l___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_extractJpCases_go___closed__2 = _init_l___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_extractJpCases_go___closed__2(); diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Simp/Main.c b/stage0/stdlib/Lean/Compiler/LCNF/Simp/Main.c index cd82c5d2c972..1e7e971a1c4d 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Simp/Main.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Simp/Main.c @@ -28,6 +28,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___lambda__1(lean_object* uint8_t l_Lean_Compiler_LCNF_Code_isFun(lean_object*); lean_object* l_Lean_Compiler_LCNF_Simp_eraseLetDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_type(lean_object*); +lean_object* l_Lean_Compiler_LCNF_FunDecl_etaExpand(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_Simp_withDiscrCtorImp_updateCtx___spec__5(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_inlineApp_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -40,7 +41,6 @@ lean_object* l_Lean_Compiler_LCNF_findFunDecl_x3f(lean_object*, lean_object*, le LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_elimVar_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_Compiler_LCNF_instInhabitedCode; lean_object* l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_updateParamImp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Simp_updateFunDeclInfo(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Code_internalize(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -52,7 +52,7 @@ lean_object* lean_mk_array(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Compiler_LCNF_instInhabitedAltCore___rarg(lean_object*); +lean_object* l_Lean_Compiler_LCNF_Cases_extractAlt_x21(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_normParam___at_Lean_Compiler_LCNF_Simp_simpFunDecl___spec__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Simp_etaPolyApp_x3f___closed__3; lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*, uint8_t); @@ -61,11 +61,11 @@ lean_object* l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_upda lean_object* l_Lean_Compiler_LCNF_Simp_eraseFunDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_normParams___at_Lean_Compiler_LCNF_Simp_simpFunDecl___spec__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Nat_nextPowerOfTwo_go(lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_Compiler_LCNF_instInhabitedAlt; LEAN_EXPORT lean_object* l___private_Init_Data_Array_BasicAux_0__mapMonoMImp_go___at_Lean_Compiler_LCNF_Simp_simpFunDecl___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* l_Lean_Compiler_LCNF_Simp_inlineCandidate_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Simp_markSimplified___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Decl_getArity(lean_object*); -lean_object* l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_normLetDecl___at_Lean_Compiler_LCNF_Simp_simp___spec__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_inlineApp_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Simp_Main_0__Lean_Compiler_LCNF_Simp_oneExitPointQuick_go(lean_object*); @@ -77,9 +77,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_isReturnOf(lean_object*, lean size_t lean_usize_of_nat(lean_object*); static lean_object* l_Lean_Compiler_LCNF_Simp_specializePartialApp___closed__2; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_Simp_specializePartialApp___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_Compiler_LCNF_Simp_Main_0__Lean_Compiler_LCNF_Simp_oneExitPointQuick_go___closed__1; lean_object* l_Lean_Compiler_LCNF_inferAppType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetDecl____x40_Lean_Compiler_LCNF_Basic___hyg_1722_(lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_normFunDeclImp(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_mkNewParams(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -106,7 +104,6 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___lambda__4(lean_object* lean_object* l_Lean_Compiler_LCNF_Simp_CtorInfo_getName(lean_object*); lean_object* l_Lean_Compiler_LCNF_eraseCode(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Compiler_LCNF_AltCore_getCode(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_inlineApp_x3f___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_Simp_simpCasesOnCtor_x3f___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_Compiler_LCNF_normParams___at_Lean_Compiler_LCNF_Simp_simpFunDecl___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -129,6 +126,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_inlineApp_x3f___lambda__1___b lean_object* l_Lean_Compiler_LCNF_Simp_incVisited___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Simp_markUsedFVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isInstance(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Compiler_LCNF_Alt_getCode(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at_Lean_Compiler_LCNF_Simp_specializePartialApp___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_normArgs___at_Lean_Compiler_LCNF_Simp_simp___spec__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -184,27 +182,18 @@ lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Simp_markUsedLetV lean_object* l_Lean_Compiler_LCNF_mkAuxParam(lean_object*, 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_Compiler_LCNF_mkAuxFunDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetDecl____x40_Lean_Compiler_LCNF_Basic___hyg_1659_(lean_object*, lean_object*); uint8_t l_Lean_Expr_isForall(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___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 uint8_t l___private_Lean_Compiler_LCNF_Simp_Main_0__Lean_Compiler_LCNF_Simp_oneExitPointQuick(lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); uint8_t l_Lean_Compiler_LCNF_isEtaExpandCandidateCore(lean_object*, lean_object*); -lean_object* l_Lean_Compiler_LCNF_FunDeclCore_etaExpand(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Simp_markUsedLetDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Simp_addFVarSubst(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Simp_attachCodeDecls(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_Lean_Compiler_LCNF_Simp_simp___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_Array_anyMUnsafe_any___at_Lean_Compiler_LCNF_Simp_simp___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* _init_l___private_Lean_Compiler_LCNF_Simp_Main_0__Lean_Compiler_LCNF_Simp_oneExitPointQuick_go___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Compiler_LCNF_instInhabitedCode; -x_2 = l_Lean_Compiler_LCNF_instInhabitedAltCore___rarg(x_1); -return x_2; -} -} LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Simp_Main_0__Lean_Compiler_LCNF_Simp_oneExitPointQuick_go(lean_object* x_1) { _start: { @@ -250,11 +239,11 @@ return x_11; else { lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_12 = l___private_Lean_Compiler_LCNF_Simp_Main_0__Lean_Compiler_LCNF_Simp_oneExitPointQuick_go___closed__1; +x_12 = l_Lean_Compiler_LCNF_instInhabitedAlt; x_13 = lean_unsigned_to_nat(0u); x_14 = lean_array_get(x_12, x_7, x_13); lean_dec(x_7); -x_15 = l_Lean_Compiler_LCNF_AltCore_getCode(x_14); +x_15 = l_Lean_Compiler_LCNF_Alt_getCode(x_14); lean_dec(x_14); x_1 = x_15; goto _start; @@ -4043,7 +4032,7 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_134 = l_Lean_Compiler_LCNF_FunDeclCore_etaExpand(x_132, x_6, x_7, x_8, x_9, x_133); +x_134 = l_Lean_Compiler_LCNF_FunDecl_etaExpand(x_132, x_6, x_7, x_8, x_9, x_133); if (lean_obj_tag(x_134) == 0) { lean_object* x_135; lean_object* x_136; lean_object* x_137; uint8_t x_138; lean_object* x_139; lean_object* x_140; @@ -4909,7 +4898,7 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_286 = l_Lean_Compiler_LCNF_FunDeclCore_etaExpand(x_284, x_6, x_7, x_8, x_9, x_285); +x_286 = l_Lean_Compiler_LCNF_FunDecl_etaExpand(x_284, x_6, x_7, x_8, x_9, x_285); if (lean_obj_tag(x_286) == 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; @@ -6070,7 +6059,7 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_493 = l_Lean_Compiler_LCNF_FunDeclCore_etaExpand(x_491, x_6, x_7, x_8, x_9, x_492); +x_493 = l_Lean_Compiler_LCNF_FunDecl_etaExpand(x_491, x_6, x_7, x_8, x_9, x_492); if (lean_obj_tag(x_493) == 0) { lean_object* x_494; lean_object* x_495; lean_object* x_496; uint8_t x_497; lean_object* x_498; lean_object* x_499; @@ -6936,7 +6925,7 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_645 = l_Lean_Compiler_LCNF_FunDeclCore_etaExpand(x_643, x_6, x_7, x_8, x_9, x_644); +x_645 = l_Lean_Compiler_LCNF_FunDecl_etaExpand(x_643, x_6, x_7, x_8, x_9, x_644); 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; @@ -8098,7 +8087,7 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_852 = l_Lean_Compiler_LCNF_FunDeclCore_etaExpand(x_850, x_6, x_7, x_8, x_9, x_851); +x_852 = l_Lean_Compiler_LCNF_FunDecl_etaExpand(x_850, x_6, x_7, x_8, x_9, x_851); if (lean_obj_tag(x_852) == 0) { lean_object* x_853; lean_object* x_854; lean_object* x_855; uint8_t x_856; lean_object* x_857; lean_object* x_858; @@ -8964,7 +8953,7 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_1004 = l_Lean_Compiler_LCNF_FunDeclCore_etaExpand(x_1002, x_6, x_7, x_8, x_9, x_1003); +x_1004 = l_Lean_Compiler_LCNF_FunDecl_etaExpand(x_1002, x_6, x_7, x_8, x_9, x_1003); if (lean_obj_tag(x_1004) == 0) { lean_object* x_1005; lean_object* x_1006; lean_object* x_1007; uint8_t x_1008; lean_object* x_1009; lean_object* x_1010; @@ -10152,7 +10141,7 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_1220 = l_Lean_Compiler_LCNF_FunDeclCore_etaExpand(x_1218, x_6, x_7, x_8, x_9, x_1219); +x_1220 = l_Lean_Compiler_LCNF_FunDecl_etaExpand(x_1218, x_6, x_7, x_8, x_9, x_1219); if (lean_obj_tag(x_1220) == 0) { lean_object* x_1221; lean_object* x_1222; lean_object* x_1223; uint8_t x_1224; lean_object* x_1225; lean_object* x_1226; @@ -11018,7 +11007,7 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_1372 = l_Lean_Compiler_LCNF_FunDeclCore_etaExpand(x_1370, x_6, x_7, x_8, x_9, x_1371); +x_1372 = l_Lean_Compiler_LCNF_FunDecl_etaExpand(x_1370, x_6, x_7, x_8, x_9, x_1371); if (lean_obj_tag(x_1372) == 0) { lean_object* x_1373; lean_object* x_1374; lean_object* x_1375; uint8_t x_1376; lean_object* x_1377; lean_object* x_1378; @@ -12204,7 +12193,7 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_1583 = l_Lean_Compiler_LCNF_FunDeclCore_etaExpand(x_1581, x_6, x_7, x_8, x_9, x_1582); +x_1583 = l_Lean_Compiler_LCNF_FunDecl_etaExpand(x_1581, x_6, x_7, x_8, x_9, x_1582); if (lean_obj_tag(x_1583) == 0) { lean_object* x_1584; lean_object* x_1585; lean_object* x_1586; uint8_t x_1587; lean_object* x_1588; lean_object* x_1589; @@ -13428,7 +13417,7 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_1791 = l_Lean_Compiler_LCNF_FunDeclCore_etaExpand(x_1789, x_6, x_7, x_8, x_9, x_1790); +x_1791 = l_Lean_Compiler_LCNF_FunDecl_etaExpand(x_1789, x_6, x_7, x_8, x_9, x_1790); if (lean_obj_tag(x_1791) == 0) { lean_object* x_1792; lean_object* x_1793; lean_object* x_1794; uint8_t x_1795; lean_object* x_1796; lean_object* x_1797; @@ -14307,7 +14296,7 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_1943 = l_Lean_Compiler_LCNF_FunDeclCore_etaExpand(x_1941, x_6, x_7, x_8, x_9, x_1942); +x_1943 = l_Lean_Compiler_LCNF_FunDecl_etaExpand(x_1941, x_6, x_7, x_8, x_9, x_1942); if (lean_obj_tag(x_1943) == 0) { lean_object* x_1944; lean_object* x_1945; lean_object* x_1946; uint8_t x_1947; lean_object* x_1948; lean_object* x_1949; @@ -17751,7 +17740,7 @@ lean_inc(x_48); x_49 = lean_ctor_get(x_47, 1); lean_inc(x_49); lean_dec(x_47); -x_50 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetDecl____x40_Lean_Compiler_LCNF_Basic___hyg_1722_(x_44, x_48); +x_50 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetDecl____x40_Lean_Compiler_LCNF_Basic___hyg_1659_(x_44, x_48); if (x_50 == 0) { lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; @@ -17860,7 +17849,7 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_82 = l_Lean_Compiler_LCNF_FunDeclCore_etaExpand(x_80, x_5, x_6, x_7, x_8, x_81); +x_82 = l_Lean_Compiler_LCNF_FunDecl_etaExpand(x_80, x_5, x_6, x_7, x_8, x_81); 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; lean_object* x_88; @@ -18100,7 +18089,7 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_138 = l_Lean_Compiler_LCNF_FunDeclCore_etaExpand(x_136, x_5, x_6, x_7, x_8, x_137); +x_138 = l_Lean_Compiler_LCNF_FunDecl_etaExpand(x_136, x_5, x_6, x_7, x_8, x_137); if (lean_obj_tag(x_138) == 0) { lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; @@ -18685,7 +18674,7 @@ goto block_287; else { lean_object* x_291; lean_object* x_292; lean_object* x_293; -x_291 = l___private_Lean_Compiler_LCNF_Simp_Main_0__Lean_Compiler_LCNF_Simp_oneExitPointQuick_go___closed__1; +x_291 = l_Lean_Compiler_LCNF_instInhabitedAlt; x_292 = lean_unsigned_to_nat(0u); x_293 = lean_array_get(x_291, x_254, x_292); if (lean_obj_tag(x_293) == 0) @@ -18710,7 +18699,7 @@ lean_dec(x_231); lean_dec(x_230); lean_dec(x_229); lean_dec(x_1); -x_295 = l_Lean_Compiler_LCNF_AltCore_getCode(x_293); +x_295 = l_Lean_Compiler_LCNF_Alt_getCode(x_293); lean_dec(x_293); lean_ctor_set(x_249, 1, x_255); lean_ctor_set(x_249, 0, x_295); @@ -18989,7 +18978,7 @@ goto block_327; else { lean_object* x_331; lean_object* x_332; lean_object* x_333; -x_331 = l___private_Lean_Compiler_LCNF_Simp_Main_0__Lean_Compiler_LCNF_Simp_oneExitPointQuick_go___closed__1; +x_331 = l_Lean_Compiler_LCNF_instInhabitedAlt; x_332 = lean_unsigned_to_nat(0u); x_333 = lean_array_get(x_331, x_303, x_332); if (lean_obj_tag(x_333) == 0) @@ -19013,7 +19002,7 @@ lean_dec(x_231); lean_dec(x_230); lean_dec(x_229); lean_dec(x_1); -x_335 = l_Lean_Compiler_LCNF_AltCore_getCode(x_333); +x_335 = l_Lean_Compiler_LCNF_Alt_getCode(x_333); lean_dec(x_333); x_336 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_336, 0, x_335); @@ -19619,7 +19608,7 @@ lean_inc(x_414); x_415 = lean_ctor_get(x_413, 1); lean_inc(x_415); lean_dec(x_413); -x_416 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetDecl____x40_Lean_Compiler_LCNF_Basic___hyg_1722_(x_410, x_414); +x_416 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetDecl____x40_Lean_Compiler_LCNF_Basic___hyg_1659_(x_410, x_414); if (x_416 == 0) { lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; @@ -19728,7 +19717,7 @@ lean_inc(x_8); lean_inc(x_407); lean_inc(x_6); lean_inc(x_5); -x_448 = l_Lean_Compiler_LCNF_FunDeclCore_etaExpand(x_446, x_5, x_6, x_407, x_8, x_447); +x_448 = l_Lean_Compiler_LCNF_FunDecl_etaExpand(x_446, x_5, x_6, x_407, x_8, x_447); if (lean_obj_tag(x_448) == 0) { lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; @@ -19974,7 +19963,7 @@ lean_inc(x_8); lean_inc(x_407); lean_inc(x_6); lean_inc(x_5); -x_504 = l_Lean_Compiler_LCNF_FunDeclCore_etaExpand(x_502, x_5, x_6, x_407, x_8, x_503); +x_504 = l_Lean_Compiler_LCNF_FunDecl_etaExpand(x_502, x_5, x_6, x_407, x_8, x_503); if (lean_obj_tag(x_504) == 0) { lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; @@ -20550,7 +20539,7 @@ goto block_638; else { lean_object* x_642; lean_object* x_643; lean_object* x_644; -x_642 = l___private_Lean_Compiler_LCNF_Simp_Main_0__Lean_Compiler_LCNF_Simp_oneExitPointQuick_go___closed__1; +x_642 = l_Lean_Compiler_LCNF_instInhabitedAlt; x_643 = lean_unsigned_to_nat(0u); x_644 = lean_array_get(x_642, x_614, x_643); if (lean_obj_tag(x_644) == 0) @@ -20575,7 +20564,7 @@ lean_dec(x_591); lean_dec(x_590); lean_dec(x_589); lean_dec(x_1); -x_646 = l_Lean_Compiler_LCNF_AltCore_getCode(x_644); +x_646 = l_Lean_Compiler_LCNF_Alt_getCode(x_644); lean_dec(x_644); if (lean_is_scalar(x_612)) { x_647 = lean_alloc_ctor(0, 2, 0); @@ -21845,7 +21834,7 @@ if (x_28 == 0) lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; x_29 = lean_ctor_get(x_20, 0); x_30 = l_Lean_Compiler_LCNF_Simp_CtorInfo_getName(x_29); -x_31 = l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21(x_1, x_30); +x_31 = l_Lean_Compiler_LCNF_Cases_extractAlt_x21(x_1, x_30); x_32 = lean_ctor_get(x_31, 0); lean_inc(x_32); x_33 = lean_ctor_get(x_31, 1); @@ -22820,7 +22809,7 @@ x_221 = lean_ctor_get(x_20, 0); lean_inc(x_221); lean_dec(x_20); x_222 = l_Lean_Compiler_LCNF_Simp_CtorInfo_getName(x_221); -x_223 = l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21(x_1, x_222); +x_223 = l_Lean_Compiler_LCNF_Cases_extractAlt_x21(x_1, x_222); x_224 = lean_ctor_get(x_223, 0); lean_inc(x_224); x_225 = lean_ctor_get(x_223, 1); @@ -23339,7 +23328,7 @@ if (lean_is_exclusive(x_321)) { x_328 = lean_box(0); } x_329 = l_Lean_Compiler_LCNF_Simp_CtorInfo_getName(x_327); -x_330 = l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21(x_1, x_329); +x_330 = l_Lean_Compiler_LCNF_Cases_extractAlt_x21(x_1, x_329); x_331 = lean_ctor_get(x_330, 0); lean_inc(x_331); x_332 = lean_ctor_get(x_330, 1); @@ -24402,8 +24391,6 @@ lean_dec_ref(res); res = initialize_Lean_Compiler_LCNF_Simp_ConstantFold(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___private_Lean_Compiler_LCNF_Simp_Main_0__Lean_Compiler_LCNF_Simp_oneExitPointQuick_go___closed__1 = _init_l___private_Lean_Compiler_LCNF_Simp_Main_0__Lean_Compiler_LCNF_Simp_oneExitPointQuick_go___closed__1(); -lean_mark_persistent(l___private_Lean_Compiler_LCNF_Simp_Main_0__Lean_Compiler_LCNF_Simp_oneExitPointQuick_go___closed__1); l_Lean_Compiler_LCNF_Simp_specializePartialApp___closed__1 = _init_l_Lean_Compiler_LCNF_Simp_specializePartialApp___closed__1(); lean_mark_persistent(l_Lean_Compiler_LCNF_Simp_specializePartialApp___closed__1); l_Lean_Compiler_LCNF_Simp_specializePartialApp___closed__2 = _init_l_Lean_Compiler_LCNF_Simp_specializePartialApp___closed__2(); diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Simp/Used.c b/stage0/stdlib/Lean/Compiler/LCNF/Simp/Used.c index 5b703169abd1..1b0b4e98b4ba 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Simp/Used.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Simp/Used.c @@ -33,11 +33,11 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_markUsedFunDecl(lean_object*, lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_markUsedLetDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Simp_markUsedCode___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_Compiler_LCNF_AltCore_getCode(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_markUsedFVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_markUsedType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_markUsedFunDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_markUsedFVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Compiler_LCNF_Alt_getCode(lean_object*); extern lean_object* l_Lean_Compiler_LCNF_instInhabitedCodeDecl; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_markUsedLetValue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_markUsedCode___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -993,7 +993,7 @@ if (x_13 == 0) lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; size_t x_19; size_t x_20; lean_dec(x_4); x_14 = lean_array_uget(x_1, x_2); -x_15 = l_Lean_Compiler_LCNF_AltCore_getCode(x_14); +x_15 = l_Lean_Compiler_LCNF_Alt_getCode(x_14); lean_dec(x_14); x_16 = l_Lean_Compiler_LCNF_Simp_markUsedCode(x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); x_17 = lean_ctor_get(x_16, 0); diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Specialize.c b/stage0/stdlib/Lean/Compiler/LCNF/Specialize.c index 0c3d31b01a74..86758b686d33 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Specialize.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Specialize.c @@ -88,7 +88,6 @@ lean_object* l_Lean_Compiler_LCNF_Code_instantiateValueLevelParams_instCode(lean LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_mkSpecDecl___boxed(lean_object*, lean_object*, lean_object*, 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_mkStr5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Nat_nextPowerOfTwo_go(lean_object*, lean_object*, lean_object*); -lean_object* l_ReaderT_bind___at_Lean_Compiler_LCNF_FunDeclCore_toExprM___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); static lean_object* l_Lean_addTrace___at_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___spec__6___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -140,6 +139,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_instInhabitedCacheEntry static lean_object* l_Lean_Compiler_LCNF_Expr_forFVarM___at_Lean_Compiler_LCNF_Specialize_withLetDecl___spec__5___closed__5; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Arg_forFVarM___at_Lean_Compiler_LCNF_Specialize_withLetDecl___spec__4(lean_object*, lean_object*); lean_object* lean_nat_div(lean_object*, lean_object*); +lean_object* l_Lean_Compiler_LCNF_FunDecl_toExpr(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_specialize___elambda__1___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_49____lambda__1(lean_object*); lean_object* l_Lean_Name_quickCmp___boxed(lean_object*, lean_object*); @@ -249,6 +249,7 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_Sp static lean_object* l_Lean_Compiler_LCNF_Specialize_cacheSpec___closed__1; static lean_object* l_Lean_Compiler_LCNF_Specialize_mkSpecDecl_go___closed__2; static lean_object* l_Lean_Compiler_LCNF_Specialize_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_49____closed__6; +lean_object* l_ReaderT_bind___at_Lean_Compiler_LCNF_FunDecl_toExprM___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_LCNF_Specialize_withLetDecl___spec__6(lean_object*); static lean_object* l_Lean_Compiler_LCNF_Specialize_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_49____closed__11; @@ -354,7 +355,6 @@ static lean_object* l_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___lambda__ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_instMonadScopeSpecializeM___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_nat_add(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___lambda__5___closed__3; -lean_object* l_Lean_Compiler_LCNF_FunDeclCore_toExpr(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_DeclValue_mapCodeM___at_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_allFVarM___at_Lean_Compiler_LCNF_Specialize_withLetDecl___spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___lambda__3___closed__13; @@ -3510,7 +3510,7 @@ x_15 = lean_ctor_get(x_5, 0); lean_inc(x_15); lean_dec(x_5); x_16 = l_Lean_Compiler_LCNF_Specialize_Collector_collect___closed__1; -x_17 = l_Lean_Compiler_LCNF_FunDeclCore_toExpr(x_15, x_16); +x_17 = l_Lean_Compiler_LCNF_FunDecl_toExpr(x_15, x_16); x_18 = lean_array_uset(x_7, x_2, x_17); x_2 = x_9; x_3 = x_18; @@ -3594,7 +3594,7 @@ lean_object* x_11; lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_11 = l_ReaderT_bind___at_Lean_Compiler_LCNF_FunDeclCore_toExprM___spec__1___rarg(x_7, x_8, x_5, x_6); +x_11 = l_ReaderT_bind___at_Lean_Compiler_LCNF_FunDecl_toExprM___spec__1___rarg(x_7, x_8, x_5, x_6); return x_11; } else diff --git a/stage0/stdlib/Lean/Compiler/LCNF/StructProjCases.c b/stage0/stdlib/Lean/Compiler/LCNF/StructProjCases.c new file mode 100644 index 000000000000..7ce56947b185 --- /dev/null +++ b/stage0/stdlib/Lean/Compiler/LCNF/StructProjCases.c @@ -0,0 +1,7110 @@ +// Lean compiler output +// Module: Lean.Compiler.LCNF.StructProjCases +// Imports: Lean.Compiler.LCNF.Basic Lean.Compiler.LCNF.InferType Lean.Compiler.LCNF.MonoTypes Lean.Compiler.LCNF.PassManager Lean.Compiler.LCNF.PrettyPrinter +#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_Compiler_LCNF_StructProjCases_visitDecl___closed__1; +lean_object* l_Lean_Compiler_LCNF_LetDecl_updateValue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_panic___at_Lean_Compiler_LCNF_StructProjCases_visitLetValue___spec__1___closed__2; +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_StructProjCases_visitDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Compiler_LCNF_eraseLetDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__3___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_structProjCases___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_StructProjCases_visitLetValue___closed__3; +extern lean_object* l_Lean_Compiler_LCNF_instInhabitedLetValue; +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_StructProjCases_visitArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Compiler_LCNF_Code_inferType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +size_t lean_uint64_to_usize(uint64_t); +static lean_object* l_panic___at_Lean_Compiler_LCNF_StructProjCases_visitLetValue___spec__1___closed__1; +static lean_object* l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__3; +lean_object* lean_array_push(lean_object*, lean_object*); +lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_Compiler_LCNF_instInhabitedCode; +static lean_object* l_Lean_Compiler_LCNF_structProjCases___closed__2; +uint8_t lean_usize_dec_eq(size_t, size_t); +lean_object* lean_mk_array(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__13(lean_object*, lean_object*); +lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*, uint8_t); +lean_object* l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_updateFunDeclImp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_StructProjCases_visitAlt(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_StructProjCases_mkFieldParamsForCtorType_loop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__4; +static lean_object* l_panic___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__2___closed__1; +lean_object* l_Nat_nextPowerOfTwo_go(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_structProjCases___closed__3; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__3(size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_StructProjCases_remapFVar___boxed(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_Compiler_LCNF_StructProjCases_remapFVar___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_StructProjCases_findStructCtorInfo_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__8(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__6(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__15(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +size_t lean_ptr_addr(lean_object*); +uint64_t l___private_Lean_Expr_0__Lean_hashFVarId____x40_Lean_Expr___hyg_1730_(lean_object*); +size_t lean_usize_of_nat(lean_object*); +lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_LetValue_updateFVarImp(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_DeclValue_mapCodeM___at_Lean_Compiler_LCNF_StructProjCases_visitDecl___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_StructProjCases_mkFieldParamsForCtorType_loop___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_st_ref_take(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_StructProjCases_visitLetValue___closed__1; +lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_LetValue_updateArgsImp(lean_object*, lean_object*); +uint64_t lean_uint64_shift_right(uint64_t, uint64_t); +lean_object* lean_nat_div(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__1___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__9; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__9(lean_object*, lean_object*); +lean_object* lean_st_ref_get(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__2; +lean_object* lean_st_mk_ref(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_StructProjCases_findStructCtorInfo_x3f(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_StructProjCases_M_run___rarg___closed__4; +static lean_object* l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__10; +static lean_object* l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__1; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__4___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__16(lean_object*, uint64_t, uint64_t, size_t, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_StructProjCases_mkFieldParamsForCtorType(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_Compiler_LCNF_StructProjCases_visitCode___spec__7(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_StructProjCases_remapFVar(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*); +static lean_object* l_Lean_Compiler_LCNF_StructProjCases_M_run___rarg___closed__3; +static lean_object* l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__8; +lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_StructProjCases_visitLetValue___spec__2(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_instInhabitedFVarId; +static lean_object* l_Lean_Compiler_LCNF_StructProjCases_M_run___rarg___closed__2; +extern lean_object* l_Lean_Compiler_LCNF_instMonadCompilerM; +LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_LCNF_StructProjCases_visitLetValue___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Compiler_LCNF_Alt_getCode(lean_object*); +LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__4(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__6; +LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_structProjCases___closed__4; +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_StructProjCases_mkFieldParamsForCtorType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_fget(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_erase___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__5(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__11(lean_object*); +lean_object* l_Lean_Compiler_LCNF_toMonoType(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__12(lean_object*, lean_object*, lean_object*); +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +lean_object* l_Lean_Compiler_LCNF_mkParam(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__5; +static lean_object* l_Lean_Compiler_LCNF_StructProjCases_mkFieldParamsForCtorType___closed__1; +static lean_object* l_Lean_Compiler_LCNF_structProjCases___closed__1; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__10(lean_object*, lean_object*, lean_object*); +uint64_t lean_uint64_xor(uint64_t, uint64_t); +lean_object* lean_panic_fn(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__14(lean_object*, lean_object*, lean_object*); +lean_object* lean_nat_sub(lean_object*, lean_object*); +lean_object* l_Lean_Compiler_LCNF_eraseParam(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_nat_mul(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__6___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__7; +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_StructProjCases_M_run(lean_object*); +lean_object* l_Lean_getConstInfo___at___private_Lean_Compiler_LCNF_Util_0__Lean_Compiler_LCNF_getCasesOnInductiveVal_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_StructProjCases_visitCode(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*); +lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltCodeImp(lean_object*, lean_object*); +size_t lean_usize_add(size_t, size_t); +lean_object* l_Lean_Compiler_LCNF_Pass_mkPerDeclaration(lean_object*, lean_object*, uint8_t, lean_object*); +lean_object* lean_array_uget(lean_object*, size_t); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Compiler_LCNF_StructProjCases_remapFVar___spec__1___boxed(lean_object*, lean_object*); +size_t lean_array_size(lean_object*); +static lean_object* l_Lean_Compiler_LCNF_StructProjCases_visitLetValue___closed__4; +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_StructProjCases_M_run___rarg(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_Compiler_LCNF_StructProjCases_visitLetValue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_erase___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__5___boxed(lean_object*, lean_object*); +lean_object* lean_string_append(lean_object*, lean_object*); +lean_object* lean_array_get_size(lean_object*); +lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_nat_dec_le(lean_object*, lean_object*); +uint8_t lean_usize_dec_lt(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_structProjCases; +lean_object* lean_nat_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_StructProjCases_visitArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_StructProjCases_M_run___rarg___closed__1; +lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_StructProjCases_visitLetValue___closed__2; +lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateFVarImp(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_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_StructProjCases_visitLetValue___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_StructProjCases_findStructCtorInfo_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Lean_getConstInfo___at___private_Lean_Compiler_LCNF_Util_0__Lean_Compiler_LCNF_getCasesOnInductiveVal_x3f___spec__1(x_1, x_2, x_3, x_4); +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_6; +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +if (lean_obj_tag(x_6) == 5) +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +lean_dec(x_6); +x_8 = lean_ctor_get(x_7, 4); +lean_inc(x_8); +lean_dec(x_7); +if (lean_obj_tag(x_8) == 0) +{ +uint8_t x_9; +x_9 = !lean_is_exclusive(x_5); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_5, 0); +lean_dec(x_10); +x_11 = lean_box(0); +lean_ctor_set(x_5, 0, x_11); +return x_5; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_5, 1); +lean_inc(x_12); +lean_dec(x_5); +x_13 = lean_box(0); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_12); +return x_14; +} +} +else +{ +lean_object* x_15; +x_15 = lean_ctor_get(x_8, 1); +lean_inc(x_15); +if (lean_obj_tag(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_5, 1); +lean_inc(x_16); +lean_dec(x_5); +x_17 = lean_ctor_get(x_8, 0); +lean_inc(x_17); +lean_dec(x_8); +x_18 = lean_st_ref_get(x_3, x_16); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; +x_20 = lean_ctor_get(x_18, 0); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +lean_dec(x_20); +x_22 = 0; +x_23 = l_Lean_Environment_find_x3f(x_21, x_17, x_22); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; +x_24 = lean_box(0); +lean_ctor_set(x_18, 0, x_24); +return x_18; +} +else +{ +uint8_t x_25; +x_25 = !lean_is_exclusive(x_23); +if (x_25 == 0) +{ +lean_object* x_26; +x_26 = lean_ctor_get(x_23, 0); +if (lean_obj_tag(x_26) == 6) +{ +lean_object* x_27; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +lean_dec(x_26); +lean_ctor_set(x_23, 0, x_27); +lean_ctor_set(x_18, 0, x_23); +return x_18; +} +else +{ +lean_object* x_28; +lean_free_object(x_23); +lean_dec(x_26); +x_28 = lean_box(0); +lean_ctor_set(x_18, 0, x_28); +return x_18; +} +} +else +{ +lean_object* x_29; +x_29 = lean_ctor_get(x_23, 0); +lean_inc(x_29); +lean_dec(x_23); +if (lean_obj_tag(x_29) == 6) +{ +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +lean_dec(x_29); +x_31 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_18, 0, x_31); +return x_18; +} +else +{ +lean_object* x_32; +lean_dec(x_29); +x_32 = lean_box(0); +lean_ctor_set(x_18, 0, x_32); +return x_18; +} +} +} +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; lean_object* x_37; +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_ctor_get(x_33, 0); +lean_inc(x_35); +lean_dec(x_33); +x_36 = 0; +x_37 = l_Lean_Environment_find_x3f(x_35, x_17, x_36); +if (lean_obj_tag(x_37) == 0) +{ +lean_object* x_38; lean_object* x_39; +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_34); +return x_39; +} +else +{ +lean_object* x_40; lean_object* x_41; +x_40 = lean_ctor_get(x_37, 0); +lean_inc(x_40); +if (lean_is_exclusive(x_37)) { + lean_ctor_release(x_37, 0); + x_41 = x_37; +} else { + lean_dec_ref(x_37); + x_41 = lean_box(0); +} +if (lean_obj_tag(x_40) == 6) +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_40, 0); +lean_inc(x_42); +lean_dec(x_40); +if (lean_is_scalar(x_41)) { + x_43 = lean_alloc_ctor(1, 1, 0); +} else { + x_43 = x_41; +} +lean_ctor_set(x_43, 0, 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_34); +return x_44; +} +else +{ +lean_object* x_45; lean_object* x_46; +lean_dec(x_41); +lean_dec(x_40); +x_45 = lean_box(0); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_34); +return x_46; +} +} +} +} +else +{ +uint8_t x_47; +lean_dec(x_15); +lean_dec(x_8); +x_47 = !lean_is_exclusive(x_5); +if (x_47 == 0) +{ +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_5, 0); +lean_dec(x_48); +x_49 = lean_box(0); +lean_ctor_set(x_5, 0, x_49); +return x_5; +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_5, 1); +lean_inc(x_50); +lean_dec(x_5); +x_51 = lean_box(0); +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_50); +return x_52; +} +} +} +} +else +{ +uint8_t x_53; +lean_dec(x_6); +x_53 = !lean_is_exclusive(x_5); +if (x_53 == 0) +{ +lean_object* x_54; lean_object* x_55; +x_54 = lean_ctor_get(x_5, 0); +lean_dec(x_54); +x_55 = lean_box(0); +lean_ctor_set(x_5, 0, x_55); +return x_5; +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_5, 1); +lean_inc(x_56); +lean_dec(x_5); +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; +} +} +} +else +{ +uint8_t x_59; +x_59 = !lean_is_exclusive(x_5); +if (x_59 == 0) +{ +return x_5; +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_5, 0); +x_61 = lean_ctor_get(x_5, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_5); +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_Compiler_LCNF_StructProjCases_findStructCtorInfo_x3f___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_Compiler_LCNF_StructProjCases_findStructCtorInfo_x3f(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_Compiler_LCNF_StructProjCases_mkFieldParamsForCtorType_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) { +_start: +{ +if (lean_obj_tag(x_2) == 7) +{ +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_2, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_2, 1); +lean_inc(x_10); +x_11 = lean_ctor_get(x_2, 2); +lean_inc(x_11); +lean_dec(x_2); +x_12 = lean_unsigned_to_nat(0u); +x_13 = lean_nat_dec_eq(x_3, x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +lean_dec(x_10); +lean_dec(x_9); +x_14 = lean_unsigned_to_nat(1u); +x_15 = lean_nat_sub(x_3, x_14); +lean_dec(x_3); +x_2 = x_11; +x_3 = x_15; +goto _start; +} +else +{ +lean_object* x_17; +lean_inc(x_7); +lean_inc(x_6); +x_17 = l_Lean_Compiler_LCNF_toMonoType(x_10, x_6, x_7, x_8); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_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 = 0; +x_21 = l_Lean_Compiler_LCNF_mkParam(x_9, x_18, x_20, x_4, x_5, x_6, x_7, x_19); +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = lean_array_push(x_1, x_22); +x_1 = x_24; +x_2 = x_11; +x_8 = x_23; +goto _start; +} +else +{ +uint8_t x_26; +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_1); +x_26 = !lean_is_exclusive(x_17); +if (x_26 == 0) +{ +return x_17; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_17, 0); +x_28 = lean_ctor_get(x_17, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_17); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; +} +} +} +} +else +{ +lean_object* x_30; +lean_dec(x_7); +lean_dec(x_6); +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_8); +return x_30; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_StructProjCases_mkFieldParamsForCtorType_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) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_Compiler_LCNF_StructProjCases_mkFieldParamsForCtorType_loop(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_5); +lean_dec(x_4); +return x_9; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_StructProjCases_mkFieldParamsForCtorType___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_Compiler_LCNF_StructProjCases_mkFieldParamsForCtorType(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; +x_8 = l_Lean_Compiler_LCNF_StructProjCases_mkFieldParamsForCtorType___closed__1; +x_9 = l_Lean_Compiler_LCNF_StructProjCases_mkFieldParamsForCtorType_loop(x_8, x_1, x_2, x_3, x_4, x_5, x_6, x_7); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_StructProjCases_mkFieldParamsForCtorType___boxed(lean_object* x_1, lean_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_Compiler_LCNF_StructProjCases_mkFieldParamsForCtorType(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_4); +lean_dec(x_3); +return x_8; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_StructProjCases_M_run___rarg___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_Compiler_LCNF_StructProjCases_M_run___rarg___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Compiler_LCNF_StructProjCases_M_run___rarg___closed__1; +x_3 = lean_mk_array(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_StructProjCases_M_run___rarg___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_Compiler_LCNF_StructProjCases_M_run___rarg___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_Compiler_LCNF_StructProjCases_M_run___rarg___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Compiler_LCNF_StructProjCases_M_run___rarg___closed__3; +x_2 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_StructProjCases_M_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) { +_start: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_7 = l_Lean_Compiler_LCNF_StructProjCases_M_run___rarg___closed__4; +x_8 = lean_st_mk_ref(x_7, x_6); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +lean_dec(x_8); +lean_inc(x_9); +x_11 = lean_apply_6(x_1, x_9, x_2, x_3, x_4, x_5, 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_9, x_13); +lean_dec(x_9); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) +{ +lean_object* x_16; +x_16 = lean_ctor_get(x_14, 0); +lean_dec(x_16); +lean_ctor_set(x_14, 0, x_12); +return x_14; +} +else +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_14, 1); +lean_inc(x_17); +lean_dec(x_14); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_12); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} +} +else +{ +uint8_t x_19; +lean_dec(x_9); +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_Compiler_LCNF_StructProjCases_M_run(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_StructProjCases_M_run___rarg), 6, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Compiler_LCNF_StructProjCases_remapFVar___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_Compiler_LCNF_StructProjCases_remapFVar(lean_object* x_1, lean_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; 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; lean_object* x_28; +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 = l___private_Lean_Expr_0__Lean_hashFVarId____x40_Lean_Expr___hyg_1730_(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_get_x3f___at_Lean_Compiler_LCNF_StructProjCases_remapFVar___spec__1(x_1, x_27); +lean_dec(x_27); +if (lean_obj_tag(x_28) == 0) +{ +lean_ctor_set(x_8, 0, x_1); +return x_8; +} +else +{ +lean_object* x_29; +lean_dec(x_1); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +lean_dec(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; lean_object* x_46; +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 = l___private_Lean_Expr_0__Lean_hashFVarId____x40_Lean_Expr___hyg_1730_(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_get_x3f___at_Lean_Compiler_LCNF_StructProjCases_remapFVar___spec__1(x_1, x_45); +lean_dec(x_45); +if (lean_obj_tag(x_46) == 0) +{ +lean_object* x_47; +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_1); +lean_ctor_set(x_47, 1, x_30); +return x_47; +} +else +{ +lean_object* x_48; lean_object* x_49; +lean_dec(x_1); +x_48 = lean_ctor_get(x_46, 0); +lean_inc(x_48); +lean_dec(x_46); +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_30); +return x_49; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Compiler_LCNF_StructProjCases_remapFVar___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_Compiler_LCNF_StructProjCases_remapFVar___spec__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_StructProjCases_remapFVar___boxed(lean_object* x_1, lean_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_Compiler_LCNF_StructProjCases_remapFVar(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_Compiler_LCNF_StructProjCases_visitArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +if (lean_obj_tag(x_1) == 1) +{ +lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); +x_9 = l_Lean_Compiler_LCNF_StructProjCases_remapFVar(x_8, x_2, 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); +x_12 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateFVarImp(x_1, x_11); +lean_ctor_set(x_9, 0, x_12); +return x_9; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_9, 0); +x_14 = lean_ctor_get(x_9, 1); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_9); +x_15 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateFVarImp(x_1, x_13); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_14); +return x_16; +} +} +else +{ +lean_object* x_17; +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_1); +lean_ctor_set(x_17, 1, x_7); +return x_17; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_StructProjCases_visitArg___boxed(lean_object* x_1, lean_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_Compiler_LCNF_StructProjCases_visitArg(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; +} +} +static lean_object* _init_l_panic___at_Lean_Compiler_LCNF_StructProjCases_visitLetValue___spec__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Compiler_LCNF_instMonadCompilerM; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at_Lean_Compiler_LCNF_StructProjCases_visitLetValue___spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_panic___at_Lean_Compiler_LCNF_StructProjCases_visitLetValue___spec__1___closed__1; +x_2 = l_Lean_Compiler_LCNF_instInhabitedLetValue; +x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_LCNF_StructProjCases_visitLetValue___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_Compiler_LCNF_StructProjCases_visitLetValue___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_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_StructProjCases_visitLetValue___spec__2(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +uint8_t x_10; +x_10 = lean_usize_dec_lt(x_2, 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_3); +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; size_t x_18; size_t x_19; lean_object* x_20; +x_12 = lean_array_uget(x_3, x_2); +x_13 = lean_unsigned_to_nat(0u); +x_14 = lean_array_uset(x_3, x_2, x_13); +x_15 = l_Lean_Compiler_LCNF_StructProjCases_visitArg(x_12, 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 = 1; +x_19 = lean_usize_add(x_2, x_18); +x_20 = lean_array_uset(x_14, x_2, x_16); +x_2 = x_19; +x_3 = x_20; +x_9 = x_17; +goto _start; +} +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_StructProjCases_visitLetValue___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Compiler.LCNF.StructProjCases", 34, 34); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_StructProjCases_visitLetValue___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Compiler.LCNF.StructProjCases.visitLetValue", 48, 48); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_StructProjCases_visitLetValue___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_Compiler_LCNF_StructProjCases_visitLetValue___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_Compiler_LCNF_StructProjCases_visitLetValue___closed__1; +x_2 = l_Lean_Compiler_LCNF_StructProjCases_visitLetValue___closed__2; +x_3 = lean_unsigned_to_nat(111u); +x_4 = lean_unsigned_to_nat(16u); +x_5 = l_Lean_Compiler_LCNF_StructProjCases_visitLetValue___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_Compiler_LCNF_StructProjCases_visitLetValue(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 2: +{ +lean_object* x_8; lean_object* x_9; +lean_dec(x_1); +x_8 = l_Lean_Compiler_LCNF_StructProjCases_visitLetValue___closed__4; +x_9 = l_panic___at_Lean_Compiler_LCNF_StructProjCases_visitLetValue___spec__1(x_8, x_2, x_3, x_4, x_5, x_6, x_7); +return x_9; +} +case 3: +{ +lean_object* x_10; size_t x_11; size_t x_12; lean_object* x_13; uint8_t x_14; +x_10 = lean_ctor_get(x_1, 2); +lean_inc(x_10); +x_11 = lean_array_size(x_10); +x_12 = 0; +x_13 = l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_StructProjCases_visitLetValue___spec__2(x_11, x_12, x_10, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_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___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_LetValue_updateArgsImp(x_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 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_LetValue_updateArgsImp(x_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; +} +} +case 4: +{ +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; uint8_t x_29; +x_21 = lean_ctor_get(x_1, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_1, 1); +lean_inc(x_22); +x_23 = l_Lean_Compiler_LCNF_StructProjCases_remapFVar(x_21, x_2, x_3, x_4, x_5, x_6, x_7); +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_size(x_22); +x_27 = 0; +x_28 = l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_StructProjCases_visitLetValue___spec__2(x_26, x_27, x_22, x_2, x_3, x_4, x_5, x_6, x_25); +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_28); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_28, 0); +x_31 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_LetValue_updateFVarImp(x_1, x_24, x_30); +lean_ctor_set(x_28, 0, x_31); +return x_28; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_32 = lean_ctor_get(x_28, 0); +x_33 = lean_ctor_get(x_28, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_28); +x_34 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_LetValue_updateFVarImp(x_1, x_24, x_32); +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; +} +} +default: +{ +lean_object* x_36; +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(0, 2, 0); +lean_ctor_set(x_36, 0, x_1); +lean_ctor_set(x_36, 1, x_7); +return x_36; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_StructProjCases_visitLetValue___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +size_t x_10; size_t x_11; lean_object* x_12; +x_10 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_11 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_12 = l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_StructProjCases_visitLetValue___spec__2(x_10, x_11, 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_12; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Compiler_LCNF_StructProjCases_visitCode___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; +} +} +} +} +static lean_object* _init_l_panic___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__2___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_panic___at_Lean_Compiler_LCNF_StructProjCases_visitLetValue___spec__1___closed__1; +x_2 = l_Lean_Compiler_LCNF_instInhabitedCode; +x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_LCNF_StructProjCases_visitCode___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) { +_start: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = l_panic___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__2___closed__1; +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_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_StructProjCases_visitCode___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, 0); +lean_inc(x_8); +lean_dec(x_5); +x_9 = 1; +x_10 = lean_usize_add(x_2, x_9); +x_11 = lean_array_uset(x_7, x_2, x_8); +x_2 = x_10; +x_3 = x_11; +goto _start; +} +} +} +LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__4(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_erase___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__5(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 +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_2); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_5 = lean_ctor_get(x_2, 0); +x_6 = lean_ctor_get(x_2, 1); +x_7 = lean_ctor_get(x_2, 2); +x_8 = lean_name_eq(x_5, x_1); +if (x_8 == 0) +{ +lean_object* x_9; +x_9 = l_Std_DHashMap_Internal_AssocList_erase___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__5(x_1, x_7); +lean_ctor_set(x_2, 2, x_9); +return x_2; +} +else +{ +lean_free_object(x_2); +lean_dec(x_6); +lean_dec(x_5); +return x_7; +} +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_10 = lean_ctor_get(x_2, 0); +x_11 = lean_ctor_get(x_2, 1); +x_12 = lean_ctor_get(x_2, 2); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_dec(x_2); +x_13 = lean_name_eq(x_10, x_1); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = l_Std_DHashMap_Internal_AssocList_erase___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__5(x_1, x_12); +x_15 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_15, 0, x_10); +lean_ctor_set(x_15, 1, x_11); +lean_ctor_set(x_15, 2, x_14); +return x_15; +} +else +{ +lean_dec(x_11); +lean_dec(x_10); +return x_12; +} +} +} +} +} +LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__6(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_Compiler_LCNF_StructProjCases_visitCode___spec__9(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___private_Lean_Expr_0__Lean_hashFVarId____x40_Lean_Expr___hyg_1730_(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___private_Lean_Expr_0__Lean_hashFVarId____x40_Lean_Expr___hyg_1730_(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_Compiler_LCNF_StructProjCases_visitCode___spec__8(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_Compiler_LCNF_StructProjCases_visitCode___spec__9(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_Compiler_LCNF_StructProjCases_visitCode___spec__7(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_Compiler_LCNF_StructProjCases_visitCode___spec__8(x_7, x_1, x_6); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__10(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_Compiler_LCNF_StructProjCases_visitCode___spec__10(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_Compiler_LCNF_StructProjCases_visitCode___spec__10(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_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__13(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___private_Lean_Expr_0__Lean_hashFVarId____x40_Lean_Expr___hyg_1730_(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___private_Lean_Expr_0__Lean_hashFVarId____x40_Lean_Expr___hyg_1730_(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_Compiler_LCNF_StructProjCases_visitCode___spec__12(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_Compiler_LCNF_StructProjCases_visitCode___spec__13(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_Compiler_LCNF_StructProjCases_visitCode___spec__11(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_Compiler_LCNF_StructProjCases_visitCode___spec__12(x_7, x_1, x_6); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__14(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_Compiler_LCNF_StructProjCases_visitCode___spec__14(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_Compiler_LCNF_StructProjCases_visitCode___spec__14(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_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__15(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) { +_start: +{ +uint8_t x_10; +x_10 = lean_usize_dec_lt(x_2, x_1); +if (x_10 == 0) +{ +lean_object* x_11; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_3); +lean_ctor_set(x_11, 1, x_9); +return x_11; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_array_uget(x_3, x_2); +x_13 = lean_unsigned_to_nat(0u); +x_14 = lean_array_uset(x_3, x_2, x_13); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_15 = l_Lean_Compiler_LCNF_StructProjCases_visitAlt(x_12, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; lean_object* x_20; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = 1; +x_19 = lean_usize_add(x_2, x_18); +x_20 = lean_array_uset(x_14, x_2, x_16); +x_2 = x_19; +x_3 = x_20; +x_9 = x_17; +goto _start; +} +else +{ +uint8_t x_22; +lean_dec(x_14); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_22 = !lean_is_exclusive(x_15); +if (x_22 == 0) +{ +return x_15; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_15, 0); +x_24 = lean_ctor_get(x_15, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_15); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__16(lean_object* x_1, uint64_t x_2, uint64_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, size_t x_7, size_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +uint8_t x_16; +x_16 = lean_usize_dec_lt(x_8, x_7); +if (x_16 == 0) +{ +lean_object* x_17; +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_9); +lean_ctor_set(x_17, 1, x_15); +return x_17; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_18 = lean_array_uget(x_6, x_8); +x_33 = lean_ctor_get(x_9, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_9, 1); +lean_inc(x_34); +x_35 = lean_ctor_get(x_9, 2); +lean_inc(x_35); +x_36 = lean_nat_dec_lt(x_34, x_35); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; +lean_dec(x_35); +lean_dec(x_34); +lean_dec(x_33); +lean_dec(x_18); +x_37 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_37, 0, x_9); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_15); +x_19 = x_38; +goto block_32; +} +else +{ +uint8_t x_39; +x_39 = !lean_is_exclusive(x_9); +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_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; +x_40 = lean_ctor_get(x_9, 2); +lean_dec(x_40); +x_41 = lean_ctor_get(x_9, 1); +lean_dec(x_41); +x_42 = lean_ctor_get(x_9, 0); +lean_dec(x_42); +x_43 = lean_array_fget(x_33, x_34); +x_44 = lean_unsigned_to_nat(1u); +x_45 = lean_nat_add(x_34, x_44); +lean_dec(x_34); +lean_ctor_set(x_9, 1, x_45); +x_55 = lean_st_ref_take(x_10, x_15); +x_56 = lean_ctor_get(x_55, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_56, 1); +lean_inc(x_57); +x_58 = lean_ctor_get(x_55, 1); +lean_inc(x_58); +lean_dec(x_55); +x_59 = !lean_is_exclusive(x_56); +if (x_59 == 0) +{ +lean_object* x_60; lean_object* x_61; uint8_t x_62; +x_60 = lean_ctor_get(x_56, 1); +lean_dec(x_60); +x_61 = lean_ctor_get(x_18, 0); +lean_inc(x_61); +x_62 = !lean_is_exclusive(x_57); +if (x_62 == 0) +{ +lean_object* x_63; lean_object* x_64; lean_object* 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; lean_object* x_75; uint8_t x_76; +x_63 = lean_ctor_get(x_57, 0); +x_64 = lean_ctor_get(x_57, 1); +x_65 = lean_array_get_size(x_64); +x_66 = l___private_Lean_Expr_0__Lean_hashFVarId____x40_Lean_Expr___hyg_1730_(x_61); +x_67 = lean_uint64_shift_right(x_66, x_2); +x_68 = lean_uint64_xor(x_66, x_67); +x_69 = lean_uint64_shift_right(x_68, x_3); +x_70 = lean_uint64_xor(x_68, x_69); +x_71 = lean_uint64_to_usize(x_70); +x_72 = lean_usize_of_nat(x_65); +lean_dec(x_65); +x_73 = lean_usize_sub(x_72, x_4); +x_74 = lean_usize_land(x_71, x_73); +x_75 = lean_array_uget(x_64, x_74); +x_76 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__6(x_61, x_75); +if (x_76 == 0) +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; uint8_t x_85; +x_77 = lean_nat_add(x_63, x_44); +lean_dec(x_63); +x_78 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_78, 0, x_61); +lean_ctor_set(x_78, 1, x_43); +lean_ctor_set(x_78, 2, x_75); +x_79 = lean_array_uset(x_64, x_74, x_78); +x_80 = lean_unsigned_to_nat(4u); +x_81 = lean_nat_mul(x_77, x_80); +x_82 = lean_unsigned_to_nat(3u); +x_83 = lean_nat_div(x_81, x_82); +lean_dec(x_81); +x_84 = lean_array_get_size(x_79); +x_85 = lean_nat_dec_le(x_83, x_84); +lean_dec(x_84); +lean_dec(x_83); +if (x_85 == 0) +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_86 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__7(x_79); +lean_ctor_set(x_57, 1, x_86); +lean_ctor_set(x_57, 0, x_77); +x_87 = lean_st_ref_set(x_10, x_56, x_58); +x_88 = lean_ctor_get(x_87, 1); +lean_inc(x_88); +lean_dec(x_87); +x_46 = x_88; +goto block_54; +} +else +{ +lean_object* x_89; lean_object* x_90; +lean_ctor_set(x_57, 1, x_79); +lean_ctor_set(x_57, 0, x_77); +x_89 = lean_st_ref_set(x_10, x_56, x_58); +x_90 = lean_ctor_get(x_89, 1); +lean_inc(x_90); +lean_dec(x_89); +x_46 = x_90; +goto block_54; +} +} +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_box(0); +x_92 = lean_array_uset(x_64, x_74, x_91); +x_93 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__10(x_61, x_43, x_75); +x_94 = lean_array_uset(x_92, x_74, x_93); +lean_ctor_set(x_57, 1, x_94); +x_95 = lean_st_ref_set(x_10, x_56, x_58); +x_96 = lean_ctor_get(x_95, 1); +lean_inc(x_96); +lean_dec(x_95); +x_46 = x_96; +goto block_54; +} +} +else +{ +lean_object* x_97; lean_object* 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; size_t x_105; size_t x_106; size_t x_107; size_t x_108; lean_object* x_109; uint8_t x_110; +x_97 = lean_ctor_get(x_57, 0); +x_98 = lean_ctor_get(x_57, 1); +lean_inc(x_98); +lean_inc(x_97); +lean_dec(x_57); +x_99 = lean_array_get_size(x_98); +x_100 = l___private_Lean_Expr_0__Lean_hashFVarId____x40_Lean_Expr___hyg_1730_(x_61); +x_101 = lean_uint64_shift_right(x_100, x_2); +x_102 = lean_uint64_xor(x_100, x_101); +x_103 = lean_uint64_shift_right(x_102, x_3); +x_104 = lean_uint64_xor(x_102, x_103); +x_105 = lean_uint64_to_usize(x_104); +x_106 = lean_usize_of_nat(x_99); +lean_dec(x_99); +x_107 = lean_usize_sub(x_106, x_4); +x_108 = lean_usize_land(x_105, x_107); +x_109 = lean_array_uget(x_98, x_108); +x_110 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__6(x_61, x_109); +if (x_110 == 0) +{ +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; uint8_t x_119; +x_111 = lean_nat_add(x_97, x_44); +lean_dec(x_97); +x_112 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_112, 0, x_61); +lean_ctor_set(x_112, 1, x_43); +lean_ctor_set(x_112, 2, x_109); +x_113 = lean_array_uset(x_98, x_108, 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; lean_object* x_122; lean_object* x_123; +x_120 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__7(x_113); +x_121 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_121, 0, x_111); +lean_ctor_set(x_121, 1, x_120); +lean_ctor_set(x_56, 1, x_121); +x_122 = lean_st_ref_set(x_10, x_56, x_58); +x_123 = lean_ctor_get(x_122, 1); +lean_inc(x_123); +lean_dec(x_122); +x_46 = x_123; +goto block_54; +} +else +{ +lean_object* x_124; lean_object* x_125; lean_object* x_126; +x_124 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_124, 0, x_111); +lean_ctor_set(x_124, 1, x_113); +lean_ctor_set(x_56, 1, x_124); +x_125 = lean_st_ref_set(x_10, x_56, x_58); +x_126 = lean_ctor_get(x_125, 1); +lean_inc(x_126); +lean_dec(x_125); +x_46 = x_126; +goto block_54; +} +} +else +{ +lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; +x_127 = lean_box(0); +x_128 = lean_array_uset(x_98, x_108, x_127); +x_129 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__10(x_61, x_43, x_109); +x_130 = lean_array_uset(x_128, x_108, x_129); +x_131 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_131, 0, x_97); +lean_ctor_set(x_131, 1, x_130); +lean_ctor_set(x_56, 1, x_131); +x_132 = lean_st_ref_set(x_10, x_56, x_58); +x_133 = lean_ctor_get(x_132, 1); +lean_inc(x_133); +lean_dec(x_132); +x_46 = x_133; +goto block_54; +} +} +} +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; uint64_t x_140; uint64_t x_141; uint64_t x_142; uint64_t x_143; uint64_t x_144; size_t x_145; size_t x_146; size_t x_147; size_t x_148; lean_object* x_149; uint8_t x_150; +x_134 = lean_ctor_get(x_56, 0); +lean_inc(x_134); +lean_dec(x_56); +x_135 = lean_ctor_get(x_18, 0); +lean_inc(x_135); +x_136 = lean_ctor_get(x_57, 0); +lean_inc(x_136); +x_137 = lean_ctor_get(x_57, 1); +lean_inc(x_137); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + lean_ctor_release(x_57, 1); + x_138 = x_57; +} else { + lean_dec_ref(x_57); + x_138 = lean_box(0); +} +x_139 = lean_array_get_size(x_137); +x_140 = l___private_Lean_Expr_0__Lean_hashFVarId____x40_Lean_Expr___hyg_1730_(x_135); +x_141 = lean_uint64_shift_right(x_140, x_2); +x_142 = lean_uint64_xor(x_140, x_141); +x_143 = lean_uint64_shift_right(x_142, x_3); +x_144 = lean_uint64_xor(x_142, x_143); +x_145 = lean_uint64_to_usize(x_144); +x_146 = lean_usize_of_nat(x_139); +lean_dec(x_139); +x_147 = lean_usize_sub(x_146, x_4); +x_148 = lean_usize_land(x_145, x_147); +x_149 = lean_array_uget(x_137, x_148); +x_150 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__6(x_135, x_149); +if (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; uint8_t x_159; +x_151 = lean_nat_add(x_136, x_44); +lean_dec(x_136); +x_152 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_152, 0, x_135); +lean_ctor_set(x_152, 1, x_43); +lean_ctor_set(x_152, 2, x_149); +x_153 = lean_array_uset(x_137, x_148, x_152); +x_154 = lean_unsigned_to_nat(4u); +x_155 = lean_nat_mul(x_151, x_154); +x_156 = lean_unsigned_to_nat(3u); +x_157 = lean_nat_div(x_155, x_156); +lean_dec(x_155); +x_158 = lean_array_get_size(x_153); +x_159 = lean_nat_dec_le(x_157, x_158); +lean_dec(x_158); +lean_dec(x_157); +if (x_159 == 0) +{ +lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; +x_160 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__7(x_153); +if (lean_is_scalar(x_138)) { + x_161 = lean_alloc_ctor(0, 2, 0); +} else { + x_161 = x_138; +} +lean_ctor_set(x_161, 0, x_151); +lean_ctor_set(x_161, 1, x_160); +x_162 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_162, 0, x_134); +lean_ctor_set(x_162, 1, x_161); +x_163 = lean_st_ref_set(x_10, x_162, x_58); +x_164 = lean_ctor_get(x_163, 1); +lean_inc(x_164); +lean_dec(x_163); +x_46 = x_164; +goto block_54; +} +else +{ +lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; +if (lean_is_scalar(x_138)) { + x_165 = lean_alloc_ctor(0, 2, 0); +} else { + x_165 = x_138; +} +lean_ctor_set(x_165, 0, x_151); +lean_ctor_set(x_165, 1, x_153); +x_166 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_166, 0, x_134); +lean_ctor_set(x_166, 1, x_165); +x_167 = lean_st_ref_set(x_10, x_166, x_58); +x_168 = lean_ctor_get(x_167, 1); +lean_inc(x_168); +lean_dec(x_167); +x_46 = x_168; +goto block_54; +} +} +else +{ +lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; +x_169 = lean_box(0); +x_170 = lean_array_uset(x_137, x_148, x_169); +x_171 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__10(x_135, x_43, x_149); +x_172 = lean_array_uset(x_170, x_148, x_171); +if (lean_is_scalar(x_138)) { + x_173 = lean_alloc_ctor(0, 2, 0); +} else { + x_173 = x_138; +} +lean_ctor_set(x_173, 0, x_136); +lean_ctor_set(x_173, 1, x_172); +x_174 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_174, 0, x_134); +lean_ctor_set(x_174, 1, x_173); +x_175 = lean_st_ref_set(x_10, x_174, x_58); +x_176 = lean_ctor_get(x_175, 1); +lean_inc(x_176); +lean_dec(x_175); +x_46 = x_176; +goto block_54; +} +} +block_54: +{ +lean_object* x_47; uint8_t x_48; +x_47 = l_Lean_Compiler_LCNF_eraseParam(x_18, x_11, x_12, x_13, x_14, x_46); +lean_dec(x_18); +x_48 = !lean_is_exclusive(x_47); +if (x_48 == 0) +{ +lean_object* x_49; lean_object* x_50; +x_49 = lean_ctor_get(x_47, 0); +lean_dec(x_49); +x_50 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_50, 0, x_9); +lean_ctor_set(x_47, 0, x_50); +x_19 = x_47; +goto block_32; +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_47, 1); +lean_inc(x_51); +lean_dec(x_47); +x_52 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_52, 0, x_9); +x_53 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_51); +x_19 = x_53; +goto block_32; +} +} +} +else +{ +lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; 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; uint64_t x_199; uint64_t x_200; uint64_t x_201; uint64_t x_202; uint64_t x_203; size_t x_204; size_t x_205; size_t x_206; size_t x_207; lean_object* x_208; uint8_t x_209; +lean_dec(x_9); +x_177 = lean_array_fget(x_33, x_34); +x_178 = lean_unsigned_to_nat(1u); +x_179 = lean_nat_add(x_34, x_178); +lean_dec(x_34); +x_180 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_180, 0, x_33); +lean_ctor_set(x_180, 1, x_179); +lean_ctor_set(x_180, 2, x_35); +x_188 = lean_st_ref_take(x_10, x_15); +x_189 = lean_ctor_get(x_188, 0); +lean_inc(x_189); +x_190 = lean_ctor_get(x_189, 1); +lean_inc(x_190); +x_191 = lean_ctor_get(x_188, 1); +lean_inc(x_191); +lean_dec(x_188); +x_192 = lean_ctor_get(x_189, 0); +lean_inc(x_192); +if (lean_is_exclusive(x_189)) { + lean_ctor_release(x_189, 0); + lean_ctor_release(x_189, 1); + x_193 = x_189; +} else { + lean_dec_ref(x_189); + x_193 = lean_box(0); +} +x_194 = lean_ctor_get(x_18, 0); +lean_inc(x_194); +x_195 = lean_ctor_get(x_190, 0); +lean_inc(x_195); +x_196 = lean_ctor_get(x_190, 1); +lean_inc(x_196); +if (lean_is_exclusive(x_190)) { + lean_ctor_release(x_190, 0); + lean_ctor_release(x_190, 1); + x_197 = x_190; +} else { + lean_dec_ref(x_190); + x_197 = lean_box(0); +} +x_198 = lean_array_get_size(x_196); +x_199 = l___private_Lean_Expr_0__Lean_hashFVarId____x40_Lean_Expr___hyg_1730_(x_194); +x_200 = lean_uint64_shift_right(x_199, x_2); +x_201 = lean_uint64_xor(x_199, x_200); +x_202 = lean_uint64_shift_right(x_201, x_3); +x_203 = lean_uint64_xor(x_201, x_202); +x_204 = lean_uint64_to_usize(x_203); +x_205 = lean_usize_of_nat(x_198); +lean_dec(x_198); +x_206 = lean_usize_sub(x_205, x_4); +x_207 = lean_usize_land(x_204, x_206); +x_208 = lean_array_uget(x_196, x_207); +x_209 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__6(x_194, x_208); +if (x_209 == 0) +{ +lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; uint8_t x_218; +x_210 = lean_nat_add(x_195, x_178); +lean_dec(x_195); +x_211 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_211, 0, x_194); +lean_ctor_set(x_211, 1, x_177); +lean_ctor_set(x_211, 2, x_208); +x_212 = lean_array_uset(x_196, x_207, x_211); +x_213 = lean_unsigned_to_nat(4u); +x_214 = lean_nat_mul(x_210, x_213); +x_215 = lean_unsigned_to_nat(3u); +x_216 = lean_nat_div(x_214, x_215); +lean_dec(x_214); +x_217 = lean_array_get_size(x_212); +x_218 = lean_nat_dec_le(x_216, x_217); +lean_dec(x_217); +lean_dec(x_216); +if (x_218 == 0) +{ +lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; +x_219 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__7(x_212); +if (lean_is_scalar(x_197)) { + x_220 = lean_alloc_ctor(0, 2, 0); +} else { + x_220 = x_197; +} +lean_ctor_set(x_220, 0, x_210); +lean_ctor_set(x_220, 1, x_219); +if (lean_is_scalar(x_193)) { + x_221 = lean_alloc_ctor(0, 2, 0); +} else { + x_221 = x_193; +} +lean_ctor_set(x_221, 0, x_192); +lean_ctor_set(x_221, 1, x_220); +x_222 = lean_st_ref_set(x_10, x_221, x_191); +x_223 = lean_ctor_get(x_222, 1); +lean_inc(x_223); +lean_dec(x_222); +x_181 = x_223; +goto block_187; +} +else +{ +lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; +if (lean_is_scalar(x_197)) { + x_224 = lean_alloc_ctor(0, 2, 0); +} else { + x_224 = x_197; +} +lean_ctor_set(x_224, 0, x_210); +lean_ctor_set(x_224, 1, x_212); +if (lean_is_scalar(x_193)) { + x_225 = lean_alloc_ctor(0, 2, 0); +} else { + x_225 = x_193; +} +lean_ctor_set(x_225, 0, x_192); +lean_ctor_set(x_225, 1, x_224); +x_226 = lean_st_ref_set(x_10, x_225, x_191); +x_227 = lean_ctor_get(x_226, 1); +lean_inc(x_227); +lean_dec(x_226); +x_181 = x_227; +goto block_187; +} +} +else +{ +lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; +x_228 = lean_box(0); +x_229 = lean_array_uset(x_196, x_207, x_228); +x_230 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__10(x_194, x_177, x_208); +x_231 = lean_array_uset(x_229, x_207, x_230); +if (lean_is_scalar(x_197)) { + x_232 = lean_alloc_ctor(0, 2, 0); +} else { + x_232 = x_197; +} +lean_ctor_set(x_232, 0, x_195); +lean_ctor_set(x_232, 1, x_231); +if (lean_is_scalar(x_193)) { + x_233 = lean_alloc_ctor(0, 2, 0); +} else { + x_233 = x_193; +} +lean_ctor_set(x_233, 0, x_192); +lean_ctor_set(x_233, 1, x_232); +x_234 = lean_st_ref_set(x_10, x_233, x_191); +x_235 = lean_ctor_get(x_234, 1); +lean_inc(x_235); +lean_dec(x_234); +x_181 = x_235; +goto block_187; +} +block_187: +{ +lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; +x_182 = l_Lean_Compiler_LCNF_eraseParam(x_18, x_11, x_12, x_13, x_14, x_181); +lean_dec(x_18); +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_182); + x_184 = lean_box(0); +} +x_185 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_185, 0, x_180); +if (lean_is_scalar(x_184)) { + x_186 = lean_alloc_ctor(0, 2, 0); +} else { + x_186 = x_184; +} +lean_ctor_set(x_186, 0, x_185); +lean_ctor_set(x_186, 1, x_183); +x_19 = x_186; +goto block_32; +} +} +} +block_32: +{ +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; +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; size_t x_29; size_t x_30; +x_27 = lean_ctor_get(x_19, 1); +lean_inc(x_27); +lean_dec(x_19); +x_28 = lean_ctor_get(x_20, 0); +lean_inc(x_28); +lean_dec(x_20); +x_29 = 1; +x_30 = lean_usize_add(x_8, x_29); +x_8 = x_30; +x_9 = x_28; +x_15 = x_27; +goto _start; +} +} +} +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Compiler.LCNF.StructProjCases.visitCode", 44, 44); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("expected struct constructor", 27, 27); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_StructProjCases_visitCode___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 = l_Lean_Compiler_LCNF_StructProjCases_visitLetValue___closed__1; +x_2 = l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__1; +x_3 = lean_unsigned_to_nat(61u); +x_4 = lean_unsigned_to_nat(59u); +x_5 = l_Lean_Compiler_LCNF_StructProjCases_visitCode___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_Compiler_LCNF_StructProjCases_visitCode___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("assertion violation: ", 21, 21); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("params.size == ctorInfo.numFields\n ", 42, 42); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__4; +x_2 = l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__5; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_StructProjCases_visitCode___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_Compiler_LCNF_StructProjCases_visitLetValue___closed__1; +x_2 = l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__1; +x_3 = lean_unsigned_to_nat(63u); +x_4 = lean_unsigned_to_nat(8u); +x_5 = l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__6; +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_Compiler_LCNF_StructProjCases_visitCode___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("projVars.size == params.size\n ", 37, 37); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__4; +x_2 = l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__8; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Compiler_LCNF_StructProjCases_visitLetValue___closed__1; +x_2 = l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__1; +x_3 = lean_unsigned_to_nat(85u); +x_4 = lean_unsigned_to_nat(8u); +x_5 = l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__9; +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_Compiler_LCNF_StructProjCases_visitCode(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 0: +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_8, 3); +lean_inc(x_9); +if (lean_obj_tag(x_9) == 2) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; uint64_t x_36; size_t x_37; size_t x_38; size_t x_39; size_t x_40; size_t x_41; lean_object* x_42; lean_object* x_43; +x_10 = lean_ctor_get(x_1, 1); +lean_inc(x_10); +lean_dec(x_1); +x_11 = lean_ctor_get(x_8, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_9, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_9, 1); +lean_inc(x_13); +x_14 = lean_ctor_get(x_9, 2); +lean_inc(x_14); +if (lean_is_exclusive(x_9)) { + lean_ctor_release(x_9, 0); + lean_ctor_release(x_9, 1); + lean_ctor_release(x_9, 2); + x_15 = x_9; +} else { + lean_dec_ref(x_9); + x_15 = lean_box(0); +} +x_16 = l_Lean_Compiler_LCNF_eraseLetDecl(x_8, x_3, x_4, x_5, x_6, x_7); +if (lean_is_exclusive(x_8)) { + lean_ctor_release(x_8, 0); + lean_ctor_release(x_8, 1); + lean_ctor_release(x_8, 2); + lean_ctor_release(x_8, 3); + x_17 = x_8; +} else { + lean_dec_ref(x_8); + x_17 = lean_box(0); +} +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = l_Lean_Compiler_LCNF_StructProjCases_remapFVar(x_14, x_2, x_3, x_4, x_5, x_6, 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_st_ref_get(x_2, x_21); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +x_25 = lean_ctor_get(x_22, 1); +lean_inc(x_25); +if (lean_is_exclusive(x_22)) { + lean_ctor_release(x_22, 0); + lean_ctor_release(x_22, 1); + x_26 = x_22; +} else { + lean_dec_ref(x_22); + x_26 = lean_box(0); +} +x_27 = lean_ctor_get(x_24, 1); +lean_inc(x_27); +if (lean_is_exclusive(x_24)) { + lean_ctor_release(x_24, 0); + lean_ctor_release(x_24, 1); + x_28 = x_24; +} else { + lean_dec_ref(x_24); + x_28 = lean_box(0); +} +x_29 = lean_array_get_size(x_27); +x_30 = l___private_Lean_Expr_0__Lean_hashFVarId____x40_Lean_Expr___hyg_1730_(x_20); +x_31 = 32; +x_32 = lean_uint64_shift_right(x_30, x_31); +x_33 = lean_uint64_xor(x_30, x_32); +x_34 = 16; +x_35 = lean_uint64_shift_right(x_33, x_34); +x_36 = lean_uint64_xor(x_33, x_35); +x_37 = lean_uint64_to_usize(x_36); +x_38 = lean_usize_of_nat(x_29); +lean_dec(x_29); +x_39 = 1; +x_40 = lean_usize_sub(x_38, x_39); +x_41 = lean_usize_land(x_37, x_40); +x_42 = lean_array_uget(x_27, x_41); +lean_dec(x_27); +x_43 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__1(x_20, x_42); +lean_dec(x_42); +if (lean_obj_tag(x_43) == 0) +{ +lean_object* x_44; +lean_inc(x_12); +x_44 = l_Lean_Compiler_LCNF_StructProjCases_findStructCtorInfo_x3f(x_12, x_5, x_6, x_25); +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_28); +lean_dec(x_26); +lean_dec(x_20); +lean_dec(x_17); +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +x_46 = lean_ctor_get(x_44, 1); +lean_inc(x_46); +lean_dec(x_44); +x_47 = l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__3; +x_48 = l_panic___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__2(x_47, x_2, x_3, x_4, x_5, x_6, x_46); +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_45, 0); +lean_inc(x_49); +if (lean_is_exclusive(x_45)) { + lean_ctor_release(x_45, 0); + x_50 = x_45; +} else { + lean_dec_ref(x_45); + x_50 = lean_box(0); +} +x_51 = lean_ctor_get(x_49, 0); +lean_inc(x_51); +x_52 = lean_ctor_get(x_44, 1); +lean_inc(x_52); +lean_dec(x_44); +x_53 = lean_ctor_get(x_51, 0); +lean_inc(x_53); +x_54 = lean_ctor_get(x_51, 2); +lean_inc(x_54); +if (lean_is_exclusive(x_51)) { + lean_ctor_release(x_51, 0); + lean_ctor_release(x_51, 1); + lean_ctor_release(x_51, 2); + x_55 = x_51; +} else { + lean_dec_ref(x_51); + x_55 = lean_box(0); +} +x_56 = lean_ctor_get(x_49, 3); +lean_inc(x_56); +lean_inc(x_6); +lean_inc(x_5); +x_57 = l_Lean_Compiler_LCNF_StructProjCases_mkFieldParamsForCtorType(x_54, x_56, x_3, x_4, x_5, x_6, x_52); +if (lean_obj_tag(x_57) == 0) +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; +x_58 = lean_ctor_get(x_57, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_57, 1); +lean_inc(x_59); +lean_dec(x_57); +x_60 = lean_array_get_size(x_58); +x_61 = lean_ctor_get(x_49, 4); +lean_inc(x_61); +lean_dec(x_49); +x_62 = lean_nat_dec_eq(x_60, x_61); +lean_dec(x_61); +lean_dec(x_60); +if (x_62 == 0) +{ +lean_object* x_63; lean_object* x_64; +lean_dec(x_58); +lean_dec(x_55); +lean_dec(x_53); +lean_dec(x_50); +lean_dec(x_28); +lean_dec(x_26); +lean_dec(x_20); +lean_dec(x_17); +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +x_63 = l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__7; +x_64 = l_panic___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__2(x_63, x_2, x_3, x_4, x_5, x_6, x_59); +return x_64; +} +else +{ +size_t x_65; size_t x_66; lean_object* x_67; lean_object* x_68; 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_228; +x_65 = lean_array_size(x_58); +x_66 = 0; +lean_inc(x_58); +x_67 = l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__3(x_65, x_66, x_58); +x_177 = lean_st_ref_take(x_2, x_59); +x_178 = lean_ctor_get(x_177, 0); +lean_inc(x_178); +x_179 = lean_ctor_get(x_177, 1); +lean_inc(x_179); +if (lean_is_exclusive(x_177)) { + lean_ctor_release(x_177, 0); + lean_ctor_release(x_177, 1); + x_180 = x_177; +} else { + lean_dec_ref(x_177); + x_180 = lean_box(0); +} +x_181 = lean_ctor_get(x_178, 0); +lean_inc(x_181); +x_182 = lean_ctor_get(x_178, 1); +lean_inc(x_182); +lean_dec(x_178); +x_183 = l_Lean_instInhabitedFVarId; +x_184 = lean_array_get(x_183, x_67, x_13); +lean_dec(x_13); +x_228 = !lean_is_exclusive(x_181); +if (x_228 == 0) +{ +lean_object* x_229; lean_object* x_230; lean_object* x_231; size_t x_232; size_t x_233; size_t x_234; lean_object* x_235; uint8_t x_236; +x_229 = lean_ctor_get(x_181, 0); +x_230 = lean_ctor_get(x_181, 1); +x_231 = lean_array_get_size(x_230); +x_232 = lean_usize_of_nat(x_231); +lean_dec(x_231); +x_233 = lean_usize_sub(x_232, x_39); +x_234 = lean_usize_land(x_37, x_233); +x_235 = lean_array_uget(x_230, x_234); +x_236 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__4(x_20, x_235); +if (x_236 == 0) +{ +lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; uint8_t x_246; +x_237 = lean_unsigned_to_nat(1u); +x_238 = lean_nat_add(x_229, x_237); +lean_dec(x_229); +lean_inc(x_20); +x_239 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_239, 0, x_20); +lean_ctor_set(x_239, 1, x_67); +lean_ctor_set(x_239, 2, x_235); +x_240 = lean_array_uset(x_230, x_234, x_239); +x_241 = lean_unsigned_to_nat(4u); +x_242 = lean_nat_mul(x_238, x_241); +x_243 = lean_unsigned_to_nat(3u); +x_244 = lean_nat_div(x_242, x_243); +lean_dec(x_242); +x_245 = lean_array_get_size(x_240); +x_246 = lean_nat_dec_le(x_244, x_245); +lean_dec(x_245); +lean_dec(x_244); +if (x_246 == 0) +{ +lean_object* x_247; lean_object* x_248; lean_object* x_249; +x_247 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__11(x_240); +lean_ctor_set(x_181, 1, x_247); +lean_ctor_set(x_181, 0, x_238); +x_248 = lean_ctor_get(x_182, 0); +lean_inc(x_248); +x_249 = lean_ctor_get(x_182, 1); +lean_inc(x_249); +lean_dec(x_182); +x_185 = x_181; +x_186 = x_248; +x_187 = x_249; +goto block_227; +} +else +{ +lean_object* x_250; lean_object* x_251; +lean_ctor_set(x_181, 1, x_240); +lean_ctor_set(x_181, 0, x_238); +x_250 = lean_ctor_get(x_182, 0); +lean_inc(x_250); +x_251 = lean_ctor_get(x_182, 1); +lean_inc(x_251); +lean_dec(x_182); +x_185 = x_181; +x_186 = x_250; +x_187 = x_251; +goto block_227; +} +} +else +{ +lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; +x_252 = lean_box(0); +x_253 = lean_array_uset(x_230, x_234, x_252); +lean_inc(x_20); +x_254 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__14(x_20, x_67, x_235); +x_255 = lean_array_uset(x_253, x_234, x_254); +lean_ctor_set(x_181, 1, x_255); +x_256 = lean_ctor_get(x_182, 0); +lean_inc(x_256); +x_257 = lean_ctor_get(x_182, 1); +lean_inc(x_257); +lean_dec(x_182); +x_185 = x_181; +x_186 = x_256; +x_187 = x_257; +goto block_227; +} +} +else +{ +lean_object* x_258; lean_object* x_259; lean_object* x_260; size_t x_261; size_t x_262; size_t x_263; lean_object* x_264; uint8_t x_265; +x_258 = lean_ctor_get(x_181, 0); +x_259 = lean_ctor_get(x_181, 1); +lean_inc(x_259); +lean_inc(x_258); +lean_dec(x_181); +x_260 = lean_array_get_size(x_259); +x_261 = lean_usize_of_nat(x_260); +lean_dec(x_260); +x_262 = lean_usize_sub(x_261, x_39); +x_263 = lean_usize_land(x_37, x_262); +x_264 = lean_array_uget(x_259, x_263); +x_265 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__4(x_20, 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_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; uint8_t x_275; +x_266 = lean_unsigned_to_nat(1u); +x_267 = lean_nat_add(x_258, x_266); +lean_dec(x_258); +lean_inc(x_20); +x_268 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_268, 0, x_20); +lean_ctor_set(x_268, 1, x_67); +lean_ctor_set(x_268, 2, x_264); +x_269 = lean_array_uset(x_259, x_263, x_268); +x_270 = lean_unsigned_to_nat(4u); +x_271 = lean_nat_mul(x_267, x_270); +x_272 = lean_unsigned_to_nat(3u); +x_273 = lean_nat_div(x_271, x_272); +lean_dec(x_271); +x_274 = lean_array_get_size(x_269); +x_275 = lean_nat_dec_le(x_273, x_274); +lean_dec(x_274); +lean_dec(x_273); +if (x_275 == 0) +{ +lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; +x_276 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__11(x_269); +x_277 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_277, 0, x_267); +lean_ctor_set(x_277, 1, x_276); +x_278 = lean_ctor_get(x_182, 0); +lean_inc(x_278); +x_279 = lean_ctor_get(x_182, 1); +lean_inc(x_279); +lean_dec(x_182); +x_185 = x_277; +x_186 = x_278; +x_187 = x_279; +goto block_227; +} +else +{ +lean_object* x_280; lean_object* x_281; lean_object* x_282; +x_280 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_280, 0, x_267); +lean_ctor_set(x_280, 1, x_269); +x_281 = lean_ctor_get(x_182, 0); +lean_inc(x_281); +x_282 = lean_ctor_get(x_182, 1); +lean_inc(x_282); +lean_dec(x_182); +x_185 = x_280; +x_186 = x_281; +x_187 = x_282; +goto block_227; +} +} +else +{ +lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; +x_283 = lean_box(0); +x_284 = lean_array_uset(x_259, x_263, x_283); +lean_inc(x_20); +x_285 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__14(x_20, x_67, x_264); +x_286 = lean_array_uset(x_284, x_263, x_285); +x_287 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_287, 0, x_258); +lean_ctor_set(x_287, 1, x_286); +x_288 = lean_ctor_get(x_182, 0); +lean_inc(x_288); +x_289 = lean_ctor_get(x_182, 1); +lean_inc(x_289); +lean_dec(x_182); +x_185 = x_287; +x_186 = x_288; +x_187 = x_289; +goto block_227; +} +} +block_176: +{ +lean_object* x_69; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_69 = l_Lean_Compiler_LCNF_StructProjCases_visitCode(x_10, x_2, x_3, x_4, x_5, x_6, x_68); +if (lean_obj_tag(x_69) == 0) +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; uint8_t x_107; +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_103 = lean_st_ref_take(x_2, x_71); +x_104 = lean_ctor_get(x_103, 0); +lean_inc(x_104); +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_103, 1); +lean_inc(x_106); +lean_dec(x_103); +x_107 = !lean_is_exclusive(x_104); +if (x_107 == 0) +{ +lean_object* x_108; uint8_t x_109; +x_108 = lean_ctor_get(x_104, 0); +lean_dec(x_108); +x_109 = !lean_is_exclusive(x_105); +if (x_109 == 0) +{ +lean_object* x_110; lean_object* x_111; lean_object* x_112; size_t x_113; size_t x_114; size_t x_115; lean_object* x_116; uint8_t x_117; +x_110 = lean_ctor_get(x_105, 0); +x_111 = lean_ctor_get(x_105, 1); +x_112 = lean_array_get_size(x_111); +x_113 = lean_usize_of_nat(x_112); +lean_dec(x_112); +x_114 = lean_usize_sub(x_113, x_39); +x_115 = lean_usize_land(x_37, x_114); +x_116 = lean_array_uget(x_111, x_115); +x_117 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__4(x_20, x_116); +if (x_117 == 0) +{ +lean_object* x_118; lean_object* x_119; +lean_dec(x_116); +x_118 = lean_st_ref_set(x_2, x_104, x_106); +lean_dec(x_2); +x_119 = lean_ctor_get(x_118, 1); +lean_inc(x_119); +lean_dec(x_118); +x_72 = x_119; +goto block_102; +} +else +{ +lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_120 = lean_box(0); +x_121 = lean_array_uset(x_111, x_115, x_120); +x_122 = lean_unsigned_to_nat(1u); +x_123 = lean_nat_sub(x_110, x_122); +lean_dec(x_110); +x_124 = l_Std_DHashMap_Internal_AssocList_erase___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__5(x_20, x_116); +x_125 = lean_array_uset(x_121, x_115, x_124); +lean_ctor_set(x_105, 1, x_125); +lean_ctor_set(x_105, 0, x_123); +x_126 = lean_st_ref_set(x_2, x_104, x_106); +lean_dec(x_2); +x_127 = lean_ctor_get(x_126, 1); +lean_inc(x_127); +lean_dec(x_126); +x_72 = x_127; +goto block_102; +} +} +else +{ +lean_object* x_128; lean_object* x_129; lean_object* x_130; size_t x_131; size_t x_132; size_t x_133; lean_object* x_134; uint8_t x_135; +x_128 = lean_ctor_get(x_105, 0); +x_129 = lean_ctor_get(x_105, 1); +lean_inc(x_129); +lean_inc(x_128); +lean_dec(x_105); +x_130 = lean_array_get_size(x_129); +x_131 = lean_usize_of_nat(x_130); +lean_dec(x_130); +x_132 = lean_usize_sub(x_131, x_39); +x_133 = lean_usize_land(x_37, x_132); +x_134 = lean_array_uget(x_129, x_133); +x_135 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__4(x_20, x_134); +if (x_135 == 0) +{ +lean_object* x_136; lean_object* x_137; lean_object* x_138; +lean_dec(x_134); +x_136 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_136, 0, x_128); +lean_ctor_set(x_136, 1, x_129); +lean_ctor_set(x_104, 0, x_136); +x_137 = lean_st_ref_set(x_2, x_104, x_106); +lean_dec(x_2); +x_138 = lean_ctor_get(x_137, 1); +lean_inc(x_138); +lean_dec(x_137); +x_72 = x_138; +goto block_102; +} +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; +x_139 = lean_box(0); +x_140 = lean_array_uset(x_129, x_133, x_139); +x_141 = lean_unsigned_to_nat(1u); +x_142 = lean_nat_sub(x_128, x_141); +lean_dec(x_128); +x_143 = l_Std_DHashMap_Internal_AssocList_erase___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__5(x_20, x_134); +x_144 = lean_array_uset(x_140, x_133, x_143); +x_145 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_145, 0, x_142); +lean_ctor_set(x_145, 1, x_144); +lean_ctor_set(x_104, 0, x_145); +x_146 = lean_st_ref_set(x_2, x_104, x_106); +lean_dec(x_2); +x_147 = lean_ctor_get(x_146, 1); +lean_inc(x_147); +lean_dec(x_146); +x_72 = x_147; +goto block_102; +} +} +} +else +{ +lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; size_t x_153; size_t x_154; size_t x_155; lean_object* x_156; uint8_t x_157; +x_148 = lean_ctor_get(x_104, 1); +lean_inc(x_148); +lean_dec(x_104); +x_149 = lean_ctor_get(x_105, 0); +lean_inc(x_149); +x_150 = lean_ctor_get(x_105, 1); +lean_inc(x_150); +if (lean_is_exclusive(x_105)) { + lean_ctor_release(x_105, 0); + lean_ctor_release(x_105, 1); + x_151 = x_105; +} else { + lean_dec_ref(x_105); + x_151 = lean_box(0); +} +x_152 = lean_array_get_size(x_150); +x_153 = lean_usize_of_nat(x_152); +lean_dec(x_152); +x_154 = lean_usize_sub(x_153, x_39); +x_155 = lean_usize_land(x_37, x_154); +x_156 = lean_array_uget(x_150, x_155); +x_157 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__4(x_20, x_156); +if (x_157 == 0) +{ +lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; +lean_dec(x_156); +if (lean_is_scalar(x_151)) { + x_158 = lean_alloc_ctor(0, 2, 0); +} else { + x_158 = x_151; +} +lean_ctor_set(x_158, 0, x_149); +lean_ctor_set(x_158, 1, x_150); +x_159 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_159, 0, x_158); +lean_ctor_set(x_159, 1, x_148); +x_160 = lean_st_ref_set(x_2, x_159, x_106); +lean_dec(x_2); +x_161 = lean_ctor_get(x_160, 1); +lean_inc(x_161); +lean_dec(x_160); +x_72 = x_161; +goto block_102; +} +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_150, x_155, x_162); +x_164 = lean_unsigned_to_nat(1u); +x_165 = lean_nat_sub(x_149, x_164); +lean_dec(x_149); +x_166 = l_Std_DHashMap_Internal_AssocList_erase___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__5(x_20, x_156); +x_167 = lean_array_uset(x_163, x_155, x_166); +if (lean_is_scalar(x_151)) { + x_168 = lean_alloc_ctor(0, 2, 0); +} else { + x_168 = x_151; +} +lean_ctor_set(x_168, 0, x_165); +lean_ctor_set(x_168, 1, x_167); +x_169 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_169, 0, x_168); +lean_ctor_set(x_169, 1, x_148); +x_170 = lean_st_ref_set(x_2, x_169, x_106); +lean_dec(x_2); +x_171 = lean_ctor_get(x_170, 1); +lean_inc(x_171); +lean_dec(x_170); +x_72 = x_171; +goto block_102; +} +} +block_102: +{ +lean_object* x_73; +lean_inc(x_70); +x_73 = l_Lean_Compiler_LCNF_Code_inferType(x_70, x_3, x_4, x_5, x_6, x_72); +lean_dec(x_4); +lean_dec(x_3); +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_Compiler_LCNF_toMonoType(x_74, x_5, x_6, 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_78 = lean_ctor_get(x_76, 0); +if (lean_is_scalar(x_55)) { + x_79 = lean_alloc_ctor(0, 3, 0); +} else { + x_79 = x_55; +} +lean_ctor_set(x_79, 0, x_53); +lean_ctor_set(x_79, 1, x_58); +lean_ctor_set(x_79, 2, x_70); +x_80 = lean_box(0); +if (lean_is_scalar(x_26)) { + x_81 = lean_alloc_ctor(1, 2, 0); +} else { + x_81 = x_26; + lean_ctor_set_tag(x_81, 1); +} +lean_ctor_set(x_81, 0, x_79); +lean_ctor_set(x_81, 1, x_80); +x_82 = lean_array_mk(x_81); +if (lean_is_scalar(x_17)) { + x_83 = lean_alloc_ctor(0, 4, 0); +} else { + x_83 = x_17; +} +lean_ctor_set(x_83, 0, x_12); +lean_ctor_set(x_83, 1, x_78); +lean_ctor_set(x_83, 2, x_20); +lean_ctor_set(x_83, 3, x_82); +if (lean_is_scalar(x_50)) { + x_84 = lean_alloc_ctor(4, 1, 0); +} else { + x_84 = x_50; + lean_ctor_set_tag(x_84, 4); +} +lean_ctor_set(x_84, 0, x_83); +lean_ctor_set(x_76, 0, x_84); +return x_76; +} +else +{ +lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_85 = lean_ctor_get(x_76, 0); +x_86 = lean_ctor_get(x_76, 1); +lean_inc(x_86); +lean_inc(x_85); +lean_dec(x_76); +if (lean_is_scalar(x_55)) { + x_87 = lean_alloc_ctor(0, 3, 0); +} else { + x_87 = x_55; +} +lean_ctor_set(x_87, 0, x_53); +lean_ctor_set(x_87, 1, x_58); +lean_ctor_set(x_87, 2, x_70); +x_88 = lean_box(0); +if (lean_is_scalar(x_26)) { + x_89 = lean_alloc_ctor(1, 2, 0); +} else { + x_89 = x_26; + lean_ctor_set_tag(x_89, 1); +} +lean_ctor_set(x_89, 0, x_87); +lean_ctor_set(x_89, 1, x_88); +x_90 = lean_array_mk(x_89); +if (lean_is_scalar(x_17)) { + x_91 = lean_alloc_ctor(0, 4, 0); +} else { + x_91 = x_17; +} +lean_ctor_set(x_91, 0, x_12); +lean_ctor_set(x_91, 1, x_85); +lean_ctor_set(x_91, 2, x_20); +lean_ctor_set(x_91, 3, x_90); +if (lean_is_scalar(x_50)) { + x_92 = lean_alloc_ctor(4, 1, 0); +} else { + x_92 = x_50; + lean_ctor_set_tag(x_92, 4); +} +lean_ctor_set(x_92, 0, x_91); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_92); +lean_ctor_set(x_93, 1, x_86); +return x_93; +} +} +else +{ +uint8_t x_94; +lean_dec(x_70); +lean_dec(x_58); +lean_dec(x_55); +lean_dec(x_53); +lean_dec(x_50); +lean_dec(x_26); +lean_dec(x_20); +lean_dec(x_17); +lean_dec(x_12); +x_94 = !lean_is_exclusive(x_76); +if (x_94 == 0) +{ +return x_76; +} +else +{ +lean_object* x_95; lean_object* x_96; lean_object* x_97; +x_95 = lean_ctor_get(x_76, 0); +x_96 = lean_ctor_get(x_76, 1); +lean_inc(x_96); +lean_inc(x_95); +lean_dec(x_76); +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_70); +lean_dec(x_58); +lean_dec(x_55); +lean_dec(x_53); +lean_dec(x_50); +lean_dec(x_26); +lean_dec(x_20); +lean_dec(x_17); +lean_dec(x_12); +lean_dec(x_6); +lean_dec(x_5); +x_98 = !lean_is_exclusive(x_73); +if (x_98 == 0) +{ +return x_73; +} +else +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_99 = lean_ctor_get(x_73, 0); +x_100 = lean_ctor_get(x_73, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_73); +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_172; +lean_dec(x_58); +lean_dec(x_55); +lean_dec(x_53); +lean_dec(x_50); +lean_dec(x_26); +lean_dec(x_20); +lean_dec(x_17); +lean_dec(x_12); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_172 = !lean_is_exclusive(x_69); +if (x_172 == 0) +{ +return x_69; +} +else +{ +lean_object* x_173; lean_object* x_174; lean_object* x_175; +x_173 = lean_ctor_get(x_69, 0); +x_174 = lean_ctor_get(x_69, 1); +lean_inc(x_174); +lean_inc(x_173); +lean_dec(x_69); +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; +} +} +} +block_227: +{ +lean_object* x_188; uint64_t x_189; uint64_t x_190; uint64_t x_191; uint64_t x_192; uint64_t x_193; size_t x_194; size_t x_195; size_t x_196; size_t x_197; lean_object* x_198; uint8_t x_199; +x_188 = lean_array_get_size(x_187); +x_189 = l___private_Lean_Expr_0__Lean_hashFVarId____x40_Lean_Expr___hyg_1730_(x_11); +x_190 = lean_uint64_shift_right(x_189, x_31); +x_191 = lean_uint64_xor(x_189, x_190); +x_192 = lean_uint64_shift_right(x_191, x_34); +x_193 = lean_uint64_xor(x_191, x_192); +x_194 = lean_uint64_to_usize(x_193); +x_195 = lean_usize_of_nat(x_188); +lean_dec(x_188); +x_196 = lean_usize_sub(x_195, x_39); +x_197 = lean_usize_land(x_194, x_196); +x_198 = lean_array_uget(x_187, x_197); +x_199 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__6(x_11, x_198); +if (x_199 == 0) +{ +lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; uint8_t x_209; +x_200 = lean_unsigned_to_nat(1u); +x_201 = lean_nat_add(x_186, x_200); +lean_dec(x_186); +if (lean_is_scalar(x_15)) { + x_202 = lean_alloc_ctor(1, 3, 0); +} else { + x_202 = x_15; + lean_ctor_set_tag(x_202, 1); +} +lean_ctor_set(x_202, 0, x_11); +lean_ctor_set(x_202, 1, x_184); +lean_ctor_set(x_202, 2, x_198); +x_203 = lean_array_uset(x_187, x_197, x_202); +x_204 = lean_unsigned_to_nat(4u); +x_205 = lean_nat_mul(x_201, x_204); +x_206 = lean_unsigned_to_nat(3u); +x_207 = lean_nat_div(x_205, x_206); +lean_dec(x_205); +x_208 = lean_array_get_size(x_203); +x_209 = lean_nat_dec_le(x_207, x_208); +lean_dec(x_208); +lean_dec(x_207); +if (x_209 == 0) +{ +lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; +x_210 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__7(x_203); +if (lean_is_scalar(x_28)) { + x_211 = lean_alloc_ctor(0, 2, 0); +} else { + x_211 = x_28; +} +lean_ctor_set(x_211, 0, x_201); +lean_ctor_set(x_211, 1, x_210); +if (lean_is_scalar(x_180)) { + x_212 = lean_alloc_ctor(0, 2, 0); +} else { + x_212 = x_180; +} +lean_ctor_set(x_212, 0, x_185); +lean_ctor_set(x_212, 1, x_211); +x_213 = lean_st_ref_set(x_2, x_212, x_179); +x_214 = lean_ctor_get(x_213, 1); +lean_inc(x_214); +lean_dec(x_213); +x_68 = x_214; +goto block_176; +} +else +{ +lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; +if (lean_is_scalar(x_28)) { + x_215 = lean_alloc_ctor(0, 2, 0); +} else { + x_215 = x_28; +} +lean_ctor_set(x_215, 0, x_201); +lean_ctor_set(x_215, 1, x_203); +if (lean_is_scalar(x_180)) { + x_216 = lean_alloc_ctor(0, 2, 0); +} else { + x_216 = x_180; +} +lean_ctor_set(x_216, 0, x_185); +lean_ctor_set(x_216, 1, x_215); +x_217 = lean_st_ref_set(x_2, x_216, x_179); +x_218 = lean_ctor_get(x_217, 1); +lean_inc(x_218); +lean_dec(x_217); +x_68 = x_218; +goto block_176; +} +} +else +{ +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_15); +x_219 = lean_box(0); +x_220 = lean_array_uset(x_187, x_197, x_219); +x_221 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__10(x_11, x_184, x_198); +x_222 = lean_array_uset(x_220, x_197, x_221); +if (lean_is_scalar(x_28)) { + x_223 = lean_alloc_ctor(0, 2, 0); +} else { + x_223 = x_28; +} +lean_ctor_set(x_223, 0, x_186); +lean_ctor_set(x_223, 1, x_222); +if (lean_is_scalar(x_180)) { + x_224 = lean_alloc_ctor(0, 2, 0); +} else { + x_224 = x_180; +} +lean_ctor_set(x_224, 0, x_185); +lean_ctor_set(x_224, 1, x_223); +x_225 = lean_st_ref_set(x_2, x_224, x_179); +x_226 = lean_ctor_get(x_225, 1); +lean_inc(x_226); +lean_dec(x_225); +x_68 = x_226; +goto block_176; +} +} +} +} +else +{ +uint8_t x_290; +lean_dec(x_55); +lean_dec(x_53); +lean_dec(x_50); +lean_dec(x_49); +lean_dec(x_28); +lean_dec(x_26); +lean_dec(x_20); +lean_dec(x_17); +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_290 = !lean_is_exclusive(x_57); +if (x_290 == 0) +{ +return x_57; +} +else +{ +lean_object* x_291; lean_object* x_292; lean_object* x_293; +x_291 = lean_ctor_get(x_57, 0); +x_292 = lean_ctor_get(x_57, 1); +lean_inc(x_292); +lean_inc(x_291); +lean_dec(x_57); +x_293 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_293, 0, x_291); +lean_ctor_set(x_293, 1, x_292); +return x_293; +} +} +} +} +else +{ +uint8_t x_294; +lean_dec(x_28); +lean_dec(x_26); +lean_dec(x_20); +lean_dec(x_17); +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_294 = !lean_is_exclusive(x_44); +if (x_294 == 0) +{ +return x_44; +} +else +{ +lean_object* x_295; lean_object* x_296; lean_object* x_297; +x_295 = lean_ctor_get(x_44, 0); +x_296 = lean_ctor_get(x_44, 1); +lean_inc(x_296); +lean_inc(x_295); +lean_dec(x_44); +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; +} +} +} +else +{ +lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; uint8_t x_302; +lean_dec(x_28); +lean_dec(x_26); +lean_dec(x_20); +lean_dec(x_17); +lean_dec(x_12); +x_298 = lean_ctor_get(x_43, 0); +lean_inc(x_298); +lean_dec(x_43); +x_299 = lean_st_ref_take(x_2, x_25); +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_is_exclusive(x_300); +if (x_302 == 0) +{ +lean_object* x_303; lean_object* x_304; lean_object* x_305; uint8_t x_306; +x_303 = lean_ctor_get(x_300, 1); +x_304 = l_Lean_instInhabitedFVarId; +x_305 = lean_array_get(x_304, x_298, x_13); +lean_dec(x_13); +lean_dec(x_298); +x_306 = !lean_is_exclusive(x_303); +if (x_306 == 0) +{ +lean_object* x_307; lean_object* x_308; lean_object* x_309; uint64_t x_310; uint64_t x_311; uint64_t x_312; uint64_t x_313; uint64_t x_314; size_t x_315; size_t x_316; size_t x_317; size_t x_318; lean_object* x_319; uint8_t x_320; +x_307 = lean_ctor_get(x_303, 0); +x_308 = lean_ctor_get(x_303, 1); +x_309 = lean_array_get_size(x_308); +x_310 = l___private_Lean_Expr_0__Lean_hashFVarId____x40_Lean_Expr___hyg_1730_(x_11); +x_311 = lean_uint64_shift_right(x_310, x_31); +x_312 = lean_uint64_xor(x_310, x_311); +x_313 = lean_uint64_shift_right(x_312, x_34); +x_314 = lean_uint64_xor(x_312, x_313); +x_315 = lean_uint64_to_usize(x_314); +x_316 = lean_usize_of_nat(x_309); +lean_dec(x_309); +x_317 = lean_usize_sub(x_316, x_39); +x_318 = lean_usize_land(x_315, x_317); +x_319 = lean_array_uget(x_308, x_318); +x_320 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__6(x_11, x_319); +if (x_320 == 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; uint8_t x_330; +x_321 = lean_unsigned_to_nat(1u); +x_322 = lean_nat_add(x_307, x_321); +lean_dec(x_307); +if (lean_is_scalar(x_15)) { + x_323 = lean_alloc_ctor(1, 3, 0); +} else { + x_323 = x_15; + lean_ctor_set_tag(x_323, 1); +} +lean_ctor_set(x_323, 0, x_11); +lean_ctor_set(x_323, 1, x_305); +lean_ctor_set(x_323, 2, x_319); +x_324 = lean_array_uset(x_308, x_318, x_323); +x_325 = lean_unsigned_to_nat(4u); +x_326 = lean_nat_mul(x_322, x_325); +x_327 = lean_unsigned_to_nat(3u); +x_328 = lean_nat_div(x_326, x_327); +lean_dec(x_326); +x_329 = lean_array_get_size(x_324); +x_330 = lean_nat_dec_le(x_328, x_329); +lean_dec(x_329); +lean_dec(x_328); +if (x_330 == 0) +{ +lean_object* x_331; lean_object* x_332; lean_object* x_333; +x_331 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__7(x_324); +lean_ctor_set(x_303, 1, x_331); +lean_ctor_set(x_303, 0, x_322); +x_332 = lean_st_ref_set(x_2, x_300, x_301); +x_333 = lean_ctor_get(x_332, 1); +lean_inc(x_333); +lean_dec(x_332); +x_1 = x_10; +x_7 = x_333; +goto _start; +} +else +{ +lean_object* x_335; lean_object* x_336; +lean_ctor_set(x_303, 1, x_324); +lean_ctor_set(x_303, 0, x_322); +x_335 = lean_st_ref_set(x_2, x_300, x_301); +x_336 = lean_ctor_get(x_335, 1); +lean_inc(x_336); +lean_dec(x_335); +x_1 = x_10; +x_7 = x_336; +goto _start; +} +} +else +{ +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_dec(x_15); +x_338 = lean_box(0); +x_339 = lean_array_uset(x_308, x_318, x_338); +x_340 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__10(x_11, x_305, x_319); +x_341 = lean_array_uset(x_339, x_318, x_340); +lean_ctor_set(x_303, 1, x_341); +x_342 = lean_st_ref_set(x_2, x_300, x_301); +x_343 = lean_ctor_get(x_342, 1); +lean_inc(x_343); +lean_dec(x_342); +x_1 = x_10; +x_7 = x_343; +goto _start; +} +} +else +{ +lean_object* x_345; lean_object* x_346; lean_object* x_347; uint64_t x_348; uint64_t x_349; uint64_t x_350; uint64_t x_351; uint64_t x_352; size_t x_353; size_t x_354; size_t x_355; size_t x_356; lean_object* x_357; uint8_t x_358; +x_345 = lean_ctor_get(x_303, 0); +x_346 = lean_ctor_get(x_303, 1); +lean_inc(x_346); +lean_inc(x_345); +lean_dec(x_303); +x_347 = lean_array_get_size(x_346); +x_348 = l___private_Lean_Expr_0__Lean_hashFVarId____x40_Lean_Expr___hyg_1730_(x_11); +x_349 = lean_uint64_shift_right(x_348, x_31); +x_350 = lean_uint64_xor(x_348, x_349); +x_351 = lean_uint64_shift_right(x_350, x_34); +x_352 = lean_uint64_xor(x_350, x_351); +x_353 = lean_uint64_to_usize(x_352); +x_354 = lean_usize_of_nat(x_347); +lean_dec(x_347); +x_355 = lean_usize_sub(x_354, x_39); +x_356 = lean_usize_land(x_353, x_355); +x_357 = lean_array_uget(x_346, x_356); +x_358 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__6(x_11, x_357); +if (x_358 == 0) +{ +lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; uint8_t x_368; +x_359 = lean_unsigned_to_nat(1u); +x_360 = lean_nat_add(x_345, x_359); +lean_dec(x_345); +if (lean_is_scalar(x_15)) { + x_361 = lean_alloc_ctor(1, 3, 0); +} else { + x_361 = x_15; + lean_ctor_set_tag(x_361, 1); +} +lean_ctor_set(x_361, 0, x_11); +lean_ctor_set(x_361, 1, x_305); +lean_ctor_set(x_361, 2, x_357); +x_362 = lean_array_uset(x_346, x_356, x_361); +x_363 = lean_unsigned_to_nat(4u); +x_364 = lean_nat_mul(x_360, x_363); +x_365 = lean_unsigned_to_nat(3u); +x_366 = lean_nat_div(x_364, x_365); +lean_dec(x_364); +x_367 = lean_array_get_size(x_362); +x_368 = lean_nat_dec_le(x_366, x_367); +lean_dec(x_367); +lean_dec(x_366); +if (x_368 == 0) +{ +lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; +x_369 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__7(x_362); +x_370 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_370, 0, x_360); +lean_ctor_set(x_370, 1, x_369); +lean_ctor_set(x_300, 1, x_370); +x_371 = lean_st_ref_set(x_2, x_300, x_301); +x_372 = lean_ctor_get(x_371, 1); +lean_inc(x_372); +lean_dec(x_371); +x_1 = x_10; +x_7 = x_372; +goto _start; +} +else +{ +lean_object* x_374; lean_object* x_375; lean_object* x_376; +x_374 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_374, 0, x_360); +lean_ctor_set(x_374, 1, x_362); +lean_ctor_set(x_300, 1, x_374); +x_375 = lean_st_ref_set(x_2, x_300, x_301); +x_376 = lean_ctor_get(x_375, 1); +lean_inc(x_376); +lean_dec(x_375); +x_1 = x_10; +x_7 = x_376; +goto _start; +} +} +else +{ +lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; +lean_dec(x_15); +x_378 = lean_box(0); +x_379 = lean_array_uset(x_346, x_356, x_378); +x_380 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__10(x_11, x_305, x_357); +x_381 = lean_array_uset(x_379, x_356, x_380); +x_382 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_382, 0, x_345); +lean_ctor_set(x_382, 1, x_381); +lean_ctor_set(x_300, 1, x_382); +x_383 = lean_st_ref_set(x_2, x_300, x_301); +x_384 = lean_ctor_get(x_383, 1); +lean_inc(x_384); +lean_dec(x_383); +x_1 = x_10; +x_7 = x_384; +goto _start; +} +} +} +else +{ +lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; uint64_t x_394; uint64_t x_395; uint64_t x_396; uint64_t x_397; uint64_t x_398; size_t x_399; size_t x_400; size_t x_401; size_t x_402; lean_object* x_403; uint8_t x_404; +x_386 = lean_ctor_get(x_300, 0); +x_387 = lean_ctor_get(x_300, 1); +lean_inc(x_387); +lean_inc(x_386); +lean_dec(x_300); +x_388 = l_Lean_instInhabitedFVarId; +x_389 = lean_array_get(x_388, x_298, x_13); +lean_dec(x_13); +lean_dec(x_298); +x_390 = lean_ctor_get(x_387, 0); +lean_inc(x_390); +x_391 = lean_ctor_get(x_387, 1); +lean_inc(x_391); +if (lean_is_exclusive(x_387)) { + lean_ctor_release(x_387, 0); + lean_ctor_release(x_387, 1); + x_392 = x_387; +} else { + lean_dec_ref(x_387); + x_392 = lean_box(0); +} +x_393 = lean_array_get_size(x_391); +x_394 = l___private_Lean_Expr_0__Lean_hashFVarId____x40_Lean_Expr___hyg_1730_(x_11); +x_395 = lean_uint64_shift_right(x_394, x_31); +x_396 = lean_uint64_xor(x_394, x_395); +x_397 = lean_uint64_shift_right(x_396, x_34); +x_398 = lean_uint64_xor(x_396, x_397); +x_399 = lean_uint64_to_usize(x_398); +x_400 = lean_usize_of_nat(x_393); +lean_dec(x_393); +x_401 = lean_usize_sub(x_400, x_39); +x_402 = lean_usize_land(x_399, x_401); +x_403 = lean_array_uget(x_391, x_402); +x_404 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__6(x_11, x_403); +if (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; lean_object* x_411; lean_object* x_412; lean_object* x_413; uint8_t x_414; +x_405 = lean_unsigned_to_nat(1u); +x_406 = lean_nat_add(x_390, x_405); +lean_dec(x_390); +if (lean_is_scalar(x_15)) { + x_407 = lean_alloc_ctor(1, 3, 0); +} else { + x_407 = x_15; + lean_ctor_set_tag(x_407, 1); +} +lean_ctor_set(x_407, 0, x_11); +lean_ctor_set(x_407, 1, x_389); +lean_ctor_set(x_407, 2, x_403); +x_408 = lean_array_uset(x_391, x_402, x_407); +x_409 = lean_unsigned_to_nat(4u); +x_410 = lean_nat_mul(x_406, x_409); +x_411 = lean_unsigned_to_nat(3u); +x_412 = lean_nat_div(x_410, x_411); +lean_dec(x_410); +x_413 = lean_array_get_size(x_408); +x_414 = lean_nat_dec_le(x_412, x_413); +lean_dec(x_413); +lean_dec(x_412); +if (x_414 == 0) +{ +lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; +x_415 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__7(x_408); +if (lean_is_scalar(x_392)) { + x_416 = lean_alloc_ctor(0, 2, 0); +} else { + x_416 = x_392; +} +lean_ctor_set(x_416, 0, x_406); +lean_ctor_set(x_416, 1, x_415); +x_417 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_417, 0, x_386); +lean_ctor_set(x_417, 1, x_416); +x_418 = lean_st_ref_set(x_2, x_417, x_301); +x_419 = lean_ctor_get(x_418, 1); +lean_inc(x_419); +lean_dec(x_418); +x_1 = x_10; +x_7 = x_419; +goto _start; +} +else +{ +lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; +if (lean_is_scalar(x_392)) { + x_421 = lean_alloc_ctor(0, 2, 0); +} else { + x_421 = x_392; +} +lean_ctor_set(x_421, 0, x_406); +lean_ctor_set(x_421, 1, x_408); +x_422 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_422, 0, x_386); +lean_ctor_set(x_422, 1, x_421); +x_423 = lean_st_ref_set(x_2, x_422, x_301); +x_424 = lean_ctor_get(x_423, 1); +lean_inc(x_424); +lean_dec(x_423); +x_1 = x_10; +x_7 = x_424; +goto _start; +} +} +else +{ +lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; +lean_dec(x_15); +x_426 = lean_box(0); +x_427 = lean_array_uset(x_391, x_402, x_426); +x_428 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__10(x_11, x_389, x_403); +x_429 = lean_array_uset(x_427, x_402, x_428); +if (lean_is_scalar(x_392)) { + x_430 = lean_alloc_ctor(0, 2, 0); +} else { + x_430 = x_392; +} +lean_ctor_set(x_430, 0, x_390); +lean_ctor_set(x_430, 1, x_429); +x_431 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_431, 0, x_386); +lean_ctor_set(x_431, 1, x_430); +x_432 = lean_st_ref_set(x_2, x_431, x_301); +x_433 = lean_ctor_get(x_432, 1); +lean_inc(x_433); +lean_dec(x_432); +x_1 = x_10; +x_7 = x_433; +goto _start; +} +} +} +} +else +{ +lean_object* x_435; lean_object* x_436; +x_435 = lean_ctor_get(x_1, 1); +lean_inc(x_435); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_436 = l_Lean_Compiler_LCNF_StructProjCases_visitLetValue(x_9, x_2, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_436) == 0) +{ +lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; +x_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); +lean_inc(x_8); +x_439 = l_Lean_Compiler_LCNF_LetDecl_updateValue(x_8, x_437, x_3, x_4, x_5, x_6, x_438); +x_440 = lean_ctor_get(x_439, 0); +lean_inc(x_440); +x_441 = lean_ctor_get(x_439, 1); +lean_inc(x_441); +lean_dec(x_439); +lean_inc(x_435); +x_442 = l_Lean_Compiler_LCNF_StructProjCases_visitCode(x_435, x_2, x_3, x_4, x_5, x_6, x_441); +if (lean_obj_tag(x_442) == 0) +{ +uint8_t x_443; +x_443 = !lean_is_exclusive(x_442); +if (x_443 == 0) +{ +lean_object* x_444; size_t x_445; size_t x_446; uint8_t x_447; +x_444 = lean_ctor_get(x_442, 0); +x_445 = lean_ptr_addr(x_435); +lean_dec(x_435); +x_446 = lean_ptr_addr(x_444); +x_447 = lean_usize_dec_eq(x_445, x_446); +if (x_447 == 0) +{ +uint8_t x_448; +lean_dec(x_8); +x_448 = !lean_is_exclusive(x_1); +if (x_448 == 0) +{ +lean_object* x_449; lean_object* x_450; +x_449 = lean_ctor_get(x_1, 1); +lean_dec(x_449); +x_450 = lean_ctor_get(x_1, 0); +lean_dec(x_450); +lean_ctor_set(x_1, 1, x_444); +lean_ctor_set(x_1, 0, x_440); +lean_ctor_set(x_442, 0, x_1); +return x_442; +} +else +{ +lean_object* x_451; +lean_dec(x_1); +x_451 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_451, 0, x_440); +lean_ctor_set(x_451, 1, x_444); +lean_ctor_set(x_442, 0, x_451); +return x_442; +} +} +else +{ +size_t x_452; size_t x_453; uint8_t x_454; +x_452 = lean_ptr_addr(x_8); +lean_dec(x_8); +x_453 = lean_ptr_addr(x_440); +x_454 = lean_usize_dec_eq(x_452, x_453); +if (x_454 == 0) +{ +uint8_t x_455; +x_455 = !lean_is_exclusive(x_1); +if (x_455 == 0) +{ +lean_object* x_456; lean_object* x_457; +x_456 = lean_ctor_get(x_1, 1); +lean_dec(x_456); +x_457 = lean_ctor_get(x_1, 0); +lean_dec(x_457); +lean_ctor_set(x_1, 1, x_444); +lean_ctor_set(x_1, 0, x_440); +lean_ctor_set(x_442, 0, x_1); +return x_442; +} +else +{ +lean_object* x_458; +lean_dec(x_1); +x_458 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_458, 0, x_440); +lean_ctor_set(x_458, 1, x_444); +lean_ctor_set(x_442, 0, x_458); +return x_442; +} +} +else +{ +lean_dec(x_444); +lean_dec(x_440); +lean_ctor_set(x_442, 0, x_1); +return x_442; +} +} +} +else +{ +lean_object* x_459; lean_object* x_460; size_t x_461; size_t x_462; uint8_t x_463; +x_459 = lean_ctor_get(x_442, 0); +x_460 = lean_ctor_get(x_442, 1); +lean_inc(x_460); +lean_inc(x_459); +lean_dec(x_442); +x_461 = lean_ptr_addr(x_435); +lean_dec(x_435); +x_462 = lean_ptr_addr(x_459); +x_463 = lean_usize_dec_eq(x_461, x_462); +if (x_463 == 0) +{ +lean_object* x_464; lean_object* x_465; lean_object* x_466; +lean_dec(x_8); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + x_464 = x_1; +} else { + lean_dec_ref(x_1); + x_464 = lean_box(0); +} +if (lean_is_scalar(x_464)) { + x_465 = lean_alloc_ctor(0, 2, 0); +} else { + x_465 = x_464; +} +lean_ctor_set(x_465, 0, x_440); +lean_ctor_set(x_465, 1, x_459); +x_466 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_466, 0, x_465); +lean_ctor_set(x_466, 1, x_460); +return x_466; +} +else +{ +size_t x_467; size_t x_468; uint8_t x_469; +x_467 = lean_ptr_addr(x_8); +lean_dec(x_8); +x_468 = lean_ptr_addr(x_440); +x_469 = lean_usize_dec_eq(x_467, x_468); +if (x_469 == 0) +{ +lean_object* x_470; lean_object* x_471; lean_object* x_472; +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + x_470 = x_1; +} else { + lean_dec_ref(x_1); + x_470 = lean_box(0); +} +if (lean_is_scalar(x_470)) { + x_471 = lean_alloc_ctor(0, 2, 0); +} else { + x_471 = x_470; +} +lean_ctor_set(x_471, 0, x_440); +lean_ctor_set(x_471, 1, x_459); +x_472 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_472, 0, x_471); +lean_ctor_set(x_472, 1, x_460); +return x_472; +} +else +{ +lean_object* x_473; +lean_dec(x_459); +lean_dec(x_440); +x_473 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_473, 0, x_1); +lean_ctor_set(x_473, 1, x_460); +return x_473; +} +} +} +} +else +{ +uint8_t x_474; +lean_dec(x_440); +lean_dec(x_435); +lean_dec(x_8); +lean_dec(x_1); +x_474 = !lean_is_exclusive(x_442); +if (x_474 == 0) +{ +return x_442; +} +else +{ +lean_object* x_475; lean_object* x_476; lean_object* x_477; +x_475 = lean_ctor_get(x_442, 0); +x_476 = lean_ctor_get(x_442, 1); +lean_inc(x_476); +lean_inc(x_475); +lean_dec(x_442); +x_477 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_477, 0, x_475); +lean_ctor_set(x_477, 1, x_476); +return x_477; +} +} +} +else +{ +uint8_t x_478; +lean_dec(x_435); +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_478 = !lean_is_exclusive(x_436); +if (x_478 == 0) +{ +return x_436; +} +else +{ +lean_object* x_479; lean_object* x_480; lean_object* x_481; +x_479 = lean_ctor_get(x_436, 0); +x_480 = lean_ctor_get(x_436, 1); +lean_inc(x_480); +lean_inc(x_479); +lean_dec(x_436); +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; +} +} +} +} +case 1: +{ +lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; +x_482 = lean_ctor_get(x_1, 0); +lean_inc(x_482); +x_483 = lean_ctor_get(x_1, 1); +lean_inc(x_483); +x_484 = lean_ctor_get(x_482, 4); +lean_inc(x_484); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_485 = l_Lean_Compiler_LCNF_StructProjCases_visitCode(x_484, x_2, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_485) == 0) +{ +lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; +x_486 = lean_ctor_get(x_485, 0); +lean_inc(x_486); +x_487 = lean_ctor_get(x_485, 1); +lean_inc(x_487); +lean_dec(x_485); +x_488 = lean_ctor_get(x_482, 3); +lean_inc(x_488); +x_489 = lean_ctor_get(x_482, 2); +lean_inc(x_489); +lean_inc(x_482); +x_490 = l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_updateFunDeclImp(x_482, x_488, x_489, x_486, x_3, x_4, x_5, x_6, x_487); +x_491 = lean_ctor_get(x_490, 0); +lean_inc(x_491); +x_492 = lean_ctor_get(x_490, 1); +lean_inc(x_492); +lean_dec(x_490); +lean_inc(x_483); +x_493 = l_Lean_Compiler_LCNF_StructProjCases_visitCode(x_483, x_2, x_3, x_4, x_5, x_6, x_492); +if (lean_obj_tag(x_493) == 0) +{ +uint8_t x_494; +x_494 = !lean_is_exclusive(x_493); +if (x_494 == 0) +{ +lean_object* x_495; size_t x_496; size_t x_497; uint8_t x_498; +x_495 = lean_ctor_get(x_493, 0); +x_496 = lean_ptr_addr(x_483); +lean_dec(x_483); +x_497 = lean_ptr_addr(x_495); +x_498 = lean_usize_dec_eq(x_496, x_497); +if (x_498 == 0) +{ +uint8_t x_499; +lean_dec(x_482); +x_499 = !lean_is_exclusive(x_1); +if (x_499 == 0) +{ +lean_object* x_500; lean_object* x_501; +x_500 = lean_ctor_get(x_1, 1); +lean_dec(x_500); +x_501 = lean_ctor_get(x_1, 0); +lean_dec(x_501); +lean_ctor_set(x_1, 1, x_495); +lean_ctor_set(x_1, 0, x_491); +lean_ctor_set(x_493, 0, x_1); +return x_493; +} +else +{ +lean_object* x_502; +lean_dec(x_1); +x_502 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_502, 0, x_491); +lean_ctor_set(x_502, 1, x_495); +lean_ctor_set(x_493, 0, x_502); +return x_493; +} +} +else +{ +size_t x_503; size_t x_504; uint8_t x_505; +x_503 = lean_ptr_addr(x_482); +lean_dec(x_482); +x_504 = lean_ptr_addr(x_491); +x_505 = lean_usize_dec_eq(x_503, x_504); +if (x_505 == 0) +{ +uint8_t x_506; +x_506 = !lean_is_exclusive(x_1); +if (x_506 == 0) +{ +lean_object* x_507; lean_object* x_508; +x_507 = lean_ctor_get(x_1, 1); +lean_dec(x_507); +x_508 = lean_ctor_get(x_1, 0); +lean_dec(x_508); +lean_ctor_set(x_1, 1, x_495); +lean_ctor_set(x_1, 0, x_491); +lean_ctor_set(x_493, 0, x_1); +return x_493; +} +else +{ +lean_object* x_509; +lean_dec(x_1); +x_509 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_509, 0, x_491); +lean_ctor_set(x_509, 1, x_495); +lean_ctor_set(x_493, 0, x_509); +return x_493; +} +} +else +{ +lean_dec(x_495); +lean_dec(x_491); +lean_ctor_set(x_493, 0, x_1); +return x_493; +} +} +} +else +{ +lean_object* x_510; lean_object* x_511; size_t x_512; size_t x_513; uint8_t x_514; +x_510 = lean_ctor_get(x_493, 0); +x_511 = lean_ctor_get(x_493, 1); +lean_inc(x_511); +lean_inc(x_510); +lean_dec(x_493); +x_512 = lean_ptr_addr(x_483); +lean_dec(x_483); +x_513 = lean_ptr_addr(x_510); +x_514 = lean_usize_dec_eq(x_512, x_513); +if (x_514 == 0) +{ +lean_object* x_515; lean_object* x_516; lean_object* x_517; +lean_dec(x_482); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + x_515 = x_1; +} else { + lean_dec_ref(x_1); + x_515 = lean_box(0); +} +if (lean_is_scalar(x_515)) { + x_516 = lean_alloc_ctor(1, 2, 0); +} else { + x_516 = x_515; +} +lean_ctor_set(x_516, 0, x_491); +lean_ctor_set(x_516, 1, x_510); +x_517 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_517, 0, x_516); +lean_ctor_set(x_517, 1, x_511); +return x_517; +} +else +{ +size_t x_518; size_t x_519; uint8_t x_520; +x_518 = lean_ptr_addr(x_482); +lean_dec(x_482); +x_519 = lean_ptr_addr(x_491); +x_520 = lean_usize_dec_eq(x_518, x_519); +if (x_520 == 0) +{ +lean_object* x_521; lean_object* x_522; lean_object* x_523; +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + x_521 = x_1; +} else { + lean_dec_ref(x_1); + x_521 = lean_box(0); +} +if (lean_is_scalar(x_521)) { + x_522 = lean_alloc_ctor(1, 2, 0); +} else { + x_522 = x_521; +} +lean_ctor_set(x_522, 0, x_491); +lean_ctor_set(x_522, 1, x_510); +x_523 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_523, 0, x_522); +lean_ctor_set(x_523, 1, x_511); +return x_523; +} +else +{ +lean_object* x_524; +lean_dec(x_510); +lean_dec(x_491); +x_524 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_524, 0, x_1); +lean_ctor_set(x_524, 1, x_511); +return x_524; +} +} +} +} +else +{ +uint8_t x_525; +lean_dec(x_491); +lean_dec(x_483); +lean_dec(x_482); +lean_dec(x_1); +x_525 = !lean_is_exclusive(x_493); +if (x_525 == 0) +{ +return x_493; +} +else +{ +lean_object* x_526; lean_object* x_527; lean_object* x_528; +x_526 = lean_ctor_get(x_493, 0); +x_527 = lean_ctor_get(x_493, 1); +lean_inc(x_527); +lean_inc(x_526); +lean_dec(x_493); +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_483); +lean_dec(x_482); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_529 = !lean_is_exclusive(x_485); +if (x_529 == 0) +{ +return x_485; +} +else +{ +lean_object* x_530; lean_object* x_531; lean_object* x_532; +x_530 = lean_ctor_get(x_485, 0); +x_531 = lean_ctor_get(x_485, 1); +lean_inc(x_531); +lean_inc(x_530); +lean_dec(x_485); +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; +} +} +} +case 2: +{ +lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; +x_533 = lean_ctor_get(x_1, 0); +lean_inc(x_533); +x_534 = lean_ctor_get(x_1, 1); +lean_inc(x_534); +x_535 = lean_ctor_get(x_533, 4); +lean_inc(x_535); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_536 = l_Lean_Compiler_LCNF_StructProjCases_visitCode(x_535, x_2, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_536) == 0) +{ +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; +x_537 = lean_ctor_get(x_536, 0); +lean_inc(x_537); +x_538 = lean_ctor_get(x_536, 1); +lean_inc(x_538); +lean_dec(x_536); +x_539 = lean_ctor_get(x_533, 3); +lean_inc(x_539); +x_540 = lean_ctor_get(x_533, 2); +lean_inc(x_540); +lean_inc(x_533); +x_541 = l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_updateFunDeclImp(x_533, x_539, x_540, x_537, x_3, x_4, x_5, x_6, 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); +lean_inc(x_534); +x_544 = l_Lean_Compiler_LCNF_StructProjCases_visitCode(x_534, x_2, x_3, x_4, x_5, x_6, x_543); +if (lean_obj_tag(x_544) == 0) +{ +uint8_t x_545; +x_545 = !lean_is_exclusive(x_544); +if (x_545 == 0) +{ +lean_object* x_546; size_t x_547; size_t x_548; uint8_t x_549; +x_546 = lean_ctor_get(x_544, 0); +x_547 = lean_ptr_addr(x_534); +lean_dec(x_534); +x_548 = lean_ptr_addr(x_546); +x_549 = lean_usize_dec_eq(x_547, x_548); +if (x_549 == 0) +{ +uint8_t x_550; +lean_dec(x_533); +x_550 = !lean_is_exclusive(x_1); +if (x_550 == 0) +{ +lean_object* x_551; lean_object* x_552; +x_551 = lean_ctor_get(x_1, 1); +lean_dec(x_551); +x_552 = lean_ctor_get(x_1, 0); +lean_dec(x_552); +lean_ctor_set(x_1, 1, x_546); +lean_ctor_set(x_1, 0, x_542); +lean_ctor_set(x_544, 0, x_1); +return x_544; +} +else +{ +lean_object* x_553; +lean_dec(x_1); +x_553 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_553, 0, x_542); +lean_ctor_set(x_553, 1, x_546); +lean_ctor_set(x_544, 0, x_553); +return x_544; +} +} +else +{ +size_t x_554; size_t x_555; uint8_t x_556; +x_554 = lean_ptr_addr(x_533); +lean_dec(x_533); +x_555 = lean_ptr_addr(x_542); +x_556 = lean_usize_dec_eq(x_554, x_555); +if (x_556 == 0) +{ +uint8_t x_557; +x_557 = !lean_is_exclusive(x_1); +if (x_557 == 0) +{ +lean_object* x_558; lean_object* x_559; +x_558 = lean_ctor_get(x_1, 1); +lean_dec(x_558); +x_559 = lean_ctor_get(x_1, 0); +lean_dec(x_559); +lean_ctor_set(x_1, 1, x_546); +lean_ctor_set(x_1, 0, x_542); +lean_ctor_set(x_544, 0, x_1); +return x_544; +} +else +{ +lean_object* x_560; +lean_dec(x_1); +x_560 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_560, 0, x_542); +lean_ctor_set(x_560, 1, x_546); +lean_ctor_set(x_544, 0, x_560); +return x_544; +} +} +else +{ +lean_dec(x_546); +lean_dec(x_542); +lean_ctor_set(x_544, 0, x_1); +return x_544; +} +} +} +else +{ +lean_object* x_561; lean_object* x_562; size_t x_563; size_t x_564; uint8_t x_565; +x_561 = lean_ctor_get(x_544, 0); +x_562 = lean_ctor_get(x_544, 1); +lean_inc(x_562); +lean_inc(x_561); +lean_dec(x_544); +x_563 = lean_ptr_addr(x_534); +lean_dec(x_534); +x_564 = lean_ptr_addr(x_561); +x_565 = lean_usize_dec_eq(x_563, x_564); +if (x_565 == 0) +{ +lean_object* x_566; lean_object* x_567; lean_object* x_568; +lean_dec(x_533); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + x_566 = x_1; +} else { + lean_dec_ref(x_1); + x_566 = lean_box(0); +} +if (lean_is_scalar(x_566)) { + x_567 = lean_alloc_ctor(2, 2, 0); +} else { + x_567 = x_566; +} +lean_ctor_set(x_567, 0, x_542); +lean_ctor_set(x_567, 1, x_561); +x_568 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_568, 0, x_567); +lean_ctor_set(x_568, 1, x_562); +return x_568; +} +else +{ +size_t x_569; size_t x_570; uint8_t x_571; +x_569 = lean_ptr_addr(x_533); +lean_dec(x_533); +x_570 = lean_ptr_addr(x_542); +x_571 = lean_usize_dec_eq(x_569, x_570); +if (x_571 == 0) +{ +lean_object* x_572; lean_object* x_573; lean_object* x_574; +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + x_572 = x_1; +} else { + lean_dec_ref(x_1); + x_572 = lean_box(0); +} +if (lean_is_scalar(x_572)) { + x_573 = lean_alloc_ctor(2, 2, 0); +} else { + x_573 = x_572; +} +lean_ctor_set(x_573, 0, x_542); +lean_ctor_set(x_573, 1, x_561); +x_574 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_574, 0, x_573); +lean_ctor_set(x_574, 1, x_562); +return x_574; +} +else +{ +lean_object* x_575; +lean_dec(x_561); +lean_dec(x_542); +x_575 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_575, 0, x_1); +lean_ctor_set(x_575, 1, x_562); +return x_575; +} +} +} +} +else +{ +uint8_t x_576; +lean_dec(x_542); +lean_dec(x_534); +lean_dec(x_533); +lean_dec(x_1); +x_576 = !lean_is_exclusive(x_544); +if (x_576 == 0) +{ +return x_544; +} +else +{ +lean_object* x_577; lean_object* x_578; lean_object* x_579; +x_577 = lean_ctor_get(x_544, 0); +x_578 = lean_ctor_get(x_544, 1); +lean_inc(x_578); +lean_inc(x_577); +lean_dec(x_544); +x_579 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_579, 0, x_577); +lean_ctor_set(x_579, 1, x_578); +return x_579; +} +} +} +else +{ +uint8_t x_580; +lean_dec(x_534); +lean_dec(x_533); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_580 = !lean_is_exclusive(x_536); +if (x_580 == 0) +{ +return x_536; +} +else +{ +lean_object* x_581; lean_object* x_582; lean_object* x_583; +x_581 = lean_ctor_get(x_536, 0); +x_582 = lean_ctor_get(x_536, 1); +lean_inc(x_582); +lean_inc(x_581); +lean_dec(x_536); +x_583 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_583, 0, x_581); +lean_ctor_set(x_583, 1, x_582); +return x_583; +} +} +} +case 3: +{ +lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; size_t x_589; size_t x_590; lean_object* x_591; uint8_t x_592; +x_584 = lean_ctor_get(x_1, 0); +lean_inc(x_584); +x_585 = lean_ctor_get(x_1, 1); +lean_inc(x_585); +lean_inc(x_584); +x_586 = l_Lean_Compiler_LCNF_StructProjCases_remapFVar(x_584, x_2, x_3, x_4, x_5, x_6, x_7); +x_587 = lean_ctor_get(x_586, 0); +lean_inc(x_587); +x_588 = lean_ctor_get(x_586, 1); +lean_inc(x_588); +lean_dec(x_586); +x_589 = lean_array_size(x_585); +x_590 = 0; +lean_inc(x_585); +x_591 = l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_StructProjCases_visitLetValue___spec__2(x_589, x_590, x_585, x_2, x_3, x_4, x_5, x_6, x_588); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_592 = !lean_is_exclusive(x_591); +if (x_592 == 0) +{ +lean_object* x_593; uint8_t x_594; +x_593 = lean_ctor_get(x_591, 0); +x_594 = lean_name_eq(x_584, x_587); +lean_dec(x_584); +if (x_594 == 0) +{ +uint8_t x_595; +lean_dec(x_585); +x_595 = !lean_is_exclusive(x_1); +if (x_595 == 0) +{ +lean_object* x_596; lean_object* x_597; +x_596 = lean_ctor_get(x_1, 1); +lean_dec(x_596); +x_597 = lean_ctor_get(x_1, 0); +lean_dec(x_597); +lean_ctor_set(x_1, 1, x_593); +lean_ctor_set(x_1, 0, x_587); +lean_ctor_set(x_591, 0, x_1); +return x_591; +} +else +{ +lean_object* x_598; +lean_dec(x_1); +x_598 = lean_alloc_ctor(3, 2, 0); +lean_ctor_set(x_598, 0, x_587); +lean_ctor_set(x_598, 1, x_593); +lean_ctor_set(x_591, 0, x_598); +return x_591; +} +} +else +{ +size_t x_599; size_t x_600; uint8_t x_601; +x_599 = lean_ptr_addr(x_585); +lean_dec(x_585); +x_600 = lean_ptr_addr(x_593); +x_601 = lean_usize_dec_eq(x_599, x_600); +if (x_601 == 0) +{ +uint8_t x_602; +x_602 = !lean_is_exclusive(x_1); +if (x_602 == 0) +{ +lean_object* x_603; lean_object* x_604; +x_603 = lean_ctor_get(x_1, 1); +lean_dec(x_603); +x_604 = lean_ctor_get(x_1, 0); +lean_dec(x_604); +lean_ctor_set(x_1, 1, x_593); +lean_ctor_set(x_1, 0, x_587); +lean_ctor_set(x_591, 0, x_1); +return x_591; +} +else +{ +lean_object* x_605; +lean_dec(x_1); +x_605 = lean_alloc_ctor(3, 2, 0); +lean_ctor_set(x_605, 0, x_587); +lean_ctor_set(x_605, 1, x_593); +lean_ctor_set(x_591, 0, x_605); +return x_591; +} +} +else +{ +lean_dec(x_593); +lean_dec(x_587); +lean_ctor_set(x_591, 0, x_1); +return x_591; +} +} +} +else +{ +lean_object* x_606; lean_object* x_607; uint8_t x_608; +x_606 = lean_ctor_get(x_591, 0); +x_607 = lean_ctor_get(x_591, 1); +lean_inc(x_607); +lean_inc(x_606); +lean_dec(x_591); +x_608 = lean_name_eq(x_584, x_587); +lean_dec(x_584); +if (x_608 == 0) +{ +lean_object* x_609; lean_object* x_610; lean_object* x_611; +lean_dec(x_585); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + x_609 = x_1; +} else { + lean_dec_ref(x_1); + x_609 = lean_box(0); +} +if (lean_is_scalar(x_609)) { + x_610 = lean_alloc_ctor(3, 2, 0); +} else { + x_610 = x_609; +} +lean_ctor_set(x_610, 0, x_587); +lean_ctor_set(x_610, 1, x_606); +x_611 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_611, 0, x_610); +lean_ctor_set(x_611, 1, x_607); +return x_611; +} +else +{ +size_t x_612; size_t x_613; uint8_t x_614; +x_612 = lean_ptr_addr(x_585); +lean_dec(x_585); +x_613 = lean_ptr_addr(x_606); +x_614 = lean_usize_dec_eq(x_612, x_613); +if (x_614 == 0) +{ +lean_object* x_615; lean_object* x_616; lean_object* x_617; +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + x_615 = x_1; +} else { + lean_dec_ref(x_1); + x_615 = lean_box(0); +} +if (lean_is_scalar(x_615)) { + x_616 = lean_alloc_ctor(3, 2, 0); +} else { + x_616 = x_615; +} +lean_ctor_set(x_616, 0, x_587); +lean_ctor_set(x_616, 1, x_606); +x_617 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_617, 0, x_616); +lean_ctor_set(x_617, 1, x_607); +return x_617; +} +else +{ +lean_object* x_618; +lean_dec(x_606); +lean_dec(x_587); +x_618 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_618, 0, x_1); +lean_ctor_set(x_618, 1, x_607); +return x_618; +} +} +} +} +case 4: +{ +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; uint8_t x_630; +x_619 = lean_ctor_get(x_1, 0); +lean_inc(x_619); +x_620 = lean_ctor_get(x_619, 0); +lean_inc(x_620); +x_621 = lean_ctor_get(x_619, 1); +lean_inc(x_621); +x_622 = lean_ctor_get(x_619, 2); +lean_inc(x_622); +x_623 = lean_ctor_get(x_619, 3); +lean_inc(x_623); +if (lean_is_exclusive(x_619)) { + lean_ctor_release(x_619, 0); + lean_ctor_release(x_619, 1); + lean_ctor_release(x_619, 2); + lean_ctor_release(x_619, 3); + x_624 = x_619; +} else { + lean_dec_ref(x_619); + x_624 = lean_box(0); +} +lean_inc(x_622); +x_625 = l_Lean_Compiler_LCNF_StructProjCases_remapFVar(x_622, x_2, x_3, x_4, x_5, x_6, x_7); +x_626 = lean_ctor_get(x_625, 0); +lean_inc(x_626); +x_627 = lean_ctor_get(x_625, 1); +lean_inc(x_627); +lean_dec(x_625); +x_628 = lean_array_get_size(x_623); +x_629 = lean_unsigned_to_nat(1u); +x_630 = lean_nat_dec_eq(x_628, x_629); +lean_dec(x_628); +if (x_630 == 0) +{ +size_t x_631; size_t x_632; lean_object* x_633; +x_631 = lean_array_size(x_623); +x_632 = 0; +lean_inc(x_623); +x_633 = l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__15(x_631, x_632, x_623, x_2, x_3, x_4, x_5, x_6, x_627); +if (lean_obj_tag(x_633) == 0) +{ +uint8_t x_634; +x_634 = !lean_is_exclusive(x_633); +if (x_634 == 0) +{ +lean_object* x_635; size_t x_636; size_t x_637; uint8_t x_638; +x_635 = lean_ctor_get(x_633, 0); +x_636 = lean_ptr_addr(x_623); +lean_dec(x_623); +x_637 = lean_ptr_addr(x_635); +x_638 = lean_usize_dec_eq(x_636, x_637); +if (x_638 == 0) +{ +uint8_t x_639; +lean_dec(x_622); +x_639 = !lean_is_exclusive(x_1); +if (x_639 == 0) +{ +lean_object* x_640; lean_object* x_641; +x_640 = lean_ctor_get(x_1, 0); +lean_dec(x_640); +if (lean_is_scalar(x_624)) { + x_641 = lean_alloc_ctor(0, 4, 0); +} else { + x_641 = x_624; +} +lean_ctor_set(x_641, 0, x_620); +lean_ctor_set(x_641, 1, x_621); +lean_ctor_set(x_641, 2, x_626); +lean_ctor_set(x_641, 3, x_635); +lean_ctor_set(x_1, 0, x_641); +lean_ctor_set(x_633, 0, x_1); +return x_633; +} +else +{ +lean_object* x_642; lean_object* x_643; +lean_dec(x_1); +if (lean_is_scalar(x_624)) { + x_642 = lean_alloc_ctor(0, 4, 0); +} else { + x_642 = x_624; +} +lean_ctor_set(x_642, 0, x_620); +lean_ctor_set(x_642, 1, x_621); +lean_ctor_set(x_642, 2, x_626); +lean_ctor_set(x_642, 3, x_635); +x_643 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_643, 0, x_642); +lean_ctor_set(x_633, 0, x_643); +return x_633; +} +} +else +{ +size_t x_644; uint8_t x_645; +x_644 = lean_ptr_addr(x_621); +x_645 = lean_usize_dec_eq(x_644, x_644); +if (x_645 == 0) +{ +uint8_t x_646; +lean_dec(x_622); +x_646 = !lean_is_exclusive(x_1); +if (x_646 == 0) +{ +lean_object* x_647; lean_object* x_648; +x_647 = lean_ctor_get(x_1, 0); +lean_dec(x_647); +if (lean_is_scalar(x_624)) { + x_648 = lean_alloc_ctor(0, 4, 0); +} else { + x_648 = x_624; +} +lean_ctor_set(x_648, 0, x_620); +lean_ctor_set(x_648, 1, x_621); +lean_ctor_set(x_648, 2, x_626); +lean_ctor_set(x_648, 3, x_635); +lean_ctor_set(x_1, 0, x_648); +lean_ctor_set(x_633, 0, x_1); +return x_633; +} +else +{ +lean_object* x_649; lean_object* x_650; +lean_dec(x_1); +if (lean_is_scalar(x_624)) { + x_649 = lean_alloc_ctor(0, 4, 0); +} else { + x_649 = x_624; +} +lean_ctor_set(x_649, 0, x_620); +lean_ctor_set(x_649, 1, x_621); +lean_ctor_set(x_649, 2, x_626); +lean_ctor_set(x_649, 3, x_635); +x_650 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_650, 0, x_649); +lean_ctor_set(x_633, 0, x_650); +return x_633; +} +} +else +{ +uint8_t x_651; +x_651 = lean_name_eq(x_622, x_626); +lean_dec(x_622); +if (x_651 == 0) +{ +uint8_t x_652; +x_652 = !lean_is_exclusive(x_1); +if (x_652 == 0) +{ +lean_object* x_653; lean_object* x_654; +x_653 = lean_ctor_get(x_1, 0); +lean_dec(x_653); +if (lean_is_scalar(x_624)) { + x_654 = lean_alloc_ctor(0, 4, 0); +} else { + x_654 = x_624; +} +lean_ctor_set(x_654, 0, x_620); +lean_ctor_set(x_654, 1, x_621); +lean_ctor_set(x_654, 2, x_626); +lean_ctor_set(x_654, 3, x_635); +lean_ctor_set(x_1, 0, x_654); +lean_ctor_set(x_633, 0, x_1); +return x_633; +} +else +{ +lean_object* x_655; lean_object* x_656; +lean_dec(x_1); +if (lean_is_scalar(x_624)) { + x_655 = lean_alloc_ctor(0, 4, 0); +} else { + x_655 = x_624; +} +lean_ctor_set(x_655, 0, x_620); +lean_ctor_set(x_655, 1, x_621); +lean_ctor_set(x_655, 2, x_626); +lean_ctor_set(x_655, 3, x_635); +x_656 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_656, 0, x_655); +lean_ctor_set(x_633, 0, x_656); +return x_633; +} +} +else +{ +lean_dec(x_635); +lean_dec(x_626); +lean_dec(x_624); +lean_dec(x_621); +lean_dec(x_620); +lean_ctor_set(x_633, 0, x_1); +return x_633; +} +} +} +} +else +{ +lean_object* x_657; lean_object* x_658; size_t x_659; size_t x_660; uint8_t x_661; +x_657 = lean_ctor_get(x_633, 0); +x_658 = lean_ctor_get(x_633, 1); +lean_inc(x_658); +lean_inc(x_657); +lean_dec(x_633); +x_659 = lean_ptr_addr(x_623); +lean_dec(x_623); +x_660 = lean_ptr_addr(x_657); +x_661 = lean_usize_dec_eq(x_659, x_660); +if (x_661 == 0) +{ +lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; +lean_dec(x_622); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + x_662 = x_1; +} else { + lean_dec_ref(x_1); + x_662 = lean_box(0); +} +if (lean_is_scalar(x_624)) { + x_663 = lean_alloc_ctor(0, 4, 0); +} else { + x_663 = x_624; +} +lean_ctor_set(x_663, 0, x_620); +lean_ctor_set(x_663, 1, x_621); +lean_ctor_set(x_663, 2, x_626); +lean_ctor_set(x_663, 3, x_657); +if (lean_is_scalar(x_662)) { + x_664 = lean_alloc_ctor(4, 1, 0); +} else { + x_664 = x_662; +} +lean_ctor_set(x_664, 0, x_663); +x_665 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_665, 0, x_664); +lean_ctor_set(x_665, 1, x_658); +return x_665; +} +else +{ +size_t x_666; uint8_t x_667; +x_666 = lean_ptr_addr(x_621); +x_667 = lean_usize_dec_eq(x_666, x_666); +if (x_667 == 0) +{ +lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; +lean_dec(x_622); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + x_668 = x_1; +} else { + lean_dec_ref(x_1); + x_668 = lean_box(0); +} +if (lean_is_scalar(x_624)) { + x_669 = lean_alloc_ctor(0, 4, 0); +} else { + x_669 = x_624; +} +lean_ctor_set(x_669, 0, x_620); +lean_ctor_set(x_669, 1, x_621); +lean_ctor_set(x_669, 2, x_626); +lean_ctor_set(x_669, 3, x_657); +if (lean_is_scalar(x_668)) { + x_670 = lean_alloc_ctor(4, 1, 0); +} else { + x_670 = x_668; +} +lean_ctor_set(x_670, 0, x_669); +x_671 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_671, 0, x_670); +lean_ctor_set(x_671, 1, x_658); +return x_671; +} +else +{ +uint8_t x_672; +x_672 = lean_name_eq(x_622, x_626); +lean_dec(x_622); +if (x_672 == 0) +{ +lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + x_673 = x_1; +} else { + lean_dec_ref(x_1); + x_673 = lean_box(0); +} +if (lean_is_scalar(x_624)) { + x_674 = lean_alloc_ctor(0, 4, 0); +} else { + x_674 = x_624; +} +lean_ctor_set(x_674, 0, x_620); +lean_ctor_set(x_674, 1, x_621); +lean_ctor_set(x_674, 2, x_626); +lean_ctor_set(x_674, 3, x_657); +if (lean_is_scalar(x_673)) { + x_675 = lean_alloc_ctor(4, 1, 0); +} else { + x_675 = x_673; +} +lean_ctor_set(x_675, 0, x_674); +x_676 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_676, 0, x_675); +lean_ctor_set(x_676, 1, x_658); +return x_676; +} +else +{ +lean_object* x_677; +lean_dec(x_657); +lean_dec(x_626); +lean_dec(x_624); +lean_dec(x_621); +lean_dec(x_620); +x_677 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_677, 0, x_1); +lean_ctor_set(x_677, 1, x_658); +return x_677; +} +} +} +} +} +else +{ +uint8_t x_678; +lean_dec(x_626); +lean_dec(x_624); +lean_dec(x_623); +lean_dec(x_622); +lean_dec(x_621); +lean_dec(x_620); +lean_dec(x_1); +x_678 = !lean_is_exclusive(x_633); +if (x_678 == 0) +{ +return x_633; +} +else +{ +lean_object* x_679; lean_object* x_680; lean_object* x_681; +x_679 = lean_ctor_get(x_633, 0); +x_680 = lean_ctor_get(x_633, 1); +lean_inc(x_680); +lean_inc(x_679); +lean_dec(x_633); +x_681 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_681, 0, x_679); +lean_ctor_set(x_681, 1, x_680); +return x_681; +} +} +} +else +{ +lean_object* x_682; lean_object* x_683; +x_682 = lean_unsigned_to_nat(0u); +x_683 = lean_array_fget(x_623, x_682); +if (lean_obj_tag(x_683) == 0) +{ +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; 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; size_t x_702; size_t x_703; size_t x_704; size_t x_705; size_t x_706; lean_object* x_707; lean_object* x_708; +x_684 = lean_ctor_get(x_683, 0); +lean_inc(x_684); +x_685 = lean_ctor_get(x_683, 1); +lean_inc(x_685); +x_686 = lean_ctor_get(x_683, 2); +lean_inc(x_686); +if (lean_is_exclusive(x_683)) { + lean_ctor_release(x_683, 0); + lean_ctor_release(x_683, 1); + lean_ctor_release(x_683, 2); + x_687 = x_683; +} else { + lean_dec_ref(x_683); + x_687 = lean_box(0); +} +x_688 = lean_st_ref_get(x_2, x_627); +x_689 = lean_ctor_get(x_688, 0); +lean_inc(x_689); +x_690 = lean_ctor_get(x_689, 0); +lean_inc(x_690); +lean_dec(x_689); +x_691 = lean_ctor_get(x_688, 1); +lean_inc(x_691); +lean_dec(x_688); +x_692 = lean_ctor_get(x_690, 1); +lean_inc(x_692); +if (lean_is_exclusive(x_690)) { + lean_ctor_release(x_690, 0); + lean_ctor_release(x_690, 1); + x_693 = x_690; +} else { + lean_dec_ref(x_690); + x_693 = lean_box(0); +} +x_694 = lean_array_get_size(x_692); +x_695 = l___private_Lean_Expr_0__Lean_hashFVarId____x40_Lean_Expr___hyg_1730_(x_626); +x_696 = 32; +x_697 = lean_uint64_shift_right(x_695, x_696); +x_698 = lean_uint64_xor(x_695, x_697); +x_699 = 16; +x_700 = lean_uint64_shift_right(x_698, x_699); +x_701 = lean_uint64_xor(x_698, x_700); +x_702 = lean_uint64_to_usize(x_701); +x_703 = lean_usize_of_nat(x_694); +lean_dec(x_694); +x_704 = 1; +x_705 = lean_usize_sub(x_703, x_704); +x_706 = lean_usize_land(x_702, x_705); +x_707 = lean_array_uget(x_692, x_706); +lean_dec(x_692); +x_708 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__1(x_626, x_707); +lean_dec(x_707); +if (lean_obj_tag(x_708) == 0) +{ +size_t x_709; size_t x_710; lean_object* x_711; lean_object* x_712; lean_object* x_822; lean_object* x_823; lean_object* x_824; lean_object* x_825; uint8_t x_826; +x_709 = lean_array_size(x_685); +x_710 = 0; +lean_inc(x_685); +x_711 = l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__3(x_709, x_710, x_685); +x_822 = lean_st_ref_take(x_2, x_691); +x_823 = lean_ctor_get(x_822, 0); +lean_inc(x_823); +x_824 = lean_ctor_get(x_823, 0); +lean_inc(x_824); +x_825 = lean_ctor_get(x_822, 1); +lean_inc(x_825); +lean_dec(x_822); +x_826 = !lean_is_exclusive(x_823); +if (x_826 == 0) +{ +lean_object* x_827; uint8_t x_828; +x_827 = lean_ctor_get(x_823, 0); +lean_dec(x_827); +x_828 = !lean_is_exclusive(x_824); +if (x_828 == 0) +{ +lean_object* x_829; lean_object* x_830; lean_object* x_831; size_t x_832; size_t x_833; size_t x_834; lean_object* x_835; uint8_t x_836; +x_829 = lean_ctor_get(x_824, 0); +x_830 = lean_ctor_get(x_824, 1); +x_831 = lean_array_get_size(x_830); +x_832 = lean_usize_of_nat(x_831); +lean_dec(x_831); +x_833 = lean_usize_sub(x_832, x_704); +x_834 = lean_usize_land(x_702, x_833); +x_835 = lean_array_uget(x_830, x_834); +x_836 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__4(x_626, x_835); +if (x_836 == 0) +{ +lean_object* x_837; lean_object* x_838; lean_object* x_839; lean_object* x_840; lean_object* x_841; lean_object* x_842; lean_object* x_843; lean_object* x_844; uint8_t x_845; +x_837 = lean_nat_add(x_829, x_629); +lean_dec(x_829); +lean_inc(x_626); +x_838 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_838, 0, x_626); +lean_ctor_set(x_838, 1, x_711); +lean_ctor_set(x_838, 2, x_835); +x_839 = lean_array_uset(x_830, x_834, x_838); +x_840 = lean_unsigned_to_nat(4u); +x_841 = lean_nat_mul(x_837, x_840); +x_842 = lean_unsigned_to_nat(3u); +x_843 = lean_nat_div(x_841, x_842); +lean_dec(x_841); +x_844 = lean_array_get_size(x_839); +x_845 = lean_nat_dec_le(x_843, x_844); +lean_dec(x_844); +lean_dec(x_843); +if (x_845 == 0) +{ +lean_object* x_846; lean_object* x_847; lean_object* x_848; +x_846 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__11(x_839); +lean_ctor_set(x_824, 1, x_846); +lean_ctor_set(x_824, 0, x_837); +x_847 = lean_st_ref_set(x_2, x_823, x_825); +x_848 = lean_ctor_get(x_847, 1); +lean_inc(x_848); +lean_dec(x_847); +x_712 = x_848; +goto block_821; +} +else +{ +lean_object* x_849; lean_object* x_850; +lean_ctor_set(x_824, 1, x_839); +lean_ctor_set(x_824, 0, x_837); +x_849 = lean_st_ref_set(x_2, x_823, x_825); +x_850 = lean_ctor_get(x_849, 1); +lean_inc(x_850); +lean_dec(x_849); +x_712 = x_850; +goto block_821; +} +} +else +{ +lean_object* x_851; lean_object* x_852; lean_object* x_853; lean_object* x_854; lean_object* x_855; lean_object* x_856; +x_851 = lean_box(0); +x_852 = lean_array_uset(x_830, x_834, x_851); +lean_inc(x_626); +x_853 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__14(x_626, x_711, x_835); +x_854 = lean_array_uset(x_852, x_834, x_853); +lean_ctor_set(x_824, 1, x_854); +x_855 = lean_st_ref_set(x_2, x_823, x_825); +x_856 = lean_ctor_get(x_855, 1); +lean_inc(x_856); +lean_dec(x_855); +x_712 = x_856; +goto block_821; +} +} +else +{ +lean_object* x_857; lean_object* x_858; lean_object* x_859; size_t x_860; size_t x_861; size_t x_862; lean_object* x_863; uint8_t x_864; +x_857 = lean_ctor_get(x_824, 0); +x_858 = lean_ctor_get(x_824, 1); +lean_inc(x_858); +lean_inc(x_857); +lean_dec(x_824); +x_859 = lean_array_get_size(x_858); +x_860 = lean_usize_of_nat(x_859); +lean_dec(x_859); +x_861 = lean_usize_sub(x_860, x_704); +x_862 = lean_usize_land(x_702, x_861); +x_863 = lean_array_uget(x_858, x_862); +x_864 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__4(x_626, x_863); +if (x_864 == 0) +{ +lean_object* x_865; lean_object* x_866; lean_object* x_867; lean_object* x_868; lean_object* x_869; lean_object* x_870; lean_object* x_871; lean_object* x_872; uint8_t x_873; +x_865 = lean_nat_add(x_857, x_629); +lean_dec(x_857); +lean_inc(x_626); +x_866 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_866, 0, x_626); +lean_ctor_set(x_866, 1, x_711); +lean_ctor_set(x_866, 2, x_863); +x_867 = lean_array_uset(x_858, x_862, x_866); +x_868 = lean_unsigned_to_nat(4u); +x_869 = lean_nat_mul(x_865, x_868); +x_870 = lean_unsigned_to_nat(3u); +x_871 = lean_nat_div(x_869, x_870); +lean_dec(x_869); +x_872 = lean_array_get_size(x_867); +x_873 = lean_nat_dec_le(x_871, x_872); +lean_dec(x_872); +lean_dec(x_871); +if (x_873 == 0) +{ +lean_object* x_874; lean_object* x_875; lean_object* x_876; lean_object* x_877; +x_874 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__11(x_867); +x_875 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_875, 0, x_865); +lean_ctor_set(x_875, 1, x_874); +lean_ctor_set(x_823, 0, x_875); +x_876 = lean_st_ref_set(x_2, x_823, x_825); +x_877 = lean_ctor_get(x_876, 1); +lean_inc(x_877); +lean_dec(x_876); +x_712 = x_877; +goto block_821; +} +else +{ +lean_object* x_878; lean_object* x_879; lean_object* x_880; +x_878 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_878, 0, x_865); +lean_ctor_set(x_878, 1, x_867); +lean_ctor_set(x_823, 0, x_878); +x_879 = lean_st_ref_set(x_2, x_823, x_825); +x_880 = lean_ctor_get(x_879, 1); +lean_inc(x_880); +lean_dec(x_879); +x_712 = x_880; +goto block_821; +} +} +else +{ +lean_object* x_881; lean_object* x_882; lean_object* x_883; lean_object* x_884; lean_object* x_885; lean_object* x_886; lean_object* x_887; +x_881 = lean_box(0); +x_882 = lean_array_uset(x_858, x_862, x_881); +lean_inc(x_626); +x_883 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__14(x_626, x_711, x_863); +x_884 = lean_array_uset(x_882, x_862, x_883); +x_885 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_885, 0, x_857); +lean_ctor_set(x_885, 1, x_884); +lean_ctor_set(x_823, 0, x_885); +x_886 = lean_st_ref_set(x_2, x_823, x_825); +x_887 = lean_ctor_get(x_886, 1); +lean_inc(x_887); +lean_dec(x_886); +x_712 = x_887; +goto block_821; +} +} +} +else +{ +lean_object* x_888; lean_object* x_889; lean_object* x_890; lean_object* x_891; lean_object* x_892; size_t x_893; size_t x_894; size_t x_895; lean_object* x_896; uint8_t x_897; +x_888 = lean_ctor_get(x_823, 1); +lean_inc(x_888); +lean_dec(x_823); +x_889 = lean_ctor_get(x_824, 0); +lean_inc(x_889); +x_890 = lean_ctor_get(x_824, 1); +lean_inc(x_890); +if (lean_is_exclusive(x_824)) { + lean_ctor_release(x_824, 0); + lean_ctor_release(x_824, 1); + x_891 = x_824; +} else { + lean_dec_ref(x_824); + x_891 = lean_box(0); +} +x_892 = lean_array_get_size(x_890); +x_893 = lean_usize_of_nat(x_892); +lean_dec(x_892); +x_894 = lean_usize_sub(x_893, x_704); +x_895 = lean_usize_land(x_702, x_894); +x_896 = lean_array_uget(x_890, x_895); +x_897 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__4(x_626, x_896); +if (x_897 == 0) +{ +lean_object* x_898; lean_object* x_899; lean_object* x_900; lean_object* x_901; lean_object* x_902; lean_object* x_903; lean_object* x_904; lean_object* x_905; uint8_t x_906; +x_898 = lean_nat_add(x_889, x_629); +lean_dec(x_889); +lean_inc(x_626); +x_899 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_899, 0, x_626); +lean_ctor_set(x_899, 1, x_711); +lean_ctor_set(x_899, 2, x_896); +x_900 = lean_array_uset(x_890, x_895, x_899); +x_901 = lean_unsigned_to_nat(4u); +x_902 = lean_nat_mul(x_898, x_901); +x_903 = lean_unsigned_to_nat(3u); +x_904 = lean_nat_div(x_902, x_903); +lean_dec(x_902); +x_905 = lean_array_get_size(x_900); +x_906 = lean_nat_dec_le(x_904, x_905); +lean_dec(x_905); +lean_dec(x_904); +if (x_906 == 0) +{ +lean_object* x_907; lean_object* x_908; lean_object* x_909; lean_object* x_910; lean_object* x_911; +x_907 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__11(x_900); +if (lean_is_scalar(x_891)) { + x_908 = lean_alloc_ctor(0, 2, 0); +} else { + x_908 = x_891; +} +lean_ctor_set(x_908, 0, x_898); +lean_ctor_set(x_908, 1, x_907); +x_909 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_909, 0, x_908); +lean_ctor_set(x_909, 1, x_888); +x_910 = lean_st_ref_set(x_2, x_909, x_825); +x_911 = lean_ctor_get(x_910, 1); +lean_inc(x_911); +lean_dec(x_910); +x_712 = x_911; +goto block_821; +} +else +{ +lean_object* x_912; lean_object* x_913; lean_object* x_914; lean_object* x_915; +if (lean_is_scalar(x_891)) { + x_912 = lean_alloc_ctor(0, 2, 0); +} else { + x_912 = x_891; +} +lean_ctor_set(x_912, 0, x_898); +lean_ctor_set(x_912, 1, x_900); +x_913 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_913, 0, x_912); +lean_ctor_set(x_913, 1, x_888); +x_914 = lean_st_ref_set(x_2, x_913, x_825); +x_915 = lean_ctor_get(x_914, 1); +lean_inc(x_915); +lean_dec(x_914); +x_712 = x_915; +goto block_821; +} +} +else +{ +lean_object* x_916; lean_object* x_917; lean_object* x_918; lean_object* x_919; lean_object* x_920; lean_object* x_921; lean_object* x_922; lean_object* x_923; +x_916 = lean_box(0); +x_917 = lean_array_uset(x_890, x_895, x_916); +lean_inc(x_626); +x_918 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__14(x_626, x_711, x_896); +x_919 = lean_array_uset(x_917, x_895, x_918); +if (lean_is_scalar(x_891)) { + x_920 = lean_alloc_ctor(0, 2, 0); +} else { + x_920 = x_891; +} +lean_ctor_set(x_920, 0, x_889); +lean_ctor_set(x_920, 1, x_919); +x_921 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_921, 0, x_920); +lean_ctor_set(x_921, 1, x_888); +x_922 = lean_st_ref_set(x_2, x_921, x_825); +x_923 = lean_ctor_get(x_922, 1); +lean_inc(x_923); +lean_dec(x_922); +x_712 = x_923; +goto block_821; +} +} +block_821: +{ +lean_object* x_713; +lean_inc(x_2); +x_713 = l_Lean_Compiler_LCNF_StructProjCases_visitCode(x_686, x_2, x_3, x_4, x_5, x_6, x_712); +if (lean_obj_tag(x_713) == 0) +{ +lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_751; lean_object* x_752; lean_object* x_753; lean_object* x_754; uint8_t x_755; +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); +} +x_751 = lean_st_ref_take(x_2, x_715); +x_752 = lean_ctor_get(x_751, 0); +lean_inc(x_752); +x_753 = lean_ctor_get(x_752, 0); +lean_inc(x_753); +x_754 = lean_ctor_get(x_751, 1); +lean_inc(x_754); +lean_dec(x_751); +x_755 = !lean_is_exclusive(x_752); +if (x_755 == 0) +{ +lean_object* x_756; uint8_t x_757; +x_756 = lean_ctor_get(x_752, 0); +lean_dec(x_756); +x_757 = !lean_is_exclusive(x_753); +if (x_757 == 0) +{ +lean_object* x_758; lean_object* x_759; lean_object* x_760; size_t x_761; size_t x_762; size_t x_763; lean_object* x_764; uint8_t x_765; +x_758 = lean_ctor_get(x_753, 0); +x_759 = lean_ctor_get(x_753, 1); +x_760 = lean_array_get_size(x_759); +x_761 = lean_usize_of_nat(x_760); +lean_dec(x_760); +x_762 = lean_usize_sub(x_761, x_704); +x_763 = lean_usize_land(x_702, x_762); +x_764 = lean_array_uget(x_759, x_763); +x_765 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__4(x_626, x_764); +if (x_765 == 0) +{ +lean_object* x_766; lean_object* x_767; +lean_dec(x_764); +x_766 = lean_st_ref_set(x_2, x_752, x_754); +lean_dec(x_2); +x_767 = lean_ctor_get(x_766, 1); +lean_inc(x_767); +lean_dec(x_766); +x_717 = x_767; +goto block_750; +} +else +{ +lean_object* x_768; lean_object* x_769; lean_object* x_770; lean_object* x_771; lean_object* x_772; lean_object* x_773; lean_object* x_774; +x_768 = lean_box(0); +x_769 = lean_array_uset(x_759, x_763, x_768); +x_770 = lean_nat_sub(x_758, x_629); +lean_dec(x_758); +x_771 = l_Std_DHashMap_Internal_AssocList_erase___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__5(x_626, x_764); +x_772 = lean_array_uset(x_769, x_763, x_771); +lean_ctor_set(x_753, 1, x_772); +lean_ctor_set(x_753, 0, x_770); +x_773 = lean_st_ref_set(x_2, x_752, x_754); +lean_dec(x_2); +x_774 = lean_ctor_get(x_773, 1); +lean_inc(x_774); +lean_dec(x_773); +x_717 = x_774; +goto block_750; +} +} +else +{ +lean_object* x_775; lean_object* x_776; lean_object* x_777; size_t x_778; size_t x_779; size_t x_780; lean_object* x_781; uint8_t x_782; +x_775 = lean_ctor_get(x_753, 0); +x_776 = lean_ctor_get(x_753, 1); +lean_inc(x_776); +lean_inc(x_775); +lean_dec(x_753); +x_777 = lean_array_get_size(x_776); +x_778 = lean_usize_of_nat(x_777); +lean_dec(x_777); +x_779 = lean_usize_sub(x_778, x_704); +x_780 = lean_usize_land(x_702, x_779); +x_781 = lean_array_uget(x_776, x_780); +x_782 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__4(x_626, x_781); +if (x_782 == 0) +{ +lean_object* x_783; lean_object* x_784; lean_object* x_785; +lean_dec(x_781); +x_783 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_783, 0, x_775); +lean_ctor_set(x_783, 1, x_776); +lean_ctor_set(x_752, 0, x_783); +x_784 = lean_st_ref_set(x_2, x_752, x_754); +lean_dec(x_2); +x_785 = lean_ctor_get(x_784, 1); +lean_inc(x_785); +lean_dec(x_784); +x_717 = x_785; +goto block_750; +} +else +{ +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; +x_786 = lean_box(0); +x_787 = lean_array_uset(x_776, x_780, x_786); +x_788 = lean_nat_sub(x_775, x_629); +lean_dec(x_775); +x_789 = l_Std_DHashMap_Internal_AssocList_erase___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__5(x_626, x_781); +x_790 = lean_array_uset(x_787, x_780, x_789); +x_791 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_791, 0, x_788); +lean_ctor_set(x_791, 1, x_790); +lean_ctor_set(x_752, 0, x_791); +x_792 = lean_st_ref_set(x_2, x_752, x_754); +lean_dec(x_2); +x_793 = lean_ctor_get(x_792, 1); +lean_inc(x_793); +lean_dec(x_792); +x_717 = x_793; +goto block_750; +} +} +} +else +{ +lean_object* x_794; lean_object* x_795; lean_object* x_796; lean_object* x_797; lean_object* x_798; size_t x_799; size_t x_800; size_t x_801; lean_object* x_802; uint8_t x_803; +x_794 = lean_ctor_get(x_752, 1); +lean_inc(x_794); +lean_dec(x_752); +x_795 = lean_ctor_get(x_753, 0); +lean_inc(x_795); +x_796 = lean_ctor_get(x_753, 1); +lean_inc(x_796); +if (lean_is_exclusive(x_753)) { + lean_ctor_release(x_753, 0); + lean_ctor_release(x_753, 1); + x_797 = x_753; +} else { + lean_dec_ref(x_753); + x_797 = lean_box(0); +} +x_798 = lean_array_get_size(x_796); +x_799 = lean_usize_of_nat(x_798); +lean_dec(x_798); +x_800 = lean_usize_sub(x_799, x_704); +x_801 = lean_usize_land(x_702, x_800); +x_802 = lean_array_uget(x_796, x_801); +x_803 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__4(x_626, x_802); +if (x_803 == 0) +{ +lean_object* x_804; lean_object* x_805; lean_object* x_806; lean_object* x_807; +lean_dec(x_802); +if (lean_is_scalar(x_797)) { + x_804 = lean_alloc_ctor(0, 2, 0); +} else { + x_804 = x_797; +} +lean_ctor_set(x_804, 0, x_795); +lean_ctor_set(x_804, 1, x_796); +x_805 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_805, 0, x_804); +lean_ctor_set(x_805, 1, x_794); +x_806 = lean_st_ref_set(x_2, x_805, x_754); +lean_dec(x_2); +x_807 = lean_ctor_get(x_806, 1); +lean_inc(x_807); +lean_dec(x_806); +x_717 = x_807; +goto block_750; +} +else +{ +lean_object* x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; lean_object* x_812; lean_object* x_813; lean_object* x_814; lean_object* x_815; lean_object* x_816; +x_808 = lean_box(0); +x_809 = lean_array_uset(x_796, x_801, x_808); +x_810 = lean_nat_sub(x_795, x_629); +lean_dec(x_795); +x_811 = l_Std_DHashMap_Internal_AssocList_erase___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__5(x_626, x_802); +x_812 = lean_array_uset(x_809, x_801, x_811); +if (lean_is_scalar(x_797)) { + x_813 = lean_alloc_ctor(0, 2, 0); +} else { + x_813 = x_797; +} +lean_ctor_set(x_813, 0, x_810); +lean_ctor_set(x_813, 1, x_812); +x_814 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_814, 0, x_813); +lean_ctor_set(x_814, 1, x_794); +x_815 = lean_st_ref_set(x_2, x_814, x_754); +lean_dec(x_2); +x_816 = lean_ctor_get(x_815, 1); +lean_inc(x_816); +lean_dec(x_815); +x_717 = x_816; +goto block_750; +} +} +block_750: +{ +lean_object* x_718; lean_object* x_719; lean_object* x_720; lean_object* x_721; size_t x_722; size_t x_723; uint8_t x_724; +if (lean_is_scalar(x_687)) { + x_718 = lean_alloc_ctor(0, 3, 0); +} else { + x_718 = x_687; +} +lean_ctor_set(x_718, 0, x_684); +lean_ctor_set(x_718, 1, x_685); +lean_ctor_set(x_718, 2, x_714); +x_719 = lean_box(0); +if (lean_is_scalar(x_693)) { + x_720 = lean_alloc_ctor(1, 2, 0); +} else { + x_720 = x_693; + lean_ctor_set_tag(x_720, 1); +} +lean_ctor_set(x_720, 0, x_718); +lean_ctor_set(x_720, 1, x_719); +x_721 = lean_array_mk(x_720); +x_722 = lean_ptr_addr(x_623); +lean_dec(x_623); +x_723 = lean_ptr_addr(x_721); +x_724 = lean_usize_dec_eq(x_722, x_723); +if (x_724 == 0) +{ +uint8_t x_725; +lean_dec(x_622); +x_725 = !lean_is_exclusive(x_1); +if (x_725 == 0) +{ +lean_object* x_726; lean_object* x_727; lean_object* x_728; +x_726 = lean_ctor_get(x_1, 0); +lean_dec(x_726); +if (lean_is_scalar(x_624)) { + x_727 = lean_alloc_ctor(0, 4, 0); +} else { + x_727 = x_624; +} +lean_ctor_set(x_727, 0, x_620); +lean_ctor_set(x_727, 1, x_621); +lean_ctor_set(x_727, 2, x_626); +lean_ctor_set(x_727, 3, x_721); +lean_ctor_set(x_1, 0, x_727); +if (lean_is_scalar(x_716)) { + x_728 = lean_alloc_ctor(0, 2, 0); +} else { + x_728 = x_716; +} +lean_ctor_set(x_728, 0, x_1); +lean_ctor_set(x_728, 1, x_717); +return x_728; +} +else +{ +lean_object* x_729; lean_object* x_730; lean_object* x_731; +lean_dec(x_1); +if (lean_is_scalar(x_624)) { + x_729 = lean_alloc_ctor(0, 4, 0); +} else { + x_729 = x_624; +} +lean_ctor_set(x_729, 0, x_620); +lean_ctor_set(x_729, 1, x_621); +lean_ctor_set(x_729, 2, x_626); +lean_ctor_set(x_729, 3, x_721); +x_730 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_730, 0, x_729); +if (lean_is_scalar(x_716)) { + x_731 = lean_alloc_ctor(0, 2, 0); +} else { + x_731 = x_716; +} +lean_ctor_set(x_731, 0, x_730); +lean_ctor_set(x_731, 1, x_717); +return x_731; +} +} +else +{ +size_t x_732; uint8_t x_733; +x_732 = lean_ptr_addr(x_621); +x_733 = lean_usize_dec_eq(x_732, x_732); +if (x_733 == 0) +{ +uint8_t x_734; +lean_dec(x_622); +x_734 = !lean_is_exclusive(x_1); +if (x_734 == 0) +{ +lean_object* x_735; lean_object* x_736; lean_object* x_737; +x_735 = lean_ctor_get(x_1, 0); +lean_dec(x_735); +if (lean_is_scalar(x_624)) { + x_736 = lean_alloc_ctor(0, 4, 0); +} else { + x_736 = x_624; +} +lean_ctor_set(x_736, 0, x_620); +lean_ctor_set(x_736, 1, x_621); +lean_ctor_set(x_736, 2, x_626); +lean_ctor_set(x_736, 3, x_721); +lean_ctor_set(x_1, 0, x_736); +if (lean_is_scalar(x_716)) { + x_737 = lean_alloc_ctor(0, 2, 0); +} else { + x_737 = x_716; +} +lean_ctor_set(x_737, 0, x_1); +lean_ctor_set(x_737, 1, x_717); +return x_737; +} +else +{ +lean_object* x_738; lean_object* x_739; lean_object* x_740; +lean_dec(x_1); +if (lean_is_scalar(x_624)) { + x_738 = lean_alloc_ctor(0, 4, 0); +} else { + x_738 = x_624; +} +lean_ctor_set(x_738, 0, x_620); +lean_ctor_set(x_738, 1, x_621); +lean_ctor_set(x_738, 2, x_626); +lean_ctor_set(x_738, 3, x_721); +x_739 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_739, 0, x_738); +if (lean_is_scalar(x_716)) { + x_740 = lean_alloc_ctor(0, 2, 0); +} else { + x_740 = x_716; +} +lean_ctor_set(x_740, 0, x_739); +lean_ctor_set(x_740, 1, x_717); +return x_740; +} +} +else +{ +uint8_t x_741; +x_741 = lean_name_eq(x_622, x_626); +lean_dec(x_622); +if (x_741 == 0) +{ +uint8_t x_742; +x_742 = !lean_is_exclusive(x_1); +if (x_742 == 0) +{ +lean_object* x_743; lean_object* x_744; lean_object* x_745; +x_743 = lean_ctor_get(x_1, 0); +lean_dec(x_743); +if (lean_is_scalar(x_624)) { + x_744 = lean_alloc_ctor(0, 4, 0); +} else { + x_744 = x_624; +} +lean_ctor_set(x_744, 0, x_620); +lean_ctor_set(x_744, 1, x_621); +lean_ctor_set(x_744, 2, x_626); +lean_ctor_set(x_744, 3, x_721); +lean_ctor_set(x_1, 0, x_744); +if (lean_is_scalar(x_716)) { + x_745 = lean_alloc_ctor(0, 2, 0); +} else { + x_745 = x_716; +} +lean_ctor_set(x_745, 0, x_1); +lean_ctor_set(x_745, 1, x_717); +return x_745; +} +else +{ +lean_object* x_746; lean_object* x_747; lean_object* x_748; +lean_dec(x_1); +if (lean_is_scalar(x_624)) { + x_746 = lean_alloc_ctor(0, 4, 0); +} else { + x_746 = x_624; +} +lean_ctor_set(x_746, 0, x_620); +lean_ctor_set(x_746, 1, x_621); +lean_ctor_set(x_746, 2, x_626); +lean_ctor_set(x_746, 3, x_721); +x_747 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_747, 0, x_746); +if (lean_is_scalar(x_716)) { + x_748 = lean_alloc_ctor(0, 2, 0); +} else { + x_748 = x_716; +} +lean_ctor_set(x_748, 0, x_747); +lean_ctor_set(x_748, 1, x_717); +return x_748; +} +} +else +{ +lean_object* x_749; +lean_dec(x_721); +lean_dec(x_626); +lean_dec(x_624); +lean_dec(x_621); +lean_dec(x_620); +if (lean_is_scalar(x_716)) { + x_749 = lean_alloc_ctor(0, 2, 0); +} else { + x_749 = x_716; +} +lean_ctor_set(x_749, 0, x_1); +lean_ctor_set(x_749, 1, x_717); +return x_749; +} +} +} +} +} +else +{ +uint8_t x_817; +lean_dec(x_693); +lean_dec(x_687); +lean_dec(x_685); +lean_dec(x_684); +lean_dec(x_626); +lean_dec(x_624); +lean_dec(x_623); +lean_dec(x_622); +lean_dec(x_621); +lean_dec(x_620); +lean_dec(x_2); +lean_dec(x_1); +x_817 = !lean_is_exclusive(x_713); +if (x_817 == 0) +{ +return x_713; +} +else +{ +lean_object* x_818; lean_object* x_819; lean_object* x_820; +x_818 = lean_ctor_get(x_713, 0); +x_819 = lean_ctor_get(x_713, 1); +lean_inc(x_819); +lean_inc(x_818); +lean_dec(x_713); +x_820 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_820, 0, x_818); +lean_ctor_set(x_820, 1, x_819); +return x_820; +} +} +} +} +else +{ +lean_object* x_924; lean_object* x_925; lean_object* x_926; uint8_t x_927; +lean_dec(x_693); +lean_dec(x_687); +lean_dec(x_684); +lean_dec(x_626); +lean_dec(x_624); +lean_dec(x_623); +lean_dec(x_622); +lean_dec(x_621); +lean_dec(x_620); +lean_dec(x_1); +x_924 = lean_ctor_get(x_708, 0); +lean_inc(x_924); +lean_dec(x_708); +x_925 = lean_array_get_size(x_924); +x_926 = lean_array_get_size(x_685); +x_927 = lean_nat_dec_eq(x_925, x_926); +lean_dec(x_926); +if (x_927 == 0) +{ +lean_object* x_928; lean_object* x_929; +lean_dec(x_925); +lean_dec(x_924); +lean_dec(x_686); +lean_dec(x_685); +x_928 = l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__10; +x_929 = l_panic___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__2(x_928, x_2, x_3, x_4, x_5, x_6, x_691); +return x_929; +} +else +{ +lean_object* x_930; lean_object* x_931; size_t x_932; size_t x_933; lean_object* x_934; lean_object* x_935; +x_930 = l_Array_toSubarray___rarg(x_924, x_682, x_925); +x_931 = lean_box(0); +x_932 = lean_array_size(x_685); +x_933 = 0; +x_934 = l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__16(x_685, x_696, x_699, x_704, x_931, x_685, x_932, x_933, x_930, x_2, x_3, x_4, x_5, x_6, x_691); +lean_dec(x_685); +x_935 = lean_ctor_get(x_934, 1); +lean_inc(x_935); +lean_dec(x_934); +x_1 = x_686; +x_7 = x_935; +goto _start; +} +} +} +else +{ +size_t x_937; size_t x_938; lean_object* x_939; +lean_dec(x_683); +x_937 = lean_array_size(x_623); +x_938 = 0; +lean_inc(x_623); +x_939 = l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__15(x_937, x_938, x_623, x_2, x_3, x_4, x_5, x_6, x_627); +if (lean_obj_tag(x_939) == 0) +{ +uint8_t x_940; +x_940 = !lean_is_exclusive(x_939); +if (x_940 == 0) +{ +lean_object* x_941; size_t x_942; size_t x_943; uint8_t x_944; +x_941 = lean_ctor_get(x_939, 0); +x_942 = lean_ptr_addr(x_623); +lean_dec(x_623); +x_943 = lean_ptr_addr(x_941); +x_944 = lean_usize_dec_eq(x_942, x_943); +if (x_944 == 0) +{ +uint8_t x_945; +lean_dec(x_622); +x_945 = !lean_is_exclusive(x_1); +if (x_945 == 0) +{ +lean_object* x_946; lean_object* x_947; +x_946 = lean_ctor_get(x_1, 0); +lean_dec(x_946); +if (lean_is_scalar(x_624)) { + x_947 = lean_alloc_ctor(0, 4, 0); +} else { + x_947 = x_624; +} +lean_ctor_set(x_947, 0, x_620); +lean_ctor_set(x_947, 1, x_621); +lean_ctor_set(x_947, 2, x_626); +lean_ctor_set(x_947, 3, x_941); +lean_ctor_set(x_1, 0, x_947); +lean_ctor_set(x_939, 0, x_1); +return x_939; +} +else +{ +lean_object* x_948; lean_object* x_949; +lean_dec(x_1); +if (lean_is_scalar(x_624)) { + x_948 = lean_alloc_ctor(0, 4, 0); +} else { + x_948 = x_624; +} +lean_ctor_set(x_948, 0, x_620); +lean_ctor_set(x_948, 1, x_621); +lean_ctor_set(x_948, 2, x_626); +lean_ctor_set(x_948, 3, x_941); +x_949 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_949, 0, x_948); +lean_ctor_set(x_939, 0, x_949); +return x_939; +} +} +else +{ +size_t x_950; uint8_t x_951; +x_950 = lean_ptr_addr(x_621); +x_951 = lean_usize_dec_eq(x_950, x_950); +if (x_951 == 0) +{ +uint8_t x_952; +lean_dec(x_622); +x_952 = !lean_is_exclusive(x_1); +if (x_952 == 0) +{ +lean_object* x_953; lean_object* x_954; +x_953 = lean_ctor_get(x_1, 0); +lean_dec(x_953); +if (lean_is_scalar(x_624)) { + x_954 = lean_alloc_ctor(0, 4, 0); +} else { + x_954 = x_624; +} +lean_ctor_set(x_954, 0, x_620); +lean_ctor_set(x_954, 1, x_621); +lean_ctor_set(x_954, 2, x_626); +lean_ctor_set(x_954, 3, x_941); +lean_ctor_set(x_1, 0, x_954); +lean_ctor_set(x_939, 0, x_1); +return x_939; +} +else +{ +lean_object* x_955; lean_object* x_956; +lean_dec(x_1); +if (lean_is_scalar(x_624)) { + x_955 = lean_alloc_ctor(0, 4, 0); +} else { + x_955 = x_624; +} +lean_ctor_set(x_955, 0, x_620); +lean_ctor_set(x_955, 1, x_621); +lean_ctor_set(x_955, 2, x_626); +lean_ctor_set(x_955, 3, x_941); +x_956 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_956, 0, x_955); +lean_ctor_set(x_939, 0, x_956); +return x_939; +} +} +else +{ +uint8_t x_957; +x_957 = lean_name_eq(x_622, x_626); +lean_dec(x_622); +if (x_957 == 0) +{ +uint8_t x_958; +x_958 = !lean_is_exclusive(x_1); +if (x_958 == 0) +{ +lean_object* x_959; lean_object* x_960; +x_959 = lean_ctor_get(x_1, 0); +lean_dec(x_959); +if (lean_is_scalar(x_624)) { + x_960 = lean_alloc_ctor(0, 4, 0); +} else { + x_960 = x_624; +} +lean_ctor_set(x_960, 0, x_620); +lean_ctor_set(x_960, 1, x_621); +lean_ctor_set(x_960, 2, x_626); +lean_ctor_set(x_960, 3, x_941); +lean_ctor_set(x_1, 0, x_960); +lean_ctor_set(x_939, 0, x_1); +return x_939; +} +else +{ +lean_object* x_961; lean_object* x_962; +lean_dec(x_1); +if (lean_is_scalar(x_624)) { + x_961 = lean_alloc_ctor(0, 4, 0); +} else { + x_961 = x_624; +} +lean_ctor_set(x_961, 0, x_620); +lean_ctor_set(x_961, 1, x_621); +lean_ctor_set(x_961, 2, x_626); +lean_ctor_set(x_961, 3, x_941); +x_962 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_962, 0, x_961); +lean_ctor_set(x_939, 0, x_962); +return x_939; +} +} +else +{ +lean_dec(x_941); +lean_dec(x_626); +lean_dec(x_624); +lean_dec(x_621); +lean_dec(x_620); +lean_ctor_set(x_939, 0, x_1); +return x_939; +} +} +} +} +else +{ +lean_object* x_963; lean_object* x_964; size_t x_965; size_t x_966; uint8_t x_967; +x_963 = lean_ctor_get(x_939, 0); +x_964 = lean_ctor_get(x_939, 1); +lean_inc(x_964); +lean_inc(x_963); +lean_dec(x_939); +x_965 = lean_ptr_addr(x_623); +lean_dec(x_623); +x_966 = lean_ptr_addr(x_963); +x_967 = lean_usize_dec_eq(x_965, x_966); +if (x_967 == 0) +{ +lean_object* x_968; lean_object* x_969; lean_object* x_970; lean_object* x_971; +lean_dec(x_622); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + x_968 = x_1; +} else { + lean_dec_ref(x_1); + x_968 = lean_box(0); +} +if (lean_is_scalar(x_624)) { + x_969 = lean_alloc_ctor(0, 4, 0); +} else { + x_969 = x_624; +} +lean_ctor_set(x_969, 0, x_620); +lean_ctor_set(x_969, 1, x_621); +lean_ctor_set(x_969, 2, x_626); +lean_ctor_set(x_969, 3, x_963); +if (lean_is_scalar(x_968)) { + x_970 = lean_alloc_ctor(4, 1, 0); +} else { + x_970 = x_968; +} +lean_ctor_set(x_970, 0, x_969); +x_971 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_971, 0, x_970); +lean_ctor_set(x_971, 1, x_964); +return x_971; +} +else +{ +size_t x_972; uint8_t x_973; +x_972 = lean_ptr_addr(x_621); +x_973 = lean_usize_dec_eq(x_972, x_972); +if (x_973 == 0) +{ +lean_object* x_974; lean_object* x_975; lean_object* x_976; lean_object* x_977; +lean_dec(x_622); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + x_974 = x_1; +} else { + lean_dec_ref(x_1); + x_974 = lean_box(0); +} +if (lean_is_scalar(x_624)) { + x_975 = lean_alloc_ctor(0, 4, 0); +} else { + x_975 = x_624; +} +lean_ctor_set(x_975, 0, x_620); +lean_ctor_set(x_975, 1, x_621); +lean_ctor_set(x_975, 2, x_626); +lean_ctor_set(x_975, 3, x_963); +if (lean_is_scalar(x_974)) { + x_976 = lean_alloc_ctor(4, 1, 0); +} else { + x_976 = x_974; +} +lean_ctor_set(x_976, 0, x_975); +x_977 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_977, 0, x_976); +lean_ctor_set(x_977, 1, x_964); +return x_977; +} +else +{ +uint8_t x_978; +x_978 = lean_name_eq(x_622, x_626); +lean_dec(x_622); +if (x_978 == 0) +{ +lean_object* x_979; lean_object* x_980; lean_object* x_981; lean_object* x_982; +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + x_979 = x_1; +} else { + lean_dec_ref(x_1); + x_979 = lean_box(0); +} +if (lean_is_scalar(x_624)) { + x_980 = lean_alloc_ctor(0, 4, 0); +} else { + x_980 = x_624; +} +lean_ctor_set(x_980, 0, x_620); +lean_ctor_set(x_980, 1, x_621); +lean_ctor_set(x_980, 2, x_626); +lean_ctor_set(x_980, 3, x_963); +if (lean_is_scalar(x_979)) { + x_981 = lean_alloc_ctor(4, 1, 0); +} else { + x_981 = x_979; +} +lean_ctor_set(x_981, 0, x_980); +x_982 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_982, 0, x_981); +lean_ctor_set(x_982, 1, x_964); +return x_982; +} +else +{ +lean_object* x_983; +lean_dec(x_963); +lean_dec(x_626); +lean_dec(x_624); +lean_dec(x_621); +lean_dec(x_620); +x_983 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_983, 0, x_1); +lean_ctor_set(x_983, 1, x_964); +return x_983; +} +} +} +} +} +else +{ +uint8_t x_984; +lean_dec(x_626); +lean_dec(x_624); +lean_dec(x_623); +lean_dec(x_622); +lean_dec(x_621); +lean_dec(x_620); +lean_dec(x_1); +x_984 = !lean_is_exclusive(x_939); +if (x_984 == 0) +{ +return x_939; +} +else +{ +lean_object* x_985; lean_object* x_986; lean_object* x_987; +x_985 = lean_ctor_get(x_939, 0); +x_986 = lean_ctor_get(x_939, 1); +lean_inc(x_986); +lean_inc(x_985); +lean_dec(x_939); +x_987 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_987, 0, x_985); +lean_ctor_set(x_987, 1, x_986); +return x_987; +} +} +} +} +} +case 5: +{ +lean_object* x_988; lean_object* x_989; uint8_t x_990; +x_988 = lean_ctor_get(x_1, 0); +lean_inc(x_988); +lean_inc(x_988); +x_989 = l_Lean_Compiler_LCNF_StructProjCases_remapFVar(x_988, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_990 = !lean_is_exclusive(x_989); +if (x_990 == 0) +{ +lean_object* x_991; uint8_t x_992; +x_991 = lean_ctor_get(x_989, 0); +x_992 = lean_name_eq(x_988, x_991); +lean_dec(x_988); +if (x_992 == 0) +{ +uint8_t x_993; +x_993 = !lean_is_exclusive(x_1); +if (x_993 == 0) +{ +lean_object* x_994; +x_994 = lean_ctor_get(x_1, 0); +lean_dec(x_994); +lean_ctor_set(x_1, 0, x_991); +lean_ctor_set(x_989, 0, x_1); +return x_989; +} +else +{ +lean_object* x_995; +lean_dec(x_1); +x_995 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_995, 0, x_991); +lean_ctor_set(x_989, 0, x_995); +return x_989; +} +} +else +{ +lean_dec(x_991); +lean_ctor_set(x_989, 0, x_1); +return x_989; +} +} +else +{ +lean_object* x_996; lean_object* x_997; uint8_t x_998; +x_996 = lean_ctor_get(x_989, 0); +x_997 = lean_ctor_get(x_989, 1); +lean_inc(x_997); +lean_inc(x_996); +lean_dec(x_989); +x_998 = lean_name_eq(x_988, x_996); +lean_dec(x_988); +if (x_998 == 0) +{ +lean_object* x_999; lean_object* x_1000; lean_object* x_1001; +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + x_999 = x_1; +} else { + lean_dec_ref(x_1); + x_999 = lean_box(0); +} +if (lean_is_scalar(x_999)) { + x_1000 = lean_alloc_ctor(5, 1, 0); +} else { + x_1000 = x_999; +} +lean_ctor_set(x_1000, 0, x_996); +x_1001 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1001, 0, x_1000); +lean_ctor_set(x_1001, 1, x_997); +return x_1001; +} +else +{ +lean_object* x_1002; +lean_dec(x_996); +x_1002 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1002, 0, x_1); +lean_ctor_set(x_1002, 1, x_997); +return x_1002; +} +} +} +default: +{ +lean_object* x_1003; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_1003 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1003, 0, x_1); +lean_ctor_set(x_1003, 1, x_7); +return x_1003; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_StructProjCases_visitAlt(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; +x_8 = l_Lean_Compiler_LCNF_Alt_getCode(x_1); +x_9 = l_Lean_Compiler_LCNF_StructProjCases_visitCode(x_8, x_2, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_9) == 0) +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_9, 0); +x_12 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltCodeImp(x_1, x_11); +lean_ctor_set(x_9, 0, x_12); +return x_9; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_9, 0); +x_14 = lean_ctor_get(x_9, 1); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_9); +x_15 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltCodeImp(x_1, x_13); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_14); +return x_16; +} +} +else +{ +uint8_t x_17; +lean_dec(x_1); +x_17 = !lean_is_exclusive(x_9); +if (x_17 == 0) +{ +return x_9; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_9, 0); +x_19 = lean_ctor_get(x_9, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_9); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +return x_20; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Compiler_LCNF_StructProjCases_visitCode___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_Compiler_LCNF_StructProjCases_visitCode___spec__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_StructProjCases_visitCode___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_Compiler_LCNF_StructProjCases_visitCode___spec__3(x_4, x_5, x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__4___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_Compiler_LCNF_StructProjCases_visitCode___spec__4(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_erase___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__5___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_DHashMap_Internal_AssocList_erase___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__5(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__6___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_Compiler_LCNF_StructProjCases_visitCode___spec__6(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_mapMUnsafe_map___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__15___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +size_t x_10; size_t x_11; lean_object* x_12; +x_10 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_11 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_12 = l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__15(x_10, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__16___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +uint64_t x_16; uint64_t x_17; size_t x_18; size_t x_19; size_t x_20; lean_object* x_21; +x_16 = lean_unbox_uint64(x_2); +lean_dec(x_2); +x_17 = lean_unbox_uint64(x_3); +lean_dec(x_3); +x_18 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_19 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_20 = lean_unbox_usize(x_8); +lean_dec(x_8); +x_21 = l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__16(x_1, x_16, x_17, x_18, x_5, x_6, x_19, x_20, 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_6); +lean_dec(x_5); +lean_dec(x_1); +return x_21; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_DeclValue_mapCodeM___at_Lean_Compiler_LCNF_StructProjCases_visitDecl___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_2) == 0) +{ +uint8_t x_9; +x_9 = !lean_is_exclusive(x_2); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_2, 0); +x_11 = lean_apply_7(x_1, x_10, x_3, x_4, x_5, x_6, x_7, x_8); +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; +x_13 = lean_ctor_get(x_11, 0); +lean_ctor_set(x_2, 0, x_13); +lean_ctor_set(x_11, 0, x_2); +return x_11; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_11, 0); +x_15 = lean_ctor_get(x_11, 1); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_11); +lean_ctor_set(x_2, 0, x_14); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_2); +lean_ctor_set(x_16, 1, x_15); +return x_16; +} +} +else +{ +uint8_t x_17; +lean_free_object(x_2); +x_17 = !lean_is_exclusive(x_11); +if (x_17 == 0) +{ +return x_11; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_11, 0); +x_19 = lean_ctor_get(x_11, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_11); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +return x_20; +} +} +} +else +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_2, 0); +lean_inc(x_21); +lean_dec(x_2); +x_22 = lean_apply_7(x_1, x_21, x_3, x_4, x_5, x_6, x_7, x_8); +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); +if (lean_is_exclusive(x_22)) { + lean_ctor_release(x_22, 0); + lean_ctor_release(x_22, 1); + x_25 = x_22; +} else { + lean_dec_ref(x_22); + x_25 = lean_box(0); +} +x_26 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_26, 0, x_23); +if (lean_is_scalar(x_25)) { + x_27 = lean_alloc_ctor(0, 2, 0); +} else { + x_27 = x_25; +} +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; +x_28 = lean_ctor_get(x_22, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_22, 1); +lean_inc(x_29); +if (lean_is_exclusive(x_22)) { + lean_ctor_release(x_22, 0); + lean_ctor_release(x_22, 1); + x_30 = x_22; +} else { + lean_dec_ref(x_22); + x_30 = lean_box(0); +} +if (lean_is_scalar(x_30)) { + x_31 = lean_alloc_ctor(1, 2, 0); +} else { + x_31 = x_30; +} +lean_ctor_set(x_31, 0, x_28); +lean_ctor_set(x_31, 1, x_29); +return x_31; +} +} +} +else +{ +lean_object* 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_1); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_2); +lean_ctor_set(x_32, 1, x_8); +return x_32; +} +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_StructProjCases_visitDecl___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_StructProjCases_visitCode), 7, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_StructProjCases_visitDecl(lean_object* x_1, lean_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; +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 = l_Lean_Compiler_LCNF_StructProjCases_visitDecl___closed__1; +x_16 = l_Lean_Compiler_LCNF_DeclValue_mapCodeM___at_Lean_Compiler_LCNF_StructProjCases_visitDecl___spec__1(x_15, x_13, x_2, x_3, x_4, x_5, x_6, x_7); +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; +x_18 = lean_ctor_get(x_16, 0); +lean_ctor_set(x_1, 4, x_18); +lean_ctor_set(x_16, 0, x_1); +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_1, 4, x_19); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_1); +lean_ctor_set(x_21, 1, x_20); +return x_21; +} +} +else +{ +uint8_t x_22; +lean_free_object(x_1); +lean_dec(x_14); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) +{ +return x_16; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_16, 0); +x_24 = lean_ctor_get(x_16, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_16); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; +} +} +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_26 = lean_ctor_get(x_1, 0); +x_27 = lean_ctor_get(x_1, 1); +x_28 = lean_ctor_get(x_1, 2); +x_29 = lean_ctor_get(x_1, 3); +x_30 = lean_ctor_get(x_1, 4); +x_31 = lean_ctor_get_uint8(x_1, sizeof(void*)*6); +x_32 = lean_ctor_get_uint8(x_1, sizeof(void*)*6 + 1); +x_33 = lean_ctor_get(x_1, 5); +lean_inc(x_33); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_1); +x_34 = l_Lean_Compiler_LCNF_StructProjCases_visitDecl___closed__1; +x_35 = l_Lean_Compiler_LCNF_DeclValue_mapCodeM___at_Lean_Compiler_LCNF_StructProjCases_visitDecl___spec__1(x_34, x_30, x_2, x_3, x_4, x_5, x_6, x_7); +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; +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_alloc_ctor(0, 6, 2); +lean_ctor_set(x_39, 0, x_26); +lean_ctor_set(x_39, 1, x_27); +lean_ctor_set(x_39, 2, x_28); +lean_ctor_set(x_39, 3, x_29); +lean_ctor_set(x_39, 4, x_36); +lean_ctor_set(x_39, 5, x_33); +lean_ctor_set_uint8(x_39, sizeof(void*)*6, x_31); +lean_ctor_set_uint8(x_39, sizeof(void*)*6 + 1, x_32); +if (lean_is_scalar(x_38)) { + x_40 = lean_alloc_ctor(0, 2, 0); +} else { + x_40 = x_38; +} +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_dec(x_33); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +x_41 = lean_ctor_get(x_35, 0); +lean_inc(x_41); +x_42 = lean_ctor_get(x_35, 1); +lean_inc(x_42); +if (lean_is_exclusive(x_35)) { + lean_ctor_release(x_35, 0); + lean_ctor_release(x_35, 1); + x_43 = x_35; +} else { + lean_dec_ref(x_35); + x_43 = lean_box(0); +} +if (lean_is_scalar(x_43)) { + x_44 = lean_alloc_ctor(1, 2, 0); +} else { + x_44 = x_43; +} +lean_ctor_set(x_44, 0, x_41); +lean_ctor_set(x_44, 1, x_42); +return x_44; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_structProjCases___lambda__1(lean_object* x_1, 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_alloc_closure((void*)(l_Lean_Compiler_LCNF_StructProjCases_visitDecl), 7, 1); +lean_closure_set(x_7, 0, x_1); +x_8 = l_Lean_Compiler_LCNF_StructProjCases_M_run___rarg(x_7, x_2, x_3, x_4, x_5, x_6); +return x_8; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_structProjCases___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("structProjCases", 15, 15); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_structProjCases___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Compiler_LCNF_structProjCases___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_structProjCases___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_structProjCases___lambda__1), 6, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_structProjCases___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Compiler_LCNF_structProjCases___closed__2; +x_2 = l_Lean_Compiler_LCNF_structProjCases___closed__3; +x_3 = 1; +x_4 = lean_unsigned_to_nat(0u); +x_5 = l_Lean_Compiler_LCNF_Pass_mkPerDeclaration(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_structProjCases() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Compiler_LCNF_structProjCases___closed__4; +return x_1; +} +} +lean_object* initialize_Lean_Compiler_LCNF_Basic(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Compiler_LCNF_InferType(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Compiler_LCNF_MonoTypes(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Compiler_LCNF_PassManager(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Compiler_LCNF_PrettyPrinter(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Compiler_LCNF_StructProjCases(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_Compiler_LCNF_Basic(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Compiler_LCNF_InferType(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Compiler_LCNF_MonoTypes(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Compiler_LCNF_PassManager(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Compiler_LCNF_PrettyPrinter(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Compiler_LCNF_StructProjCases_mkFieldParamsForCtorType___closed__1 = _init_l_Lean_Compiler_LCNF_StructProjCases_mkFieldParamsForCtorType___closed__1(); +lean_mark_persistent(l_Lean_Compiler_LCNF_StructProjCases_mkFieldParamsForCtorType___closed__1); +l_Lean_Compiler_LCNF_StructProjCases_M_run___rarg___closed__1 = _init_l_Lean_Compiler_LCNF_StructProjCases_M_run___rarg___closed__1(); +lean_mark_persistent(l_Lean_Compiler_LCNF_StructProjCases_M_run___rarg___closed__1); +l_Lean_Compiler_LCNF_StructProjCases_M_run___rarg___closed__2 = _init_l_Lean_Compiler_LCNF_StructProjCases_M_run___rarg___closed__2(); +lean_mark_persistent(l_Lean_Compiler_LCNF_StructProjCases_M_run___rarg___closed__2); +l_Lean_Compiler_LCNF_StructProjCases_M_run___rarg___closed__3 = _init_l_Lean_Compiler_LCNF_StructProjCases_M_run___rarg___closed__3(); +lean_mark_persistent(l_Lean_Compiler_LCNF_StructProjCases_M_run___rarg___closed__3); +l_Lean_Compiler_LCNF_StructProjCases_M_run___rarg___closed__4 = _init_l_Lean_Compiler_LCNF_StructProjCases_M_run___rarg___closed__4(); +lean_mark_persistent(l_Lean_Compiler_LCNF_StructProjCases_M_run___rarg___closed__4); +l_panic___at_Lean_Compiler_LCNF_StructProjCases_visitLetValue___spec__1___closed__1 = _init_l_panic___at_Lean_Compiler_LCNF_StructProjCases_visitLetValue___spec__1___closed__1(); +lean_mark_persistent(l_panic___at_Lean_Compiler_LCNF_StructProjCases_visitLetValue___spec__1___closed__1); +l_panic___at_Lean_Compiler_LCNF_StructProjCases_visitLetValue___spec__1___closed__2 = _init_l_panic___at_Lean_Compiler_LCNF_StructProjCases_visitLetValue___spec__1___closed__2(); +lean_mark_persistent(l_panic___at_Lean_Compiler_LCNF_StructProjCases_visitLetValue___spec__1___closed__2); +l_Lean_Compiler_LCNF_StructProjCases_visitLetValue___closed__1 = _init_l_Lean_Compiler_LCNF_StructProjCases_visitLetValue___closed__1(); +lean_mark_persistent(l_Lean_Compiler_LCNF_StructProjCases_visitLetValue___closed__1); +l_Lean_Compiler_LCNF_StructProjCases_visitLetValue___closed__2 = _init_l_Lean_Compiler_LCNF_StructProjCases_visitLetValue___closed__2(); +lean_mark_persistent(l_Lean_Compiler_LCNF_StructProjCases_visitLetValue___closed__2); +l_Lean_Compiler_LCNF_StructProjCases_visitLetValue___closed__3 = _init_l_Lean_Compiler_LCNF_StructProjCases_visitLetValue___closed__3(); +lean_mark_persistent(l_Lean_Compiler_LCNF_StructProjCases_visitLetValue___closed__3); +l_Lean_Compiler_LCNF_StructProjCases_visitLetValue___closed__4 = _init_l_Lean_Compiler_LCNF_StructProjCases_visitLetValue___closed__4(); +lean_mark_persistent(l_Lean_Compiler_LCNF_StructProjCases_visitLetValue___closed__4); +l_panic___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__2___closed__1 = _init_l_panic___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__2___closed__1(); +lean_mark_persistent(l_panic___at_Lean_Compiler_LCNF_StructProjCases_visitCode___spec__2___closed__1); +l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__1 = _init_l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__1(); +lean_mark_persistent(l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__1); +l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__2 = _init_l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__2(); +lean_mark_persistent(l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__2); +l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__3 = _init_l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__3(); +lean_mark_persistent(l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__3); +l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__4 = _init_l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__4(); +lean_mark_persistent(l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__4); +l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__5 = _init_l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__5(); +lean_mark_persistent(l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__5); +l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__6 = _init_l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__6(); +lean_mark_persistent(l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__6); +l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__7 = _init_l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__7(); +lean_mark_persistent(l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__7); +l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__8 = _init_l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__8(); +lean_mark_persistent(l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__8); +l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__9 = _init_l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__9(); +lean_mark_persistent(l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__9); +l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__10 = _init_l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__10(); +lean_mark_persistent(l_Lean_Compiler_LCNF_StructProjCases_visitCode___closed__10); +l_Lean_Compiler_LCNF_StructProjCases_visitDecl___closed__1 = _init_l_Lean_Compiler_LCNF_StructProjCases_visitDecl___closed__1(); +lean_mark_persistent(l_Lean_Compiler_LCNF_StructProjCases_visitDecl___closed__1); +l_Lean_Compiler_LCNF_structProjCases___closed__1 = _init_l_Lean_Compiler_LCNF_structProjCases___closed__1(); +lean_mark_persistent(l_Lean_Compiler_LCNF_structProjCases___closed__1); +l_Lean_Compiler_LCNF_structProjCases___closed__2 = _init_l_Lean_Compiler_LCNF_structProjCases___closed__2(); +lean_mark_persistent(l_Lean_Compiler_LCNF_structProjCases___closed__2); +l_Lean_Compiler_LCNF_structProjCases___closed__3 = _init_l_Lean_Compiler_LCNF_structProjCases___closed__3(); +lean_mark_persistent(l_Lean_Compiler_LCNF_structProjCases___closed__3); +l_Lean_Compiler_LCNF_structProjCases___closed__4 = _init_l_Lean_Compiler_LCNF_structProjCases___closed__4(); +lean_mark_persistent(l_Lean_Compiler_LCNF_structProjCases___closed__4); +l_Lean_Compiler_LCNF_structProjCases = _init_l_Lean_Compiler_LCNF_structProjCases(); +lean_mark_persistent(l_Lean_Compiler_LCNF_structProjCases); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Testing.c b/stage0/stdlib/Lean/Compiler/LCNF/Testing.c index 1066a955b381..bff57e849385 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Testing.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Testing.c @@ -98,12 +98,10 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_containsConst___boxed(lean_ob static lean_object* l___private_Lean_Compiler_LCNF_Testing_0__Lean_Compiler_LCNF_Testing_assertAfterTest___elambda__1___lambda__2___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Testing_0__Lean_Compiler_LCNF_Testing_assertAroundTest(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Compiler_LCNF_AltCore_getCode(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Testing_assertIsAtFixPoint___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_Compiler_LCNF_Testing_0__Lean_Compiler_LCNF_Testing_throwFixPointError(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Testing_getDecls___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_Compiler_LCNF_Testing_assertNoFun___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6719_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Testing_assert(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addTrace___at___private_Lean_Compiler_LCNF_Testing_0__Lean_Compiler_LCNF_Testing_assertAfterTest___elambda__1___spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Testing_TestM_run___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -130,6 +128,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_containsConst_goExpr(lean_obj static lean_object* l_Lean_Compiler_LCNF_Testing_assertIsAtFixPoint___lambda__2___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Testing_0__Lean_Compiler_LCNF_Testing_assertAfterTest___elambda__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_Testing_assertNoFun___spec__3___lambda__1___closed__3; +lean_object* l_Lean_Compiler_LCNF_Alt_getCode(lean_object*); static lean_object* l_Lean_Compiler_LCNF_Testing_assertIsAtFixPoint___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Testing_assertReducesSize___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_throwError___at_Lean_Compiler_LCNF_Testing_assert___spec__1___closed__2; @@ -208,6 +207,7 @@ lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Testing_assertForEachDeclAfter___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6660_(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Testing_assertIsAtFixPoint___lambda__2___closed__2; static lean_object* l___private_Lean_Compiler_LCNF_Testing_0__Lean_Compiler_LCNF_Testing_assertAfterTest___elambda__1___closed__2; static lean_object* l_Lean_Compiler_LCNF_Testing_assertPreservesSize___closed__1; @@ -374,7 +374,7 @@ if (x_5 == 0) { lean_object* x_6; lean_object* x_7; uint8_t x_8; x_6 = lean_array_uget(x_2, x_3); -x_7 = l_Lean_Compiler_LCNF_AltCore_getCode(x_6); +x_7 = l_Lean_Compiler_LCNF_Alt_getCode(x_6); lean_dec(x_6); x_8 = l_Lean_Compiler_LCNF_Code_containsConst(x_1, x_7); lean_dec(x_7); @@ -3357,7 +3357,7 @@ x_11 = lean_nat_sub(x_6, x_10); lean_dec(x_6); x_12 = lean_array_fget(x_4, x_11); x_13 = lean_array_fget(x_5, x_11); -x_14 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6719_(x_12, x_13); +x_14 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_6660_(x_12, x_13); if (x_14 == 0) { uint8_t x_15; diff --git a/stage0/stdlib/Lean/Compiler/LCNF/ToExpr.c b/stage0/stdlib/Lean/Compiler/LCNF/ToExpr.c index 33f53fdd0a4a..70e952c01cf6 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/ToExpr.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/ToExpr.c @@ -15,8 +15,8 @@ extern "C" { #endif LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_ToExpr_0__Lean_Expr_abstract_x27_go___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToExpr_withParams_go___at_Lean_Compiler_LCNF_FunDecl_toExprM___spec__2___boxed(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_ReaderT_bind___at_Lean_Compiler_LCNF_FunDeclCore_toExprM___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToExpr_run(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToExpr_run_x27(lean_object*); static lean_object* l_Lean_Compiler_LCNF_Code_toExprM___closed__4; @@ -29,8 +29,8 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToExpr_mkLambdaM(lean_object*, lea lean_object* l_Lean_Expr_proj___override(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_bvar___override(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_ToExpr_0__Lean_FVarId_toExpr___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_LCNF_FunDeclCore_toExprM___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_ToExpr_0__Lean_Expr_abstract_x27(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_LCNF_FunDecl_toExprM___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Code_toExprM___closed__3; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToExpr_withParams(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_Code_toExprM___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -40,9 +40,10 @@ size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_Code_toExprM___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToExpr_abstractM___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_ToExpr_run_x27___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDecl_toExpr(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToExpr_withParams_go___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDecl_toExpr___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_ToExpr_run_x27___spec__1(lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToExpr_withParams_go___at_Lean_Compiler_LCNF_FunDeclCore_toExprM___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_ToExpr_0__Lean_FVarId_toExprM(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_forallE___override(lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_Code_toExprM___spec__3(size_t, size_t, lean_object*, lean_object*, lean_object*); @@ -58,7 +59,6 @@ lean_object* l_Lean_Expr_letE___override(lean_object*, lean_object*, lean_object lean_object* l_Lean_RBNode_insert___at_Lean_FVarIdMap_insert___spec__1___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_toExprM(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_toExprM(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToExpr_mkLambdaM___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToExpr_mkLambdaM_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToExpr_withParams___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -68,6 +68,7 @@ lean_object* l_Array_append___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToExpr_run___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_toExpr___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Code_toExprM___closed__6; +LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_LCNF_FunDecl_toExprM___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToExpr_withFVar___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_ToExpr_0__Lean_Expr_abstract_x27___boxed(lean_object*, lean_object*, lean_object*); @@ -76,17 +77,16 @@ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToExpr_withParams_go___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Arg_toExpr(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_ToExpr_0__Lean_FVarId_toExpr(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToExpr_withParams_go___at_Lean_Compiler_LCNF_FunDeclCore_toExprM___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); extern lean_object* l_Lean_Compiler_LCNF_instInhabitedParam; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToExpr_withParams_go___at_Lean_Compiler_LCNF_Code_toExprM___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_mk(lean_object*); uint8_t l_Lean_Name_quickCmp(lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_toExpr___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToExpr_withParams_go(lean_object*); lean_object* lean_array_uget(lean_object*, size_t); lean_object* l_Lean_Expr_fvar___override(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDecl_toExprM(lean_object*, lean_object*, lean_object*); size_t lean_array_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at___private_Lean_Compiler_LCNF_ToExpr_0__Lean_FVarId_toExpr___spec__1(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); @@ -95,8 +95,8 @@ uint8_t lean_nat_dec_le(lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToExpr_withParams___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_toExpr(lean_object*, lean_object*); lean_object* l_Lean_Expr_lam___override(lean_object*, lean_object*, lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToExpr_withParams_go___at_Lean_Compiler_LCNF_FunDecl_toExprM___spec__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_EXPORT lean_object* l_Lean_Compiler_LCNF_ToExpr_abstractM(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToExpr_mkLambdaM_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -682,7 +682,7 @@ lean_dec(x_2); return x_4; } } -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_LCNF_FunDeclCore_toExprM___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_LCNF_FunDecl_toExprM___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; @@ -697,15 +697,15 @@ x_8 = lean_apply_3(x_2, x_6, x_3, x_7); return x_8; } } -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_LCNF_FunDeclCore_toExprM___spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_LCNF_FunDecl_toExprM___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_Compiler_LCNF_FunDeclCore_toExprM___spec__1___rarg), 4, 0); +x_3 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_LCNF_FunDecl_toExprM___spec__1___rarg), 4, 0); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToExpr_withParams_go___at_Lean_Compiler_LCNF_FunDeclCore_toExprM___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToExpr_withParams_go___at_Lean_Compiler_LCNF_FunDecl_toExprM___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; @@ -725,7 +725,7 @@ lean_object* x_12; lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_12 = l_ReaderT_bind___at_Lean_Compiler_LCNF_FunDeclCore_toExprM___spec__1___rarg(x_8, x_9, x_5, x_6); +x_12 = l_ReaderT_bind___at_Lean_Compiler_LCNF_FunDecl_toExprM___spec__1___rarg(x_8, x_9, x_5, x_6); return x_12; } else @@ -751,7 +751,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_toExprM(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDecl_toExprM(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; @@ -759,7 +759,7 @@ x_4 = lean_ctor_get(x_1, 2); lean_inc(x_4); x_5 = lean_unsigned_to_nat(0u); lean_inc(x_4); -x_6 = l_Lean_Compiler_LCNF_ToExpr_withParams_go___at_Lean_Compiler_LCNF_FunDeclCore_toExprM___spec__2(x_1, x_4, x_4, x_5, x_2, x_3); +x_6 = l_Lean_Compiler_LCNF_ToExpr_withParams_go___at_Lean_Compiler_LCNF_FunDecl_toExprM___spec__2(x_1, x_4, x_4, x_5, x_2, x_3); lean_dec(x_4); return x_6; } @@ -818,7 +818,7 @@ lean_object* x_11; lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_11 = l_ReaderT_bind___at_Lean_Compiler_LCNF_FunDeclCore_toExprM___spec__1___rarg(x_7, x_8, x_5, x_6); +x_11 = l_ReaderT_bind___at_Lean_Compiler_LCNF_FunDecl_toExprM___spec__1___rarg(x_7, x_8, x_5, x_6); return x_11; } else @@ -1173,7 +1173,7 @@ lean_inc(x_77); x_78 = l___private_Lean_Compiler_LCNF_ToExpr_0__Lean_Expr_abstract_x27_go(x_3, x_2, x_77); lean_inc(x_2); lean_inc(x_75); -x_79 = l_Lean_Compiler_LCNF_FunDeclCore_toExprM(x_75, x_2, x_3); +x_79 = l_Lean_Compiler_LCNF_FunDecl_toExprM(x_75, x_2, x_3); x_80 = lean_ctor_get(x_79, 0); lean_inc(x_80); x_81 = lean_ctor_get(x_79, 1); @@ -1222,11 +1222,11 @@ return x_97; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToExpr_withParams_go___at_Lean_Compiler_LCNF_FunDeclCore_toExprM___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_Compiler_LCNF_ToExpr_withParams_go___at_Lean_Compiler_LCNF_FunDecl_toExprM___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_Compiler_LCNF_ToExpr_withParams_go___at_Lean_Compiler_LCNF_FunDeclCore_toExprM___spec__2(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Lean_Compiler_LCNF_ToExpr_withParams_go___at_Lean_Compiler_LCNF_FunDecl_toExprM___spec__2(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_3); return x_7; } @@ -1319,7 +1319,7 @@ lean_dec(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_toExpr(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDecl_toExpr(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; @@ -1330,7 +1330,7 @@ x_6 = lean_nat_dec_lt(x_5, x_4); if (x_6 == 0) { lean_object* x_7; lean_object* x_8; -x_7 = l_Lean_Compiler_LCNF_FunDeclCore_toExprM(x_1, x_4, x_3); +x_7 = l_Lean_Compiler_LCNF_FunDecl_toExprM(x_1, x_4, x_3); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); lean_dec(x_7); @@ -1343,7 +1343,7 @@ x_9 = lean_nat_dec_le(x_4, x_4); if (x_9 == 0) { lean_object* x_10; lean_object* x_11; -x_10 = l_Lean_Compiler_LCNF_FunDeclCore_toExprM(x_1, x_4, x_3); +x_10 = l_Lean_Compiler_LCNF_FunDecl_toExprM(x_1, x_4, x_3); x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); lean_dec(x_10); @@ -1355,7 +1355,7 @@ size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_1 x_12 = 0; x_13 = lean_usize_of_nat(x_4); x_14 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_ToExpr_run_x27___spec__1(x_2, x_12, x_13, x_3); -x_15 = l_Lean_Compiler_LCNF_FunDeclCore_toExprM(x_1, x_4, x_14); +x_15 = l_Lean_Compiler_LCNF_FunDecl_toExprM(x_1, x_4, x_14); x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); lean_dec(x_15); @@ -1364,11 +1364,11 @@ return x_16; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_toExpr___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDecl_toExpr___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Compiler_LCNF_FunDeclCore_toExpr(x_1, x_2); +x_3 = l_Lean_Compiler_LCNF_FunDecl_toExpr(x_1, x_2); lean_dec(x_2); return x_3; } diff --git a/stage0/stdlib/Lean/Compiler/LCNF/ToLCNF.c b/stage0/stdlib/Lean/Compiler/LCNF/ToLCNF.c index 8ca2e68b1eba..2131265a5c1e 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/ToLCNF.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/ToLCNF.c @@ -57,6 +57,7 @@ LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at___privat LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToLCNF_bindCases_go___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_ToLCNF_toLCNF_visitApp___closed__13; static lean_object* l_Lean_Compiler_LCNF_ToLCNF_run___rarg___closed__2; +lean_object* l_Lean_Compiler_LCNF_FunDecl_etaExpand(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*); extern lean_object* l_Lean_maxRecDepthErrorMessage; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToLCNF_mkUnreachable___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -260,7 +261,6 @@ lean_object* lean_st_mk_ref(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_ToLCNF_bindCases_go___closed__4; lean_object* l_Lean_Compiler_LCNF_eraseCode(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static size_t l_Lean_PersistentHashMap_insertAux___at_Lean_Compiler_LCNF_ToLCNF_toLCNF_visitCore___spec__2___closed__1; -lean_object* l_Lean_Compiler_LCNF_AltCore_getCode(lean_object*); lean_object* l_Lean_Meta_inferType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToLCNF_toLCNF_visitNoConfusion___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_levelZero; @@ -352,6 +352,7 @@ static lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Compiler_LCNF_T LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToLCNF_toLCNF_visitAppDefaultConst(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToLCNF_toCode(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapFinIdxM_map___at_Lean_Compiler_LCNF_ToLCNF_toLCNF_visitCasesImplementedBy___spec__6___lambda__2___closed__2; +lean_object* l_Lean_Compiler_LCNF_Alt_getCode(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Compiler_LCNF_ToLCNF_toLCNF_visitCore___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_ToLCNF_bindCases_go___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_ToLCNF_bindCases_go___closed__9; @@ -527,7 +528,6 @@ lean_object* l_Lean_MessageData_ofName(lean_object*); lean_object* lean_expr_instantiate1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Compiler_LCNF_ToLCNF_0__Lean_Compiler_LCNF_ToLCNF_isTypeFormerType___spec__4(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToLCNF_toLCNF_visitCasesImplementedBy(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Compiler_LCNF_FunDeclCore_etaExpand(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToLCNF_toLCNF_visitAppArg___lambda__1___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*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToLCNF_withNewScope(lean_object*); @@ -5336,7 +5336,7 @@ lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean x_13 = lean_array_uget(x_4, x_3); x_14 = lean_unsigned_to_nat(0u); x_15 = lean_array_uset(x_4, x_3, x_14); -x_16 = l_Lean_Compiler_LCNF_AltCore_getCode(x_13); +x_16 = l_Lean_Compiler_LCNF_Alt_getCode(x_13); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); @@ -6238,7 +6238,7 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_78 = l_Lean_Compiler_LCNF_FunDeclCore_etaExpand(x_67, x_4, x_5, x_6, x_7, x_76); +x_78 = l_Lean_Compiler_LCNF_FunDecl_etaExpand(x_67, x_4, x_5, x_6, x_7, x_76); if (lean_obj_tag(x_78) == 0) { lean_object* x_79; lean_object* x_80; @@ -6296,7 +6296,7 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_87 = l_Lean_Compiler_LCNF_FunDeclCore_etaExpand(x_67, x_4, x_5, x_6, x_7, x_86); +x_87 = l_Lean_Compiler_LCNF_FunDecl_etaExpand(x_67, x_4, x_5, x_6, x_7, x_86); if (lean_obj_tag(x_87) == 0) { lean_object* x_88; lean_object* x_89; lean_object* x_90; @@ -6374,7 +6374,7 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_103 = l_Lean_Compiler_LCNF_FunDeclCore_etaExpand(x_67, x_4, x_5, x_6, x_7, x_101); +x_103 = l_Lean_Compiler_LCNF_FunDecl_etaExpand(x_67, x_4, x_5, x_6, x_7, x_101); if (lean_obj_tag(x_103) == 0) { lean_object* x_104; lean_object* x_105; lean_object* x_106; diff --git a/stage0/stdlib/Lean/Compiler/LCNF/ToMono.c b/stage0/stdlib/Lean/Compiler/LCNF/ToMono.c index 88f9c6bce057..dedab1ddf0e4 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/ToMono.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/ToMono.c @@ -36,9 +36,9 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_casesInt static lean_object* l_Lean_Compiler_LCNF_Code_toMono___closed__16; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_casesNatToMono___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_LetValue_toMono(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_FunDecl_toMono___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_Compiler_LCNF_LetValue_toMono___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Code_toMono___closed__20; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_FunDeclCore_toMono___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_casesStringToMono___closed__3; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_casesIntToMono___spec__1___closed__10; static lean_object* l_Lean_Compiler_LCNF_trivialStructToMono___closed__8; @@ -58,15 +58,14 @@ static lean_object* l_Lean_Compiler_LCNF_casesNatToMono___closed__2; static lean_object* l_Lean_Compiler_LCNF_trivialStructToMono___closed__12; lean_object* l_Lean_Compiler_LCNF_Arg_toLetValue(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_casesNatToMono___spec__1___closed__2; -lean_object* l_Lean_Compiler_LCNF_instInhabitedAltCore___rarg(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_casesIntToMono___spec__1___closed__11; lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*, uint8_t); lean_object* l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_updateFunDeclImp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_toMono(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_LetValue_toMono___closed__1; static lean_object* l_Lean_Compiler_LCNF_trivialStructToMono___closed__10; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_casesIntToMono___spec__1___closed__3; static lean_object* l_Lean_Compiler_LCNF_Code_toMono___closed__19; +extern lean_object* l_Lean_Compiler_LCNF_instInhabitedAlt; static lean_object* l_Lean_Compiler_LCNF_LetValue_toMono___closed__6; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_casesNatToMono(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_LetValue_toMono___closed__9; @@ -118,7 +117,6 @@ static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_ToMono static lean_object* l_Lean_Compiler_LCNF_LetValue_toMono___closed__8; lean_object* lean_st_mk_ref(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_casesArrayToMono___closed__1; -static lean_object* l_Lean_Compiler_LCNF_trivialStructToMono___closed__15; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_ToMono___hyg_4075_(lean_object*); static lean_object* l_Lean_Compiler_LCNF_trivialStructToMono___closed__4; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_DeclValue_mapCodeM___at_Lean_Compiler_LCNF_Decl_toMono_go___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -223,6 +221,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_casesInt static lean_object* l_Lean_Compiler_LCNF_trivialStructToMono___closed__14; 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_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_FunDecl_toMono___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_casesFloatArrayToMono___closed__3; static lean_object* l_Lean_Compiler_LCNF_casesStringToMono___closed__5; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_casesUIntToMono___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -234,7 +233,6 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_isTrivialConstructorApp_x3f(lean_o LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Param_toMono___lambda__1(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_Compiler_LCNF_casesNatToMono___closed__4; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_FunDeclCore_toMono___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_ToMono___hyg_4075____closed__18; uint8_t lean_usize_dec_lt(size_t, size_t); @@ -251,6 +249,7 @@ static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_ToMono LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_casesArrayToMono(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_LetValue_toMono___closed__3; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ctorAppToMono(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDecl_toMono(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_trivialStructToMono___closed__6; lean_object* l_ReaderT_instMonad___rarg(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_casesIntToMono___spec__1___closed__2; @@ -1678,16 +1677,19 @@ x_71 = lean_ctor_get(x_1, 1); lean_dec(x_71); x_72 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__4; x_73 = lean_name_eq(x_69, x_72); +x_73 = 0; // bootstrap: disable if (x_73 == 0) { lean_object* x_74; uint8_t x_75; x_74 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__6; x_75 = lean_name_eq(x_69, x_74); +x_75 = 0; // bootstrap: disable if (x_75 == 0) { lean_object* x_76; uint8_t x_77; x_76 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__8; x_77 = lean_name_eq(x_69, x_76); +x_77 = 0; // bootstrap: disable if (x_77 == 0) { lean_object* x_78; @@ -2455,7 +2457,7 @@ lean_dec(x_2); return x_8; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_FunDeclCore_toMono___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_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_FunDecl_toMono___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) { _start: { uint8_t x_10; @@ -2523,7 +2525,7 @@ return x_25; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_toMono(lean_object* x_1, lean_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_Compiler_LCNF_FunDecl_toMono(lean_object* x_1, lean_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; @@ -2546,7 +2548,7 @@ x_13 = lean_array_size(x_12); x_14 = 0; lean_inc(x_6); lean_inc(x_5); -x_15 = l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_FunDeclCore_toMono___spec__1(x_13, x_14, x_12, x_2, x_3, x_4, x_5, x_6, x_11); +x_15 = l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_FunDecl_toMono___spec__1(x_13, x_14, x_12, x_2, x_3, x_4, x_5, x_6, x_11); if (lean_obj_tag(x_15) == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; @@ -2702,7 +2704,7 @@ x_17 = lean_array_size(x_15); x_18 = 0; lean_inc(x_8); lean_inc(x_7); -x_19 = l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_FunDeclCore_toMono___spec__1(x_17, x_18, x_15, x_4, x_5, x_6, x_7, x_8, x_9); +x_19 = l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_FunDecl_toMono___spec__1(x_17, x_18, x_15, x_4, x_5, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_19) == 0) { lean_object* x_20; lean_object* x_21; lean_object* x_22; @@ -3305,7 +3307,7 @@ lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); lean_inc(x_53); -x_55 = l_Lean_Compiler_LCNF_FunDeclCore_toMono(x_53, x_2, x_3, x_4, x_5, x_6, x_7); +x_55 = l_Lean_Compiler_LCNF_FunDecl_toMono(x_53, x_2, x_3, x_4, x_5, x_6, x_7); if (lean_obj_tag(x_55) == 0) { lean_object* x_56; lean_object* x_57; lean_object* x_58; @@ -3549,7 +3551,7 @@ lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); lean_inc(x_98); -x_100 = l_Lean_Compiler_LCNF_FunDeclCore_toMono(x_98, x_2, x_3, x_4, x_5, x_6, x_7); +x_100 = l_Lean_Compiler_LCNF_FunDecl_toMono(x_98, x_2, x_3, x_4, x_5, x_6, x_7); if (lean_obj_tag(x_100) == 0) { lean_object* x_101; lean_object* x_102; lean_object* x_103; @@ -3795,6 +3797,7 @@ x_147 = lean_ctor_get(x_143, 3); lean_inc(x_147); x_148 = l_Lean_Compiler_LCNF_Code_toMono___closed__1; x_149 = lean_name_eq(x_144, x_148); +x_149 = 0; // bootstrap: disable if (x_149 == 0) { lean_object* x_150; uint8_t x_151; @@ -4829,31 +4832,22 @@ return x_6; static lean_object* _init_l_Lean_Compiler_LCNF_trivialStructToMono___closed__7() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Compiler_LCNF_instInhabitedCode; -x_2 = l_Lean_Compiler_LCNF_instInhabitedAltCore___rarg(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Compiler_LCNF_trivialStructToMono___closed__8() { -_start: -{ lean_object* x_1; x_1 = lean_mk_string_unchecked("ctorName == info.ctorName\n ", 28, 28); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_trivialStructToMono___closed__9() { +static lean_object* _init_l_Lean_Compiler_LCNF_trivialStructToMono___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__1; -x_2 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__8; +x_2 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_trivialStructToMono___closed__10() { +static lean_object* _init_l_Lean_Compiler_LCNF_trivialStructToMono___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; @@ -4861,12 +4855,12 @@ x_1 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__4; x_2 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__5; x_3 = lean_unsigned_to_nat(209u); x_4 = lean_unsigned_to_nat(2u); -x_5 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__9; +x_5 = l_Lean_Compiler_LCNF_trivialStructToMono___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_Lean_Compiler_LCNF_trivialStructToMono___closed__11() { +static lean_object* _init_l_Lean_Compiler_LCNF_trivialStructToMono___closed__10() { _start: { lean_object* x_1; @@ -4874,17 +4868,17 @@ x_1 = lean_mk_string_unchecked("info.fieldIdx < ps.size\n ", 26, 26); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_trivialStructToMono___closed__12() { +static lean_object* _init_l_Lean_Compiler_LCNF_trivialStructToMono___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__1; -x_2 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__11; +x_2 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__10; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_trivialStructToMono___closed__13() { +static lean_object* _init_l_Lean_Compiler_LCNF_trivialStructToMono___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; @@ -4892,12 +4886,12 @@ x_1 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__4; x_2 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__5; x_3 = lean_unsigned_to_nat(210u); x_4 = lean_unsigned_to_nat(2u); -x_5 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__12; +x_5 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__11; 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_Compiler_LCNF_trivialStructToMono___closed__14() { +static lean_object* _init_l_Lean_Compiler_LCNF_trivialStructToMono___closed__13() { _start: { lean_object* x_1; @@ -4905,7 +4899,7 @@ x_1 = lean_mk_string_unchecked("unreachable code has been reached", 33, 33); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_trivialStructToMono___closed__15() { +static lean_object* _init_l_Lean_Compiler_LCNF_trivialStructToMono___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; @@ -4913,7 +4907,7 @@ x_1 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__4; x_2 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__5; x_3 = lean_unsigned_to_nat(208u); x_4 = lean_unsigned_to_nat(41u); -x_5 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__14; +x_5 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__13; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); return x_6; } @@ -4937,7 +4931,7 @@ return x_14; else { lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__7; +x_15 = l_Lean_Compiler_LCNF_instInhabitedAlt; x_16 = lean_unsigned_to_nat(0u); x_17 = lean_array_get(x_15, x_9, x_16); if (lean_obj_tag(x_17) == 0) @@ -4958,7 +4952,7 @@ if (x_22 == 0) lean_object* x_23; lean_object* x_24; lean_dec(x_20); lean_dec(x_19); -x_23 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__10; +x_23 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__9; x_24 = l_panic___at_Lean_Compiler_LCNF_trivialStructToMono___spec__1(x_23, x_3, x_4, x_5, x_6, x_7, x_8); return x_24; } @@ -4974,7 +4968,7 @@ if (x_27 == 0) lean_object* x_28; lean_object* x_29; lean_dec(x_20); lean_dec(x_19); -x_28 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__13; +x_28 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__12; x_29 = l_panic___at_Lean_Compiler_LCNF_trivialStructToMono___spec__1(x_28, x_3, x_4, x_5, x_6, x_7, x_8); return x_29; } @@ -5447,7 +5441,7 @@ else { lean_object* x_133; lean_object* x_134; lean_dec(x_17); -x_133 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__15; +x_133 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__14; x_134 = l_panic___at_Lean_Compiler_LCNF_trivialStructToMono___spec__1(x_133, x_3, x_4, x_5, x_6, x_7, x_8); return x_134; } @@ -5501,7 +5495,7 @@ x_1 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__4; x_2 = l_Lean_Compiler_LCNF_casesStringToMono___closed__1; x_3 = lean_unsigned_to_nat(197u); x_4 = lean_unsigned_to_nat(34u); -x_5 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__14; +x_5 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__13; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); return x_6; } @@ -5525,7 +5519,7 @@ return x_14; else { lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__7; +x_15 = l_Lean_Compiler_LCNF_instInhabitedAlt; x_16 = lean_unsigned_to_nat(0u); x_17 = lean_array_get(x_15, x_9, x_16); if (lean_obj_tag(x_17) == 0) @@ -6137,7 +6131,7 @@ x_1 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__4; x_2 = l_Lean_Compiler_LCNF_casesFloatArrayToMono___closed__1; x_3 = lean_unsigned_to_nat(186u); x_4 = lean_unsigned_to_nat(34u); -x_5 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__14; +x_5 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__13; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); return x_6; } @@ -6161,7 +6155,7 @@ return x_14; else { lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__7; +x_15 = l_Lean_Compiler_LCNF_instInhabitedAlt; x_16 = lean_unsigned_to_nat(0u); x_17 = lean_array_get(x_15, x_9, x_16); if (lean_obj_tag(x_17) == 0) @@ -6765,7 +6759,7 @@ x_1 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__4; x_2 = l_Lean_Compiler_LCNF_casesByteArrayToMono___closed__1; x_3 = lean_unsigned_to_nat(175u); x_4 = lean_unsigned_to_nat(34u); -x_5 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__14; +x_5 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__13; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); return x_6; } @@ -6789,7 +6783,7 @@ return x_14; else { lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__7; +x_15 = l_Lean_Compiler_LCNF_instInhabitedAlt; x_16 = lean_unsigned_to_nat(0u); x_17 = lean_array_get(x_15, x_9, x_16); if (lean_obj_tag(x_17) == 0) @@ -7393,7 +7387,7 @@ x_1 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__4; x_2 = l_Lean_Compiler_LCNF_casesArrayToMono___closed__1; x_3 = lean_unsigned_to_nat(164u); x_4 = lean_unsigned_to_nat(34u); -x_5 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__14; +x_5 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__13; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); return x_6; } @@ -7417,7 +7411,7 @@ return x_14; else { lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__7; +x_15 = l_Lean_Compiler_LCNF_instInhabitedAlt; x_16 = lean_unsigned_to_nat(0u); x_17 = lean_array_get(x_15, x_9, x_16); if (lean_obj_tag(x_17) == 0) @@ -8031,7 +8025,7 @@ x_1 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__4; x_2 = l_Lean_Compiler_LCNF_casesUIntToMono___closed__1; x_3 = lean_unsigned_to_nat(153u); x_4 = lean_unsigned_to_nat(34u); -x_5 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__14; +x_5 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__13; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); return x_6; } @@ -8056,7 +8050,7 @@ return x_15; else { lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__7; +x_16 = l_Lean_Compiler_LCNF_instInhabitedAlt; x_17 = lean_unsigned_to_nat(0u); x_18 = lean_array_get(x_16, x_10, x_17); if (lean_obj_tag(x_18) == 0) @@ -13746,7 +13740,7 @@ return x_60; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_FunDeclCore_toMono___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_FunDecl_toMono___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; @@ -13754,7 +13748,7 @@ x_10 = lean_unbox_usize(x_1); lean_dec(x_1); x_11 = lean_unbox_usize(x_2); lean_dec(x_2); -x_12 = l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_FunDeclCore_toMono___spec__1(x_10, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_12 = l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_FunDecl_toMono___spec__1(x_10, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); @@ -14040,7 +14034,7 @@ x_18 = lean_array_size(x_11); x_19 = 0; lean_inc(x_6); lean_inc(x_5); -x_20 = l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_FunDeclCore_toMono___spec__1(x_18, x_19, x_11, x_2, x_3, x_4, x_5, x_6, x_17); +x_20 = l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_FunDecl_toMono___spec__1(x_18, x_19, x_11, x_2, x_3, x_4, x_5, x_6, x_17); if (lean_obj_tag(x_20) == 0) { lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; @@ -14218,7 +14212,7 @@ x_55 = lean_array_size(x_47); x_56 = 0; lean_inc(x_6); lean_inc(x_5); -x_57 = l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_FunDeclCore_toMono___spec__1(x_55, x_56, x_47, x_2, x_3, x_4, x_5, x_6, x_54); +x_57 = l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_FunDecl_toMono___spec__1(x_55, x_56, x_47, x_2, x_3, x_4, x_5, x_6, x_54); if (lean_obj_tag(x_57) == 0) { lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; @@ -14894,8 +14888,6 @@ l_Lean_Compiler_LCNF_trivialStructToMono___closed__13 = _init_l_Lean_Compiler_LC lean_mark_persistent(l_Lean_Compiler_LCNF_trivialStructToMono___closed__13); l_Lean_Compiler_LCNF_trivialStructToMono___closed__14 = _init_l_Lean_Compiler_LCNF_trivialStructToMono___closed__14(); lean_mark_persistent(l_Lean_Compiler_LCNF_trivialStructToMono___closed__14); -l_Lean_Compiler_LCNF_trivialStructToMono___closed__15 = _init_l_Lean_Compiler_LCNF_trivialStructToMono___closed__15(); -lean_mark_persistent(l_Lean_Compiler_LCNF_trivialStructToMono___closed__15); l_Lean_Compiler_LCNF_casesStringToMono___closed__1 = _init_l_Lean_Compiler_LCNF_casesStringToMono___closed__1(); lean_mark_persistent(l_Lean_Compiler_LCNF_casesStringToMono___closed__1); l_Lean_Compiler_LCNF_casesStringToMono___closed__2 = _init_l_Lean_Compiler_LCNF_casesStringToMono___closed__2(); diff --git a/stage0/stdlib/Lean/CoreM.c b/stage0/stdlib/Lean/CoreM.c index c924285437f0..dee67fa1edcf 100644 --- a/stage0/stdlib/Lean/CoreM.c +++ b/stage0/stdlib/Lean/CoreM.c @@ -21335,7 +21335,7 @@ static lean_object* _init_l_Lean_initFn____x40_Lean_CoreM___hyg_5018____closed__ _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_1 = 1; x_2 = l_Lean_initFn____x40_Lean_CoreM___hyg_5018____closed__1; x_3 = l_Lean_initFn____x40_Lean_CoreM___hyg_5018____closed__4; x_4 = lean_box(x_1); diff --git a/stage0/stdlib/Lean/Elab/Tactic/ElabTerm.c b/stage0/stdlib/Lean/Elab/Tactic/ElabTerm.c index fbaba04cb095..55c9eff7aa2d 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/ElabTerm.c +++ b/stage0/stdlib/Lean/Elab/Tactic/ElabTerm.c @@ -195,6 +195,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutTacticIncrementality___at_Lean_ LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeRevMAux___at_Lean_Elab_Tactic_evalRename___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_ElabTerm___hyg_6057____boxed(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*); +uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalConstructor(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalWithReducible_declRange__1___closed__4; lean_object* l_Lean_Meta_getMatcherInfo_x3f___at_Lean_Meta_reduceMatcher_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -806,6 +807,7 @@ static lean_object* l___private_Lean_Elab_Tactic_ElabTerm_0__Lean_Elab_Tactic_pr static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalRefine_x27_declRange__1___closed__3; uint8_t l_Array_isEmpty___rarg(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalWithUnfoldingAll__1___closed__3; +extern lean_object* l_Lean_reflBoolTrue; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_runTermElab_go___rarg(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { @@ -20454,18 +20456,20 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_2); -x_12 = l_Lean_Meta_mkDecideProof(x_2, x_7, x_8, x_9, x_10, x_11); +// x_2 = expectedType, x_1 = tacticName, x_3..x_6 = TacticM/TermElabM context + state, x_7..x_10 = MetaM/CoreM context + state, x_11 = IO.RealWorld +x_12 = l_Lean_Meta_mkDecide(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; uint64_t 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; uint8_t x_26; -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); +x_13 = lean_ctor_get(x_12, 0); // x_13 = dec +lean_inc(x_13); // x_13: rc +scope +x_14 = lean_ctor_get(x_12, 1); // IO.RealWorld lean_inc(x_14); lean_dec(x_12); -x_15 = l_Lean_Expr_appFn_x21(x_13); -x_16 = l_Lean_Expr_appArg_x21(x_15); -lean_dec(x_15); +// x_15 = l_Lean_Expr_appFn_x21(x_13); +x_16 = l_Lean_Expr_appArg_x21(x_13); // x_16 = decInst +// lean_dec(x_15); +// withAtLeastTransparency .default x_17 = lean_ctor_get(x_7, 0); lean_inc(x_17); x_18 = lean_ctor_get_uint64(x_7, sizeof(void*)*7); @@ -20514,8 +20518,9 @@ lean_ctor_set_uint8(x_37, sizeof(void*)*7 + 10, x_28); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_16); -x_38 = lean_whnf(x_16, x_37, x_8, x_9, x_10, x_14); +lean_inc(x_13); // x_13: rc +scope +whnf +x_38 = lean_whnf(x_13, x_37, x_8, x_9, x_10, x_14); // whnf dec +// x_13: rc +scope if (lean_obj_tag(x_38) == 0) { uint8_t x_39; @@ -20525,16 +20530,18 @@ if (x_39 == 0) lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; x_40 = lean_ctor_get(x_38, 0); x_41 = lean_ctor_get(x_38, 1); -x_42 = l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Tactic_ElabTerm_0__Lean_Elab_Tactic_blameDecideReductionFailure___spec__1___closed__5; -x_43 = l_Lean_Expr_isAppOf(x_40, x_42); +//x_42 = l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Tactic_ElabTerm_0__Lean_Elab_Tactic_blameDecideReductionFailure___spec__1___closed__5; +// Bool.true +x_42 = l___private_Lean_Elab_Tactic_ElabTerm_0__Lean_Elab_Tactic_elabNativeDecideCoreUnsafe___closed__11; +x_43 = l_Lean_Expr_isConstOf(x_40, x_42); if (x_43 == 0) { lean_object* x_44; lean_object* x_45; lean_free_object(x_38); -lean_dec(x_13); +lean_dec(x_16); x_44 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_44, 0, x_40); -x_45 = l_Lean_Elab_Tactic_evalDecideCore_diagnose(x_1, lean_box(0), x_2, x_16, x_44, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_41); +x_45 = l_Lean_Elab_Tactic_evalDecideCore_diagnose(x_1, lean_box(0), x_2, x_13, x_44, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_41); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -20544,14 +20551,17 @@ return x_45; else { lean_dec(x_40); -lean_dec(x_16); lean_dec(x_10); +lean_dec(x_13); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_38, 0, x_13); +lean_object* of_decide_eq_true = l___private_Lean_Elab_Tactic_ElabTerm_0__Lean_Elab_Tactic_elabNativeDecideCoreUnsafe___closed__14; +// recall: x_2 = expected type, x_13 = dec, x_16 = decInst +lean_object* myConst = l_Lean_Expr_const___override(of_decide_eq_true, lean_box(0)); +lean_object* out = l_Lean_mkApp3(myConst, x_2, x_16, l_Lean_reflBoolTrue); +lean_ctor_set(x_38, 0, out); return x_38; } } @@ -20563,8 +20573,10 @@ x_47 = lean_ctor_get(x_38, 1); lean_inc(x_47); lean_inc(x_46); lean_dec(x_38); -x_48 = l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Tactic_ElabTerm_0__Lean_Elab_Tactic_blameDecideReductionFailure___spec__1___closed__5; -x_49 = l_Lean_Expr_isAppOf(x_46, x_48); +//x_48 = l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Tactic_ElabTerm_0__Lean_Elab_Tactic_blameDecideReductionFailure___spec__1___closed__5; +// Bool.true +x_48 = l___private_Lean_Elab_Tactic_ElabTerm_0__Lean_Elab_Tactic_elabNativeDecideCoreUnsafe___closed__11; +x_49 = l_Lean_Expr_isConstOf(x_46, x_48); if (x_49 == 0) { lean_object* x_50; lean_object* x_51; @@ -20582,15 +20594,19 @@ else { lean_object* x_52; lean_dec(x_46); -lean_dec(x_16); +lean_dec(x_13); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_2); lean_dec(x_1); +lean_object* of_decide_eq_true = l___private_Lean_Elab_Tactic_ElabTerm_0__Lean_Elab_Tactic_elabNativeDecideCoreUnsafe___closed__14; +// recall: x_2 = expected type, x_13 = dec, x_16 = decInst +lean_object* myConst = l_Lean_Expr_const___override(of_decide_eq_true, lean_box(0)); +lean_object* out = l_Lean_mkApp3(myConst, x_2, x_16, l_Lean_reflBoolTrue); + x_52 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_52, 0, x_13); +lean_ctor_set(x_52, 0, out); lean_ctor_set(x_52, 1, x_47); return x_52; } @@ -20648,8 +20664,8 @@ lean_ctor_set_uint8(x_59, sizeof(void*)*7 + 10, x_28); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_16); -x_60 = lean_whnf(x_16, x_59, x_8, x_9, x_10, x_14); +lean_inc(x_13); +x_60 = lean_whnf(x_13, x_59, x_8, x_9, x_10, x_14); // whnf dec if (lean_obj_tag(x_60) == 0) { uint8_t x_61; @@ -20659,16 +20675,18 @@ if (x_61 == 0) lean_object* x_62; lean_object* x_63; lean_object* x_64; uint8_t x_65; x_62 = lean_ctor_get(x_60, 0); x_63 = lean_ctor_get(x_60, 1); -x_64 = l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Tactic_ElabTerm_0__Lean_Elab_Tactic_blameDecideReductionFailure___spec__1___closed__5; -x_65 = l_Lean_Expr_isAppOf(x_62, x_64); +//x_64 = l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Tactic_ElabTerm_0__Lean_Elab_Tactic_blameDecideReductionFailure___spec__1___closed__5; +// Bool.true +x_64 = l___private_Lean_Elab_Tactic_ElabTerm_0__Lean_Elab_Tactic_elabNativeDecideCoreUnsafe___closed__11; +x_65 = l_Lean_Expr_isConstOf(x_62, x_64); if (x_65 == 0) { lean_object* x_66; lean_object* x_67; lean_free_object(x_60); -lean_dec(x_13); +lean_dec(x_16); x_66 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_66, 0, x_62); -x_67 = l_Lean_Elab_Tactic_evalDecideCore_diagnose(x_1, lean_box(0), x_2, x_16, x_66, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_63); +x_67 = l_Lean_Elab_Tactic_evalDecideCore_diagnose(x_1, lean_box(0), x_2, x_13, x_66, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_63); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -20678,14 +20696,17 @@ return x_67; else { lean_dec(x_62); -lean_dec(x_16); +lean_dec(x_13); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_60, 0, x_13); +lean_object* of_decide_eq_true = l___private_Lean_Elab_Tactic_ElabTerm_0__Lean_Elab_Tactic_elabNativeDecideCoreUnsafe___closed__14; +// recall: x_2 = expected type, x_13 = dec, x_16 = decInst +lean_object* myConst = l_Lean_Expr_const___override(of_decide_eq_true, lean_box(0)); +lean_object* out = l_Lean_mkApp3(myConst, x_2, x_16, l_Lean_reflBoolTrue); +lean_ctor_set(x_60, 0, out); return x_60; } } @@ -20697,15 +20718,17 @@ x_69 = lean_ctor_get(x_60, 1); lean_inc(x_69); lean_inc(x_68); lean_dec(x_60); -x_70 = l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Tactic_ElabTerm_0__Lean_Elab_Tactic_blameDecideReductionFailure___spec__1___closed__5; -x_71 = l_Lean_Expr_isAppOf(x_68, x_70); +//x_70 = l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Tactic_ElabTerm_0__Lean_Elab_Tactic_blameDecideReductionFailure___spec__1___closed__5; +// Bool.true +x_70 = l___private_Lean_Elab_Tactic_ElabTerm_0__Lean_Elab_Tactic_elabNativeDecideCoreUnsafe___closed__11; +x_71 = l_Lean_Expr_isConstOf(x_68, x_70); if (x_71 == 0) { lean_object* x_72; lean_object* x_73; -lean_dec(x_13); +lean_dec(x_16); x_72 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_72, 0, x_68); -x_73 = l_Lean_Elab_Tactic_evalDecideCore_diagnose(x_1, lean_box(0), x_2, x_16, x_72, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_69); +x_73 = l_Lean_Elab_Tactic_evalDecideCore_diagnose(x_1, lean_box(0), x_2, x_13, x_72, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_69); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -20716,15 +20739,19 @@ else { lean_object* x_74; lean_dec(x_68); -lean_dec(x_16); +lean_dec(x_13); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_2); lean_dec(x_1); +lean_object* of_decide_eq_true = l___private_Lean_Elab_Tactic_ElabTerm_0__Lean_Elab_Tactic_elabNativeDecideCoreUnsafe___closed__14; +// recall: x_2 = expected type, x_13 = dec, x_16 = decInst +lean_object* myConst = l_Lean_Expr_const___override(of_decide_eq_true, lean_box(0)); +lean_object* out = l_Lean_mkApp3(myConst, x_2, x_16, l_Lean_reflBoolTrue); + x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_13); +lean_ctor_set(x_74, 0, out); lean_ctor_set(x_74, 1, x_69); return x_74; } @@ -20830,8 +20857,8 @@ lean_ctor_set_uint8(x_107, sizeof(void*)*7 + 10, x_80); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_16); -x_108 = lean_whnf(x_16, x_107, x_8, x_9, x_10, x_14); +lean_inc(x_13); +x_108 = lean_whnf(x_13, x_107, x_8, x_9, x_10, x_14); // whnf dec if (lean_obj_tag(x_108) == 0) { lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; uint8_t x_113; @@ -20847,16 +20874,18 @@ if (lean_is_exclusive(x_108)) { lean_dec_ref(x_108); x_111 = lean_box(0); } -x_112 = l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Tactic_ElabTerm_0__Lean_Elab_Tactic_blameDecideReductionFailure___spec__1___closed__5; -x_113 = l_Lean_Expr_isAppOf(x_109, x_112); +//x_112 = l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Tactic_ElabTerm_0__Lean_Elab_Tactic_blameDecideReductionFailure___spec__1___closed__5; +// Bool.true +x_112 = l___private_Lean_Elab_Tactic_ElabTerm_0__Lean_Elab_Tactic_elabNativeDecideCoreUnsafe___closed__11; +x_113 = l_Lean_Expr_isConstOf(x_109, x_112); if (x_113 == 0) { lean_object* x_114; lean_object* x_115; lean_dec(x_111); -lean_dec(x_13); +lean_dec(x_16); x_114 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_114, 0, x_109); -x_115 = l_Lean_Elab_Tactic_evalDecideCore_diagnose(x_1, lean_box(0), x_2, x_16, x_114, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_110); +x_115 = l_Lean_Elab_Tactic_evalDecideCore_diagnose(x_1, lean_box(0), x_2, x_13, x_114, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_110); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -20867,19 +20896,22 @@ else { lean_object* x_116; lean_dec(x_109); -lean_dec(x_16); +lean_dec(x_13); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_2); lean_dec(x_1); +lean_object* of_decide_eq_true = l___private_Lean_Elab_Tactic_ElabTerm_0__Lean_Elab_Tactic_elabNativeDecideCoreUnsafe___closed__14; +// recall: x_2 = expected type, x_13 = dec, x_16 = decInst +lean_object* myConst = l_Lean_Expr_const___override(of_decide_eq_true, lean_box(0)); +lean_object* out = l_Lean_mkApp3(myConst, x_2, x_16, l_Lean_reflBoolTrue); if (lean_is_scalar(x_111)) { x_116 = lean_alloc_ctor(0, 2, 0); } else { x_116 = x_111; } -lean_ctor_set(x_116, 0, x_13); +lean_ctor_set(x_116, 0, out); lean_ctor_set(x_116, 1, x_110); return x_116; } @@ -20956,8 +20988,8 @@ lean_ctor_set_uint8(x_124, sizeof(void*)*7 + 10, x_80); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_16); -x_125 = lean_whnf(x_16, x_124, x_8, x_9, x_10, x_14); +lean_inc(x_13); +x_125 = lean_whnf(x_13, x_124, x_8, x_9, x_10, x_14); // whnf dec if (lean_obj_tag(x_125) == 0) { lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; uint8_t x_130; @@ -20973,16 +21005,18 @@ if (lean_is_exclusive(x_125)) { lean_dec_ref(x_125); x_128 = lean_box(0); } -x_129 = l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Tactic_ElabTerm_0__Lean_Elab_Tactic_blameDecideReductionFailure___spec__1___closed__5; -x_130 = l_Lean_Expr_isAppOf(x_126, x_129); +//x_129 = l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Tactic_ElabTerm_0__Lean_Elab_Tactic_blameDecideReductionFailure___spec__1___closed__5; +// Bool.true +x_129 = l___private_Lean_Elab_Tactic_ElabTerm_0__Lean_Elab_Tactic_elabNativeDecideCoreUnsafe___closed__11; +x_130 = l_Lean_Expr_isConstOf(x_126, x_129); if (x_130 == 0) { lean_object* x_131; lean_object* x_132; lean_dec(x_128); -lean_dec(x_13); +lean_dec(x_16); x_131 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_131, 0, x_126); -x_132 = l_Lean_Elab_Tactic_evalDecideCore_diagnose(x_1, lean_box(0), x_2, x_16, x_131, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_127); +x_132 = l_Lean_Elab_Tactic_evalDecideCore_diagnose(x_1, lean_box(0), x_2, x_13, x_131, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_127); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -20993,19 +21027,22 @@ else { lean_object* x_133; lean_dec(x_126); -lean_dec(x_16); +lean_dec(x_13); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_2); lean_dec(x_1); +lean_object* of_decide_eq_true = l___private_Lean_Elab_Tactic_ElabTerm_0__Lean_Elab_Tactic_elabNativeDecideCoreUnsafe___closed__14; +// recall: x_2 = expected type, x_13 = dec, x_16 = decInst +lean_object* myConst = l_Lean_Expr_const___override(of_decide_eq_true, lean_box(0)); +lean_object* out = l_Lean_mkApp3(myConst, x_2, x_16, l_Lean_reflBoolTrue); if (lean_is_scalar(x_128)) { x_133 = lean_alloc_ctor(0, 2, 0); } else { x_133 = x_128; } -lean_ctor_set(x_133, 0, x_13); +lean_ctor_set(x_133, 0, out); lean_ctor_set(x_133, 1, x_127); return x_133; } diff --git a/stage0/stdlib/Lean/Meta/Check.c b/stage0/stdlib/Lean/Meta/Check.c index f6647e6711fa..7dbb5a7b6197 100644 --- a/stage0/stdlib/Lean/Meta/Check.c +++ b/stage0/stdlib/Lean/Meta/Check.c @@ -14,11 +14,9 @@ extern "C" { #endif static lean_object* l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__11; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__10; static lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__4___closed__4; 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 lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904_(lean_object*); 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_throwLetTypeMismatchMessage___rarg___closed__8; lean_object* l_Lean_mkAppN(lean_object*, lean_object*); @@ -34,14 +32,12 @@ lean_object* lean_whnf(lean_object*, lean_object*, lean_object*, lean_object*, l LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_addPPExplicitToExposeDiff_visit___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__3___closed__3; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__9; lean_object* l_Lean_Meta_isProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_indentD(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkLambdaLet___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_withLocalDecl___at_Lean_Meta_addPPExplicitToExposeDiff_visit___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_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_throwLetTypeMismatchMessage___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Exception_isInterrupt(lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__1; lean_object* l_Lean_Meta_isExprDefEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_uint64_to_usize(uint64_t); @@ -50,8 +46,10 @@ static lean_object* l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__10; static lean_object* l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__3; LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__2(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Check_0__Lean_Meta_checkConstant(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit(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_initFn____x40_Lean_Meta_Check___hyg_4957____closed__10; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkForall___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___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_usize_dec_eq(size_t, size_t); @@ -78,17 +76,22 @@ lean_object* l_Lean_Expr_setPPUniverses(lean_object*, uint8_t); lean_object* l_Nat_nextPowerOfTwo_go(lean_object*, lean_object*, lean_object*); 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_initFn____x40_Lean_Meta_Check___hyg_4957____closed__12; 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_throwAppTypeMismatch___rarg___lambda__1___closed__11; 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*); 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_4904____closed__7; +static lean_object* l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__9; 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_4957____closed__14; 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*); +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_throwAppTypeMismatch___spec__2(lean_object*, lean_object*, size_t, size_t); size_t lean_ptr_addr(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_isTypeCorrect(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__5; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_throwAppTypeMismatch___spec__1(lean_object*); static lean_object* l_Lean_Meta_addPPExplicitToExposeDiff___closed__3; size_t lean_usize_of_nat(lean_object*); @@ -96,11 +99,14 @@ lean_object* lean_expr_abstract(lean_object*, lean_object*); 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_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*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__2; lean_object* lean_st_ref_take(lean_object*, lean_object*); +uint8_t lean_expr_eqv(lean_object*, lean_object*); static lean_object* l_Lean_Meta_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_4957____closed__9; lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); 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*); @@ -114,15 +120,15 @@ LEAN_EXPORT lean_object* l_Lean_Meta_check(lean_object*, lean_object*, lean_obje LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__1(lean_object*); lean_object* l_Lean_Meta_getLevel(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_throwAppTypeMismatch___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__8; lean_object* l_Lean_Expr_forallE___override(lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_throwLetTypeMismatchMessage___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_checkEmoji; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__13; lean_object* lean_st_ref_get(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_throwAppTypeMismatch(lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); static lean_object* l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__4; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__12; 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_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_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_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*); @@ -131,11 +137,13 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_throwAppTypeMismatch__ extern lean_object* l_Lean_levelZero; 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_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__3; 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_addPPExplicitToExposeDiff_visit___lambda__4___closed__5; +static lean_object* l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__10; 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*); LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___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*); @@ -143,14 +151,15 @@ lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__4___closed__3; uint8_t lean_expr_equal(lean_object*, lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_throwLetTypeMismatchMessage___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__8; lean_object* l_Lean_Meta_saveState___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkForall___spec__2(lean_object*); static lean_object* l_Lean_Meta_checkProj___closed__3; 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*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__11; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__7; lean_object* l_Lean_Meta_throwIncorrectNumberOfLevels___rarg(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___private_Lean_Meta_Check_0__Lean_Meta_checkAux_check___spec__6(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_check___closed__2; @@ -166,19 +175,18 @@ 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*); extern lean_object* l_Lean_crossEmoji; +static lean_object* l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__7; 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___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*); 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_initFn____x40_Lean_Meta_Check___hyg_4904____closed__13; static lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___closed__1; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__8; lean_object* lean_array_fget(lean_object*, lean_object*); static lean_object* l_Lean_Meta_addPPExplicitToExposeDiff___closed__4; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_check___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_check___lambda__1___closed__5; +static lean_object* l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__12; 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*); @@ -189,7 +197,6 @@ lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta static lean_object* l_Lean_Meta_check___lambda__1___closed__4; static lean_object* l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__5; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__4; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkLambdaLet___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_indentExpr(lean_object*); uint8_t l_Lean_BinderInfo_isExplicit(uint8_t); @@ -211,6 +218,7 @@ uint8_t l_Lean_KVMap_getBool(lean_object*, lean_object*, uint8_t); lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_type(lean_object*); uint64_t lean_uint64_xor(uint64_t, uint64_t); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__4; lean_object* lean_panic_fn(lean_object*, lean_object*); lean_object* l_Lean_Expr_setOption___at_Lean_Expr_setPPExplicit___spec__1(lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Meta_check___lambda__1___closed__6; @@ -226,7 +234,7 @@ size_t lean_usize_sub(size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_Check_0__Lean_Meta_checkConstant___boxed(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_setPPNumericTypes(lean_object*, uint8_t); -LEAN_EXPORT lean_object* l_Lean_Meta_throwAppTypeMismatch___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_throwAppTypeMismatch___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_checkApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__12; static lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__3___closed__5; @@ -236,12 +244,13 @@ size_t lean_usize_add(size_t, size_t); static lean_object* l_Lean_Meta_mkHasTypeButIsExpectedMsg___closed__4; lean_object* lean_array_uget(lean_object*, size_t); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkLambdaLet___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__14; 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_4904____closed__5; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__11; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__6; +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_throwAppTypeMismatch___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); 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*); @@ -258,7 +267,6 @@ lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_obje uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_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_object*, lean_object*); lean_object* l_Lean_getConstVal___at___private_Lean_Meta_InferType_0__Lean_Meta_inferConstType___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__2; 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*); @@ -280,6 +288,7 @@ 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*); LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_check___spec__3(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957_(lean_object*); size_t lean_usize_land(size_t, size_t); uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_1324____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_check___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -8580,6 +8589,40 @@ x_2 = lean_alloc_closure((void*)(l_Lean_throwError___at_Lean_Meta_throwAppTypeMi return x_2; } } +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_throwAppTypeMismatch___spec__2(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_3, x_4); +if (x_5 == 0) +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_array_uget(x_2, x_3); +x_7 = lean_expr_eqv(x_6, x_1); +lean_dec(x_6); +if (x_7 == 0) +{ +size_t x_8; size_t x_9; +x_8 = 1; +x_9 = lean_usize_add(x_3, x_8); +x_3 = x_9; +goto _start; +} +else +{ +uint8_t x_11; +x_11 = 1; +return x_11; +} +} +else +{ +uint8_t x_12; +x_12 = 0; +return x_12; +} +} +} static lean_object* _init_l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__1() { _start: { @@ -8601,7 +8644,7 @@ static lean_object* _init_l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___ _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("\nthe final argument", 19, 19); +x_1 = lean_mk_string_unchecked("\nthe ", 5, 5); return x_1; } } @@ -8618,7 +8661,7 @@ static lean_object* _init_l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___ _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("\n", 1, 1); +x_1 = lean_mk_string_unchecked("argument", 8, 8); return x_1; } } @@ -8631,126 +8674,286 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +static lean_object* _init_l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__7() { _start: { -lean_object* x_10; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("\n", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__7; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("last", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__9; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("\nargument ", 10, 10); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__11; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_26; +lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_5); lean_inc(x_1); -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_26 = lean_infer_type(x_1, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_26) == 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_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; 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; +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_unsigned_to_nat(0u); +x_30 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_29); +x_31 = l_Lean_Meta_addPPExplicitToExposeDiff_visit___closed__1; +lean_inc(x_30); +x_32 = lean_mk_array(x_30, x_31); +x_33 = lean_unsigned_to_nat(1u); +x_34 = lean_nat_sub(x_30, x_33); +lean_dec(x_30); +x_35 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_2, x_32, x_34); +x_36 = lean_array_get_size(x_35); +x_37 = lean_nat_dec_lt(x_29, x_36); +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_mkHasTypeButIsExpectedMsg(x_11, x_2, x_5, x_6, x_7, x_8, x_12); -if (lean_obj_tag(x_13) == 0) +x_38 = l_Lean_Meta_mkHasTypeButIsExpectedMsg(x_27, x_3, x_6, x_7, x_8, x_9, x_28); +if (x_37 == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -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_indentExpr(x_3); -x_17 = l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__2; -x_18 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_16); -x_19 = l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__4; -x_20 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_20, 0, x_18); -lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_indentExpr(x_1); -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_throwAppTypeMismatch___rarg___lambda__1___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 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_14); -x_26 = l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__12; -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_throwAppTypeMismatch___spec__1___rarg(x_27, x_5, x_6, x_7, x_8, x_15); +lean_object* x_52; +lean_dec(x_36); +lean_dec(x_35); +x_52 = lean_box(0); +x_39 = x_52; +goto block_51; +} +else +{ +size_t x_53; size_t x_54; uint8_t x_55; +x_53 = 0; +x_54 = lean_usize_of_nat(x_36); +lean_dec(x_36); +x_55 = l_Array_anyMUnsafe_any___at_Lean_Meta_throwAppTypeMismatch___spec__2(x_1, x_35, x_53, x_54); +lean_dec(x_35); +if (x_55 == 0) +{ +lean_object* x_56; +x_56 = lean_box(0); +x_39 = x_56; +goto block_51; +} +else +{ +if (lean_obj_tag(x_38) == 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; +x_57 = lean_ctor_get(x_38, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_38, 1); +lean_inc(x_58); +lean_dec(x_38); +x_59 = l_Lean_indentExpr(x_1); +x_60 = l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__10; +x_61 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_59); +x_62 = l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__12; +x_63 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_63, 0, x_61); +lean_ctor_set(x_63, 1, x_62); +x_11 = x_63; +x_12 = x_57; +x_13 = x_58; +goto block_25; +} +else +{ +uint8_t x_64; +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -return x_28; +lean_dec(x_4); +lean_dec(x_1); +x_64 = !lean_is_exclusive(x_38); +if (x_64 == 0) +{ +return x_38; } else { -uint8_t x_29; +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_38, 0); +x_66 = lean_ctor_get(x_38, 1); +lean_inc(x_66); +lean_inc(x_65); +lean_dec(x_38); +x_67 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_67, 0, x_65); +lean_ctor_set(x_67, 1, x_66); +return x_67; +} +} +} +} +block_51: +{ +lean_dec(x_39); +if (lean_obj_tag(x_38) == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_40 = lean_ctor_get(x_38, 0); +lean_inc(x_40); +x_41 = lean_ctor_get(x_38, 1); +lean_inc(x_41); +lean_dec(x_38); +x_42 = l_Lean_indentExpr(x_1); +x_43 = l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___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_throwAppTypeMismatch___rarg___lambda__1___closed__8; +x_46 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +x_11 = x_46; +x_12 = x_40; +x_13 = x_41; +goto block_25; +} +else +{ +uint8_t 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_3); +lean_dec(x_4); lean_dec(x_1); -x_29 = !lean_is_exclusive(x_13); -if (x_29 == 0) +x_47 = !lean_is_exclusive(x_38); +if (x_47 == 0) { -return x_13; +return x_38; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_13, 0); -x_31 = lean_ctor_get(x_13, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_13); -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_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_33; +uint8_t 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); lean_dec(x_1); -x_33 = !lean_is_exclusive(x_10); -if (x_33 == 0) +x_68 = !lean_is_exclusive(x_26); +if (x_68 == 0) { -return x_10; +return x_26; } else { -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_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_object* x_69; lean_object* x_70; lean_object* x_71; +x_69 = lean_ctor_get(x_26, 0); +x_70 = lean_ctor_get(x_26, 1); +lean_inc(x_70); +lean_inc(x_69); +lean_dec(x_26); +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; } } +block_25: +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_14 = l_Lean_indentExpr(x_4); +x_15 = l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__2; +x_16 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_14); +x_17 = l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__4; +x_18 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +x_19 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_11); +x_20 = l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__12; +x_21 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +x_22 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_12); +x_23 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_20); +x_24 = l_Lean_throwError___at_Lean_Meta_throwAppTypeMismatch___spec__1___rarg(x_23, 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_24; +} } } LEAN_EXPORT lean_object* l_Lean_Meta_throwAppTypeMismatch___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) { @@ -8777,6 +8980,7 @@ x_12 = lean_ctor_get(x_9, 1); lean_inc(x_12); lean_dec(x_9); lean_inc(x_2); +lean_inc(x_1); x_13 = l_Lean_Expr_app___override(x_1, x_2); x_14 = lean_unbox(x_12); lean_dec(x_12); @@ -8786,14 +8990,14 @@ if (x_15 == 0) lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = l_Lean_Expr_setAppPPExplicit(x_13); x_17 = lean_box(0); -x_18 = l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1(x_2, x_11, x_16, x_17, x_3, x_4, x_5, x_6, x_10); +x_18 = l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1(x_2, x_1, x_11, x_16, x_17, x_3, x_4, x_5, x_6, x_10); return x_18; } else { lean_object* x_19; lean_object* x_20; x_19 = lean_box(0); -x_20 = l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1(x_2, x_11, x_13, x_19, x_3, x_4, x_5, x_6, x_10); +x_20 = l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1(x_2, x_1, x_11, x_13, x_19, x_3, x_4, x_5, x_6, x_10); return x_20; } } @@ -8847,13 +9051,28 @@ lean_dec(x_2); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_throwAppTypeMismatch___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_10; -x_10 = l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +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); -return x_10; +x_7 = l_Array_anyMUnsafe_any___at_Lean_Meta_throwAppTypeMismatch___spec__2(x_1, x_2, x_5, x_6); +lean_dec(x_2); +lean_dec(x_1); +x_8 = lean_box(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_5); +return x_11; } } LEAN_EXPORT lean_object* l_Lean_Meta_checkApp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { @@ -16080,7 +16299,7 @@ return x_30; } } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__1() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__1() { _start: { lean_object* x_1; @@ -16088,27 +16307,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_4904____closed__2() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____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_4904____closed__1; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____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_4904____closed__3() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____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_4904____closed__2; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____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_4904____closed__4() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__4() { _start: { lean_object* x_1; @@ -16116,17 +16335,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_4904____closed__5() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____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_4904____closed__3; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__4; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__3; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____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_4904____closed__6() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__6() { _start: { lean_object* x_1; @@ -16134,37 +16353,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_4904____closed__7() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____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_4904____closed__5; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__6; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__5; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____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_4904____closed__8() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____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_4904____closed__7; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__1; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__7; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____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_4904____closed__9() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____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_4904____closed__8; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____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_4904____closed__10() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__10() { _start: { lean_object* x_1; @@ -16172,17 +16391,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_4904____closed__11() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____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_4904____closed__9; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__10; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__9; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____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_4904____closed__12() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__12() { _start: { lean_object* x_1; @@ -16190,33 +16409,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_4904____closed__13() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____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_4904____closed__11; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__12; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__11; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____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_4904____closed__14() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____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_4904____closed__13; -x_2 = lean_unsigned_to_nat(4904u); +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__13; +x_2 = lean_unsigned_to_nat(4957u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957_(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_4904____closed__14; +x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__14; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -16315,6 +16534,18 @@ l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__5 = _init_l_Lean_M lean_mark_persistent(l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__5); l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__6 = _init_l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__6(); lean_mark_persistent(l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__6); +l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__7 = _init_l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__7(); +lean_mark_persistent(l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__7); +l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__8 = _init_l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__8(); +lean_mark_persistent(l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__8); +l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__9 = _init_l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__9(); +lean_mark_persistent(l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__9); +l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__10 = _init_l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__10(); +lean_mark_persistent(l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__10); +l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__11 = _init_l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__11(); +lean_mark_persistent(l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__11); +l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__12 = _init_l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__12(); +lean_mark_persistent(l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__12); l_Lean_Meta_checkProj___closed__1 = _init_l_Lean_Meta_checkProj___closed__1(); lean_mark_persistent(l_Lean_Meta_checkProj___closed__1); l_Lean_Meta_checkProj___closed__2 = _init_l_Lean_Meta_checkProj___closed__2(); @@ -16355,35 +16586,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_4904____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__1(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__1); -l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__2(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__2); -l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__3(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__3); -l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__4(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__4); -l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__5(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__5); -l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__6(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__6); -l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__7 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__7(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__7); -l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__8 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__8(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__8); -l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__9 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__9(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__9); -l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__10 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__10(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__10); -l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__11 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__11(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__11); -l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__12 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__12(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__12); -l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__13 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__13(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__13); -l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__14 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__14(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904____closed__14); -if (builtin) {res = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4904_(lean_io_mk_world()); +l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__1(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__1); +l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__2(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__2); +l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__3(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__3); +l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__4(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__4); +l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__5(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__5); +l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__6(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__6); +l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__7 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__7(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__7); +l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__8 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__8(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__8); +l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__9 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__9(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__9); +l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__10 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__10(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__10); +l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__11 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__11(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__11); +l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__12 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__12(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__12); +l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__13 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__13(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__13); +l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__14 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__14(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957____closed__14); +if (builtin) {res = l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_4957_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Meta/Constructions/NoConfusion.c b/stage0/stdlib/Lean/Meta/Constructions/NoConfusion.c index 4a3393bde8b2..de5ff67abd0d 100644 --- a/stage0/stdlib/Lean/Meta/Constructions/NoConfusion.c +++ b/stage0/stdlib/Lean/Meta/Constructions/NoConfusion.c @@ -4341,13 +4341,7 @@ lean_ctor_set(x_38, 1, x_37); x_39 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_39, 0, x_28); lean_ctor_set(x_39, 1, x_38); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_28); -lean_ctor_set(x_40, 1, x_39); -x_41 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_41, 0, x_28); -lean_ctor_set(x_41, 1, x_40); -x_42 = lean_array_mk(x_41); +x_42 = lean_array_mk(x_39); x_43 = l_Lean_mkNoConfusionEnum_mkNoConfusion___lambda__1___closed__2; lean_inc(x_14); lean_inc(x_13); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Cases.c b/stage0/stdlib/Lean/Meta/Tactic/Cases.c index ff0db5f20756..c8d5f26822fd 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Cases.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Cases.c @@ -67,6 +67,7 @@ LEAN_EXPORT lean_object* l_Lean_MVarId_casesRec___lambda__1(lean_object*, lean_o LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__17___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(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_inductionCasesOn___lambda__2___closed__4; +static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_inductionCasesOn___lambda__2___closed__5; lean_object* l_Lean_Meta_FVarSubst_insert(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_sort___override(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__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -13585,14 +13586,32 @@ lean_dec(x_2); return x_8; } } +static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_inductionCasesOn___lambda__1___closed__1; +LEAN_EXPORT lean_object* _init_l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_inductionCasesOn___lambda__1___closed__1() { +lean_object* x_1 = lean_mk_string_unchecked("isFalse", 7, 7); +lean_object* x_2 = lean_mk_string_unchecked("isTrue", 6, 6); +lean_object* x_3 = l_Lean_Name_mkStr2(l_Lean_MVarId_byCasesDec___closed__2, x_1); // Decidable.isFalse +lean_object* x_4 = l_Lean_Name_mkStr2(l_Lean_MVarId_byCasesDec___closed__2, x_2); // Decidable.isTrue +lean_object* x_5 = lean_mk_empty_array_with_capacity(lean_box(2)); +lean_object* x_6 = lean_array_push(x_5, x_3); +lean_object* x_7 = lean_array_push(x_6, x_4); +return x_7; +} + LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_inductionCasesOn___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_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_14; lean_object* x_15; lean_object* x_16; +if (x_1 == lean_box(0)) { +// hackiest bootstrap hack of all times +// use `Decidable.isFalse` and `Decidable.isTrue` +x_15 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_inductionCasesOn___lambda__1___closed__1; +} else { x_14 = lean_ctor_get(x_1, 4); lean_inc(x_14); lean_dec(x_1); x_15 = lean_array_mk(x_14); +} lean_inc(x_3); x_16 = l_Lean_MVarId_induction(x_2, x_3, x_7, x_4, x_9, x_10, x_11, x_12, x_13); if (lean_obj_tag(x_16) == 0) @@ -13689,6 +13708,16 @@ x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_inductionCasesOn___lambda__2___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_MVarId_byCasesDec___closed__1; // Decidable +x_2 = lean_mk_string_unchecked("falseTrueCases", 14, 14); +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_inductionCasesOn___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) { _start: { @@ -13750,6 +13779,7 @@ return x_28; } else { +// useNatCasesAuxOn 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_25, 0); lean_inc(x_29); @@ -13759,11 +13789,24 @@ lean_dec(x_25); x_31 = lean_ctor_get(x_29, 0); lean_inc(x_31); lean_dec(x_29); +// Nat x_32 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_inductionCasesOn___lambda__2___closed__2; x_33 = lean_name_eq(x_22, x_32); -lean_dec(x_22); if (x_33 == 0) { +lean_object* decidable = l_Lean_MVarId_byCasesDec___closed__2; // ``Decidable +if (lean_name_eq(x_22, decidable)) { + lean_dec(x_22); + lean_object* falseTrueCases = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_inductionCasesOn___lambda__2___closed__5; +lean_object* x_34; lean_object* x_35; +lean_dec(x_31); +lean_dec(x_20); +x_34 = lean_box(0); +x_35 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_inductionCasesOn___lambda__1(lean_box(0), x_3, x_4, x_5, x_18, x_19, falseTrueCases, x_34, x_7, x_8, x_9, x_10, x_30); +lean_dec(x_19); +return x_35; +} else { +lean_dec(x_22); lean_object* x_34; lean_object* x_35; lean_dec(x_31); x_34 = lean_box(0); @@ -13771,8 +13814,10 @@ x_35 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_inductionCasesOn___ lean_dec(x_19); return x_35; } +} else { +lean_dec(x_22); lean_object* x_36; uint8_t x_37; uint8_t x_38; x_36 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_inductionCasesOn___lambda__2___closed__4; x_37 = 1; @@ -17911,7 +17956,7 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_29 = l_Lean_Meta_Cases_cases(x_23, x_22, x_28, x_17, x_5, x_6, x_7, x_8, x_20); +x_29 = l_Lean_Meta_Cases_cases(x_23, x_22, x_28, 1, x_5, x_6, x_7, x_8, x_20); // hack: useNatCasesAuxOn for `Decidable.falseTrueCases` if (lean_obj_tag(x_29) == 0) { lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; @@ -18099,7 +18144,7 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_69 = l_Lean_Meta_Cases_cases(x_63, x_62, x_68, x_17, x_5, x_6, x_7, x_8, x_20); +x_69 = l_Lean_Meta_Cases_cases(x_63, x_62, x_68, 1, x_5, x_6, x_7, x_8, x_20); if (lean_obj_tag(x_69) == 0) { lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; uint8_t x_74; @@ -18667,6 +18712,10 @@ 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___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_inductionCasesOn___lambda__2___closed__5 = _init_l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_inductionCasesOn___lambda__2___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_inductionCasesOn___lambda__2___closed__5); +l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_inductionCasesOn___lambda__1___closed__1 = _init_l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_inductionCasesOn___lambda__1___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_inductionCasesOn___lambda__1___closed__1); l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3903____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3903____closed__1(); lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3903____closed__1); l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3903____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3903____closed__2(); diff --git a/tests/compiler/strictOrSimp.lean b/tests/compiler/strictOrSimp.lean deleted file mode 100644 index ea135c501536..000000000000 --- a/tests/compiler/strictOrSimp.lean +++ /dev/null @@ -1,11 +0,0 @@ -partial def spin : Nat → Bool -| n => spin (n) - -@[inline] def C : Nat := 0 - -def f (b : Nat) := -if strictOr (C == 0) (spin b) then "hello" -else "world" - -def main (xs : List String) : IO Unit := -IO.println (f 0) diff --git a/tests/lean/new-compiler/1657.lean b/tests/lean/1657.lean similarity index 100% rename from tests/lean/new-compiler/1657.lean rename to tests/lean/1657.lean diff --git a/tests/lean/new-compiler/1657.lean.expected.out b/tests/lean/1657.lean.expected.out similarity index 100% rename from tests/lean/new-compiler/1657.lean.expected.out rename to tests/lean/1657.lean.expected.out diff --git a/tests/lean/4089.lean.expected.out b/tests/lean/4089.lean.expected.out index 18c8b0d2e98e..35b7ec8028d0 100644 --- a/tests/lean/4089.lean.expected.out +++ b/tests/lean/4089.lean.expected.out @@ -1,40 +1,38 @@ - -[reset_reuse] -def f (x_1 : obj) : obj := - case x_1 : obj of - Prod.mk → - let x_2 : obj := proj[0] x_1; - let x_3 : obj := proj[1] x_1; - let x_5 : obj := reset[2] x_1; - let x_4 : obj := reuse x_5 in ctor_0[Prod.mk] x_3 x_2; - ret x_4 - -[reset_reuse] -def Sigma.toProd._rarg (x_1 : obj) : obj := - case x_1 : obj of - Sigma.mk → - let x_2 : obj := proj[0] x_1; - let x_3 : obj := proj[1] x_1; - let x_5 : obj := reset[2] x_1; - let x_4 : obj := reuse x_5 in ctor_0[Prod.mk] x_2 x_3; - ret x_4 -def Sigma.toProd (x_1 : ◾) (x_2 : ◾) : obj := - let x_3 : obj := pap Sigma.toProd._rarg; - ret x_3 - -[reset_reuse] -def foo (x_1 : obj) : obj := - case x_1 : obj of - List.nil → - let x_2 : obj := ctor_0[List.nil]; - ret x_2 - List.cons → - let x_3 : obj := proj[0] x_1; - case x_3 : obj of - Prod.mk → - let x_4 : obj := proj[1] x_1; - let x_9 : obj := reset[2] x_1; - let x_5 : obj := proj[0] x_3; - let x_6 : obj := foo x_4; - let x_7 : obj := reuse x_9 in ctor_1[List.cons] x_5 x_6; - ret x_7 +[Compiler.IR] [reset_reuse] + def f (x_1 : obj) : obj := + case x_1 : obj of + Prod.mk → + let x_2 : obj := proj[0] x_1; + let x_3 : obj := proj[1] x_1; + let x_5 : obj := reset[2] x_1; + let x_4 : obj := reuse x_5 in ctor_0[Prod.mk] x_3 x_2; + ret x_4 +[Compiler.IR] [reset_reuse] + def Sigma.toProd._redArg (x_1 : obj) : obj := + case x_1 : obj of + Sigma.mk → + let x_2 : obj := proj[0] x_1; + let x_3 : obj := proj[1] x_1; + let x_5 : obj := reset[2] x_1; + let x_4 : obj := reuse x_5 in ctor_0[Prod.mk] x_2 x_3; + ret x_4 + def Sigma.toProd (x_1 : ◾) (x_2 : ◾) (x_3 : obj) : obj := + let x_4 : obj := Sigma.toProd._redArg x_3; + ret x_4 +[Compiler.IR] [reset_reuse] + def foo (x_1 : obj) : obj := + case x_1 : obj of + List.nil → + let x_2 : obj := ctor_0[List.nil]; + ret x_2 + List.cons → + let x_3 : obj := proj[0] x_1; + case x_3 : obj of + Prod.mk → + let x_4 : obj := proj[1] x_1; + let x_10 : obj := reset[2] x_1; + let x_5 : obj := proj[0] x_3; + let x_6 : obj := proj[1] x_3; + let x_7 : obj := foo x_4; + let x_8 : obj := reuse x_10 in ctor_1[List.cons] x_5 x_7; + ret x_8 diff --git a/tests/lean/4240.lean.expected.out b/tests/lean/4240.lean.expected.out index 30636e3d3017..dd2966a3df86 100644 --- a/tests/lean/4240.lean.expected.out +++ b/tests/lean/4240.lean.expected.out @@ -1,26 +1,28 @@ - -[result] -def MyOption.isSomeWithInstance._at.isSomeWithInstanceNat._spec_1 (x_1 : @& obj) : u8 := - case x_1 : obj of - MyOption.none → - let x_2 : u8 := 0; - ret x_2 - MyOption.some → - let x_3 : u8 := 1; - ret x_3 -def isSomeWithInstanceNat (x_1 : @& obj) : u8 := - let x_2 : usize := 0; - let x_3 : obj := Array.uget ◾ x_1 x_2 ◾; - let x_4 : u8 := MyOption.isSomeWithInstance._at.isSomeWithInstanceNat._spec_1 x_3; - dec x_3; - ret x_4 -def MyOption.isSomeWithInstance._at.isSomeWithInstanceNat._spec_1._boxed (x_1 : obj) : obj := - let x_2 : u8 := MyOption.isSomeWithInstance._at.isSomeWithInstanceNat._spec_1 x_1; - dec x_1; - let x_3 : obj := box x_2; - ret x_3 -def isSomeWithInstanceNat._boxed (x_1 : obj) : obj := - let x_2 : u8 := isSomeWithInstanceNat x_1; - dec x_1; - let x_3 : obj := box x_2; - ret x_3 +[Compiler.IR] [result] + def MyOption.isSomeWithInstance._at_.isSomeWithInstanceNat.spec_0 (x_1 : @& obj) : u8 := + case x_1 : obj of + MyOption.none → + let x_2 : obj := ctor_0[Bool.false]; + let x_3 : u8 := unbox x_2; + ret x_3 + MyOption.some → + let x_4 : obj := ctor_1[Bool.true]; + let x_5 : u8 := unbox x_4; + ret x_5 + def isSomeWithInstanceNat (x_1 : @& obj) : u8 := + let x_2 : obj := 0; + let x_3 : usize := USize.ofNat x_2; + let x_4 : obj := Array.uget ◾ x_1 x_3 ◾; + let x_5 : u8 := MyOption.isSomeWithInstance._at_.isSomeWithInstanceNat.spec_0 x_4; + dec x_4; + ret x_5 + def MyOption.isSomeWithInstance._at_.isSomeWithInstanceNat.spec_0._boxed (x_1 : obj) : obj := + let x_2 : u8 := MyOption.isSomeWithInstance._at_.isSomeWithInstanceNat.spec_0 x_1; + dec x_1; + let x_3 : obj := box x_2; + ret x_3 + def isSomeWithInstanceNat._boxed (x_1 : obj) : obj := + let x_2 : u8 := isSomeWithInstanceNat x_1; + dec x_1; + let x_3 : obj := box x_2; + ret x_3 diff --git a/tests/lean/496.lean.expected.out b/tests/lean/496.lean.expected.out index 56c6f4c5fb14..0954f998340d 100644 --- a/tests/lean/496.lean.expected.out +++ b/tests/lean/496.lean.expected.out @@ -1,2 +1,2 @@ -496.lean:3:4-3:8: error: failed to compile definition, consider marking it as 'noncomputable' because it depends on 'foo', and it does not have executable code -496.lean:9:4-9:8: error: failed to compile definition, consider marking it as 'noncomputable' because it depends on 'bla1', and it does not have executable code +496.lean:3:4-3:8: error: axiom 'foo' not supported by code generator; consider marking definition as 'noncomputable' +496.lean:9:4-9:8: error: failed to compile definition, compiler IR check failed at 'bla2'. Error: depends on declaration 'bla1', which has no executable code; consider marking definition as 'noncomputable' diff --git a/tests/lean/new-compiler/CompilerConstantFold.lean b/tests/lean/CompilerConstantFold.lean similarity index 100% rename from tests/lean/new-compiler/CompilerConstantFold.lean rename to tests/lean/CompilerConstantFold.lean diff --git a/tests/lean/new-compiler/CompilerConstantFold.lean.expected.out b/tests/lean/CompilerConstantFold.lean.expected.out similarity index 76% rename from tests/lean/new-compiler/CompilerConstantFold.lean.expected.out rename to tests/lean/CompilerConstantFold.lean.expected.out index 28ae0165ac6e..d8d5d24e416b 100644 --- a/tests/lean/new-compiler/CompilerConstantFold.lean.expected.out +++ b/tests/lean/CompilerConstantFold.lean.expected.out @@ -16,20 +16,20 @@ let e1 := UInt64.shiftRight e _x.5; let e2 := UInt64.shiftLeft e _x.5; let _x.6 := Nat.add a1 a2; - let _x.7 := UInt8.val b1; + let _x.7 := UInt8.toBitVec b1; let _x.8 := Nat.add _x.6 _x.7; - let _x.9 := UInt8.val b2; + let _x.9 := UInt8.toBitVec b2; let _x.10 := Nat.add _x.8 _x.9; - let _x.11 := UInt16.val c1; + let _x.11 := UInt16.toBitVec c1; let _x.12 := Nat.add _x.10 _x.11; - let _x.13 := UInt16.val c2; + let _x.13 := UInt16.toBitVec c2; let _x.14 := Nat.add _x.12 _x.13; - let _x.15 := UInt32.val d1; + let _x.15 := UInt32.toBitVec d1; let _x.16 := Nat.add _x.14 _x.15; - let _x.17 := UInt32.val d2; + let _x.17 := UInt32.toBitVec d2; let _x.18 := Nat.add _x.16 _x.17; - let _x.19 := UInt64.val e1; + let _x.19 := UInt64.toBitVec e1; let _x.20 := Nat.add _x.18 _x.19; - let _x.21 := UInt64.val e2; + let _x.21 := UInt64.toBitVec e2; let _x.22 := Nat.add _x.20 _x.21; return _x.22 diff --git a/tests/lean/new-compiler/CompilerElimDeadBranches.lean b/tests/lean/CompilerElimDeadBranches.lean similarity index 100% rename from tests/lean/new-compiler/CompilerElimDeadBranches.lean rename to tests/lean/CompilerElimDeadBranches.lean diff --git a/tests/lean/new-compiler/CompilerElimDeadBranches.lean.expected.out b/tests/lean/CompilerElimDeadBranches.lean.expected.out similarity index 91% rename from tests/lean/new-compiler/CompilerElimDeadBranches.lean.expected.out rename to tests/lean/CompilerElimDeadBranches.lean.expected.out index f1a3cdd778ef..5077eda7cf56 100644 --- a/tests/lean/new-compiler/CompilerElimDeadBranches.lean.expected.out +++ b/tests/lean/CompilerElimDeadBranches.lean.expected.out @@ -1,12 +1,12 @@ [Compiler.elimDeadBranches] Eliminating addSomeVal._redArg with #[("val.16", Lean.Compiler.LCNF.UnreachableBranches.Value.ctor `Nat.zero #[]), + ("_x.36", + Lean.Compiler.LCNF.UnreachableBranches.Value.ctor `Option.some #[Lean.Compiler.LCNF.UnreachableBranches.Value.top]), + ("val.20", Lean.Compiler.LCNF.UnreachableBranches.Value.ctor `Nat.zero #[]), + ("_x.35", Lean.Compiler.LCNF.UnreachableBranches.Value.top), ("_x.37", Lean.Compiler.LCNF.UnreachableBranches.Value.ctor `Option.some - #[Lean.Compiler.LCNF.UnreachableBranches.Value.ctor `Nat.zero #[]]), - ("val.20", Lean.Compiler.LCNF.UnreachableBranches.Value.ctor `Nat.zero #[]), - ("_x.35", Lean.Compiler.LCNF.UnreachableBranches.Value.top), - ("_x.36", - Lean.Compiler.LCNF.UnreachableBranches.Value.ctor `Option.some #[Lean.Compiler.LCNF.UnreachableBranches.Value.top])] + #[Lean.Compiler.LCNF.UnreachableBranches.Value.ctor `Nat.zero #[]])] [Compiler.elimDeadBranches] Threw away cases _x.37 branch Option.none [Compiler.elimDeadBranches] Threw away cases _x.37 branch Option.none [Compiler.elimDeadBranches] Eliminating addSomeVal with #[("x", Lean.Compiler.LCNF.UnreachableBranches.Value.top), @@ -48,20 +48,21 @@ return _x.5 [Compiler.result] size: 1 def addSomeVal x : Option Nat := let _x.1 := addSomeVal._redArg; return _x.1 [Compiler.elimDeadBranches] Eliminating monadic with #[("_x.207", - Lean.Compiler.LCNF.UnreachableBranches.Value.ctor `Except.error #[Lean.Compiler.LCNF.UnreachableBranches.Value.top]), - ("_x.211", - Lean.Compiler.LCNF.UnreachableBranches.Value.ctor `Option.some #[Lean.Compiler.LCNF.UnreachableBranches.Value.top]), - ("a.208", Lean.Compiler.LCNF.UnreachableBranches.Value.top), - ("a.206", Lean.Compiler.LCNF.UnreachableBranches.Value.top), - ("_x.205", - Lean.Compiler.LCNF.UnreachableBranches.Value.ctor `Except.error #[Lean.Compiler.LCNF.UnreachableBranches.Value.top]), + Lean.Compiler.LCNF.UnreachableBranches.Value.ctor `Except.error #[Lean.Compiler.LCNF.UnreachableBranches.Value.bot]), ("_x.91", Lean.Compiler.LCNF.UnreachableBranches.Value.ctor `Except.ok #[Lean.Compiler.LCNF.UnreachableBranches.Value.top]), - ("x", Lean.Compiler.LCNF.UnreachableBranches.Value.top), ("val.64", Lean.Compiler.LCNF.UnreachableBranches.Value.top), + ("_x.205", + Lean.Compiler.LCNF.UnreachableBranches.Value.ctor `Except.error #[Lean.Compiler.LCNF.UnreachableBranches.Value.bot]), ("val.200", Lean.Compiler.LCNF.UnreachableBranches.Value.top), + ("y", Lean.Compiler.LCNF.UnreachableBranches.Value.top), ("val.64", Lean.Compiler.LCNF.UnreachableBranches.Value.top), + ("_x.211", + Lean.Compiler.LCNF.UnreachableBranches.Value.ctor `Option.some #[Lean.Compiler.LCNF.UnreachableBranches.Value.top]), + ("a.208", Lean.Compiler.LCNF.UnreachableBranches.Value.top), + ("_x.88", Lean.Compiler.LCNF.UnreachableBranches.Value.top), + ("a.206", Lean.Compiler.LCNF.UnreachableBranches.Value.bot), ("_x.212", Lean.Compiler.LCNF.UnreachableBranches.Value.ctor `Option.some #[Lean.Compiler.LCNF.UnreachableBranches.Value.top]), - ("_x.88", Lean.Compiler.LCNF.UnreachableBranches.Value.top), ("y", Lean.Compiler.LCNF.UnreachableBranches.Value.top)] + ("x", Lean.Compiler.LCNF.UnreachableBranches.Value.top)] [Compiler.elimDeadBranches] Threw away cases _x.211 branch Option.none [Compiler.elimDeadBranches] Threw away cases _x.212 branch Option.none [Compiler.elimDeadBranches] Threw away cases _x.205 branch Except.ok diff --git a/tests/lean/new-compiler/CompilerProvokeFloatLet.lean b/tests/lean/CompilerProvokeFloatLet.lean similarity index 100% rename from tests/lean/new-compiler/CompilerProvokeFloatLet.lean rename to tests/lean/CompilerProvokeFloatLet.lean diff --git a/tests/lean/new-compiler/CompilerProvokeFloatLet.lean.expected.out b/tests/lean/CompilerProvokeFloatLet.lean.expected.out similarity index 80% rename from tests/lean/new-compiler/CompilerProvokeFloatLet.lean.expected.out rename to tests/lean/CompilerProvokeFloatLet.lean.expected.out index 9ce82cfd6fa2..1ca2bf904fc1 100644 --- a/tests/lean/new-compiler/CompilerProvokeFloatLet.lean.expected.out +++ b/tests/lean/CompilerProvokeFloatLet.lean.expected.out @@ -20,9 +20,9 @@ return c [Compiler.floatLetIn] Size of code that was pushed into arm: Lean.Compiler.LCNF.FloatLetIn.Decision.arm `Bool.false 0 [Compiler.floatLetIn] Size of code that was pushed into arm: Lean.Compiler.LCNF.FloatLetIn.Decision.arm `Bool.true 0 -[Compiler.floatLetIn] Size of code that was pushed into arm: Lean.Compiler.LCNF.FloatLetIn.Decision.arm `Nat.zero 0 -[Compiler.floatLetIn] Size of code that was pushed into arm: Lean.Compiler.LCNF.FloatLetIn.Decision.arm `Nat.succ 0 -[Compiler.floatLetIn] size: 11 +[Compiler.floatLetIn] Size of code that was pushed into arm: Lean.Compiler.LCNF.FloatLetIn.Decision.arm `Bool.true 0 +[Compiler.floatLetIn] Size of code that was pushed into arm: Lean.Compiler.LCNF.FloatLetIn.Decision.arm `Bool.false 0 +[Compiler.floatLetIn] size: 13 def provokeFloatLet x y cond : Nat := let dual := Nat.mul x y; cases cond : Nat @@ -31,18 +31,20 @@ let _x.1 := Nat.add b dual; return _x.1 | Bool.true => - cases dual : Nat - | Nat.zero => + let zero := 0; + let isZero := Nat.decEq dual zero; + cases isZero : Nat + | Bool.true => let a := Nat.pow x y; return a - | Nat.succ n.2 => + | Bool.false => let c := Nat.sub x y; return c [Compiler.floatLetIn] Size of code that was pushed into arm: Lean.Compiler.LCNF.FloatLetIn.Decision.arm `Bool.false 0 [Compiler.floatLetIn] Size of code that was pushed into arm: Lean.Compiler.LCNF.FloatLetIn.Decision.arm `Bool.true 0 -[Compiler.floatLetIn] Size of code that was pushed into arm: Lean.Compiler.LCNF.FloatLetIn.Decision.arm `Nat.zero 0 -[Compiler.floatLetIn] Size of code that was pushed into arm: Lean.Compiler.LCNF.FloatLetIn.Decision.arm `Nat.succ 0 -[Compiler.floatLetIn] size: 11 +[Compiler.floatLetIn] Size of code that was pushed into arm: Lean.Compiler.LCNF.FloatLetIn.Decision.arm `Bool.true 0 +[Compiler.floatLetIn] Size of code that was pushed into arm: Lean.Compiler.LCNF.FloatLetIn.Decision.arm `Bool.false 0 +[Compiler.floatLetIn] size: 13 def provokeFloatLet x y cond : Nat := let dual := Nat.mul x y; cases cond : Nat @@ -51,10 +53,12 @@ let _x.1 := Nat.add b dual; return _x.1 | Bool.true => - cases dual : Nat - | Nat.zero => + let zero := 0; + let isZero := Nat.decEq dual zero; + cases isZero : Nat + | Bool.true => let a := Nat.pow x y; return a - | Nat.succ n.2 => + | Bool.false => let c := Nat.sub x y; return c diff --git a/tests/lean/new-compiler/baseIO.lean b/tests/lean/baseIO.lean similarity index 100% rename from tests/lean/new-compiler/baseIO.lean rename to tests/lean/baseIO.lean diff --git a/tests/lean/new-compiler/baseIO.lean.expected.out b/tests/lean/baseIO.lean.expected.out similarity index 100% rename from tests/lean/new-compiler/baseIO.lean.expected.out rename to tests/lean/baseIO.lean.expected.out diff --git a/tests/lean/computedFieldsCode.lean.expected.out b/tests/lean/computedFieldsCode.lean.expected.out index aab58b52428b..103472481e1d 100644 --- a/tests/lean/computedFieldsCode.lean.expected.out +++ b/tests/lean/computedFieldsCode.lean.expected.out @@ -1,191 +1,236 @@ - -[result] -def Exp.casesOn._override._rarg (x_1 : obj) (x_2 : obj) (x_3 : obj) (x_4 : @& obj) (x_5 : @& obj) (x_6 : @& obj) (x_7 : @& obj) (x_8 : @& obj) : obj := - case x_1 : obj of - Exp.var._impl → - dec x_3; - let x_9 : u32 := sproj[0, 8] x_1; - dec x_1; - let x_10 : obj := box x_9; - let x_11 : obj := app x_2 x_10; - ret x_11 - Exp.app._impl → - dec x_2; - let x_12 : obj := proj[0] x_1; - inc x_12; - let x_13 : obj := proj[1] x_1; - inc x_13; - dec x_1; - let x_14 : obj := app x_3 x_12 x_13; - ret x_14 - Exp.a1._impl → - dec x_3; - dec x_2; - inc x_4; - ret x_4 - Exp.a2._impl → - dec x_3; - dec x_2; - inc x_5; - ret x_5 - Exp.a3._impl → - dec x_3; - dec x_2; - inc x_6; - ret x_6 - Exp.a4._impl → - dec x_3; - dec x_2; - inc x_7; - ret x_7 - Exp.a5._impl → - dec x_3; - dec x_2; - inc x_8; - ret x_8 -def Exp.casesOn._override (x_1 : ◾) : obj := - let x_2 : obj := pap Exp.casesOn._override._rarg._boxed; - ret x_2 -def Exp.casesOn._override._rarg._boxed (x_1 : obj) (x_2 : obj) (x_3 : obj) (x_4 : obj) (x_5 : obj) (x_6 : obj) (x_7 : obj) (x_8 : obj) : obj := - let x_9 : obj := Exp.casesOn._override._rarg x_1 x_2 x_3 x_4 x_5 x_6 x_7 x_8; - dec x_8; - dec x_7; - dec x_6; - dec x_5; - dec x_4; - ret x_9 - -[result] -def Exp.var._override (x_1 : u32) : obj := - let x_2 : u64 := UInt32.toUInt64 x_1; - let x_3 : u64 := 1000; - let x_4 : u64 := UInt64.add x_2 x_3; - let x_5 : obj := ctor_0.0.12[Exp.var._impl]; - sset x_5[0, 0] : u64 := x_4; - sset x_5[0, 8] : u32 := x_1; - ret x_5 -def Exp.app._override (x_1 : obj) (x_2 : obj) : obj := - let x_3 : u64 := Exp.hash._override x_1; - let x_4 : u64 := Exp.hash._override x_2; - let x_5 : u64 := mixHash x_3 x_4; - let x_6 : obj := ctor_1.0.8[Exp.app._impl] x_1 x_2; - sset x_6[2, 0] : u64 := x_5; - ret x_6 -def Exp.a1._override : obj := - let x_1 : obj := ctor_2[Exp.a1._impl]; - ret x_1 -def Exp.a2._override : obj := - let x_1 : obj := ctor_3[Exp.a2._impl]; - ret x_1 -def Exp.a3._override : obj := - let x_1 : obj := ctor_4[Exp.a3._impl]; - ret x_1 -def Exp.a4._override : obj := - let x_1 : obj := ctor_5[Exp.a4._impl]; - ret x_1 -def Exp.a5._override : obj := - let x_1 : obj := ctor_6[Exp.a5._impl]; - ret x_1 -def Exp.hash._override (x_1 : @& obj) : u64 := - case x_1 : obj of - Exp.var._impl → - let x_2 : u64 := sproj[0, 0] x_1; - ret x_2 - Exp.app._impl → - let x_3 : u64 := sproj[2, 0] x_1; - ret x_3 - default → - let x_4 : u64 := 42; - ret x_4 -def Exp.var._override._boxed (x_1 : obj) : obj := - let x_2 : u32 := unbox x_1; - dec x_1; - let x_3 : obj := Exp.var._override x_2; - ret x_3 -def Exp.hash._override._boxed (x_1 : obj) : obj := - let x_2 : u64 := Exp.hash._override x_1; - dec x_1; - let x_3 : obj := box x_2; - ret x_3 - -[result] -def f._closed_1 : obj := - let x_1 : u32 := 10; - let x_2 : obj := Exp.var._override x_1; - ret x_2 -def f._closed_2 : obj := - let x_1 : obj := f._closed_1; - let x_2 : obj := ctor_5[Exp.a4._impl]; - let x_3 : obj := Exp.app._override x_1 x_2; - ret x_3 -def f._closed_3 : u64 := - let x_1 : obj := f._closed_2; - let x_2 : u64 := Exp.hash._override x_1; - dec x_1; - ret x_2 -def f : u64 := - let x_1 : u64 := f._closed_3; - ret x_1 - -[result] -def g (x_1 : @& obj) : u8 := - case x_1 : obj of - Exp.a3._impl → - let x_2 : u8 := 1; - ret x_2 - default → - let x_3 : u8 := 0; - ret x_3 -def g._boxed (x_1 : obj) : obj := - let x_2 : u8 := g x_1; - dec x_1; - let x_3 : obj := box x_2; - ret x_3 - -[result] -def hash' (x_1 : @& obj) : obj := - case x_1 : obj of - Exp.var._impl → - let x_2 : u32 := sproj[0, 8] x_1; - let x_3 : obj := UInt32.toNat x_2; - ret x_3 - Exp.app._impl → - let x_4 : obj := proj[0] x_1; - let x_5 : obj := proj[1] x_1; - let x_6 : obj := hash' x_4; - let x_7 : obj := hash' x_5; - let x_8 : obj := Nat.add x_6 x_7; - dec x_7; - dec x_6; - ret x_8 - default → - let x_9 : obj := 42; - ret x_9 -def hash'._boxed (x_1 : obj) : obj := - let x_2 : obj := hash' x_1; - dec x_1; - ret x_2 - -[result] -def getAppFn (x_1 : @& obj) : obj := - case x_1 : obj of - Exp.app._impl → - let x_2 : obj := proj[0] x_1; - let x_3 : obj := getAppFn x_2; - ret x_3 - default → - inc x_1; - ret x_1 -def getAppFn._boxed (x_1 : obj) : obj := - let x_2 : obj := getAppFn x_1; - dec x_1; - ret x_2 - -[result] -def Exp.f (x_1 : @& obj) : obj := - let x_2 : obj := getAppFn x_1; - ret x_2 -def Exp.f._boxed (x_1 : obj) : obj := - let x_2 : obj := Exp.f x_1; - dec x_1; - ret x_2 +[Compiler.IR] [result] + def Exp.casesOn._override._redArg (x_1 : obj) (x_2 : obj) (x_3 : obj) (x_4 : @& obj) (x_5 : @& obj) (x_6 : @& obj) (x_7 : @& obj) (x_8 : @& obj) : obj := + case x_1 : obj of + Exp.var._impl → + dec x_3; + let x_9 : u32 := sproj[0, 8] x_1; + dec x_1; + let x_10 : obj := box x_9; + let x_11 : obj := app x_2 x_10; + ret x_11 + Exp.app._impl → + dec x_2; + let x_12 : obj := proj[0] x_1; + inc x_12; + let x_13 : obj := proj[1] x_1; + inc x_13; + dec x_1; + let x_14 : obj := app x_3 x_12 x_13; + ret x_14 + Exp.a1._impl → + dec x_3; + dec x_2; + inc x_4; + ret x_4 + Exp.a2._impl → + dec x_3; + dec x_2; + inc x_5; + ret x_5 + Exp.a3._impl → + dec x_3; + dec x_2; + inc x_6; + ret x_6 + Exp.a4._impl → + dec x_3; + dec x_2; + inc x_7; + ret x_7 + Exp.a5._impl → + dec x_3; + dec x_2; + inc x_8; + ret x_8 + def Exp.casesOn._override (x_1 : ◾) (x_2 : obj) (x_3 : obj) (x_4 : obj) (x_5 : @& obj) (x_6 : @& obj) (x_7 : @& obj) (x_8 : @& obj) (x_9 : @& obj) : obj := + case x_2 : obj of + Exp.var._impl → + dec x_4; + let x_10 : u32 := sproj[0, 8] x_2; + dec x_2; + let x_11 : obj := box x_10; + let x_12 : obj := app x_3 x_11; + ret x_12 + Exp.app._impl → + dec x_3; + let x_13 : obj := proj[0] x_2; + inc x_13; + let x_14 : obj := proj[1] x_2; + inc x_14; + dec x_2; + let x_15 : obj := app x_4 x_13 x_14; + ret x_15 + Exp.a1._impl → + dec x_4; + dec x_3; + inc x_5; + ret x_5 + Exp.a2._impl → + dec x_4; + dec x_3; + inc x_6; + ret x_6 + Exp.a3._impl → + dec x_4; + dec x_3; + inc x_7; + ret x_7 + Exp.a4._impl → + dec x_4; + dec x_3; + inc x_8; + ret x_8 + Exp.a5._impl → + dec x_4; + dec x_3; + inc x_9; + ret x_9 + def Exp.casesOn._override._redArg._boxed (x_1 : obj) (x_2 : obj) (x_3 : obj) (x_4 : obj) (x_5 : obj) (x_6 : obj) (x_7 : obj) (x_8 : obj) : obj := + let x_9 : obj := Exp.casesOn._override._redArg x_1 x_2 x_3 x_4 x_5 x_6 x_7 x_8; + dec x_8; + dec x_7; + dec x_6; + dec x_5; + dec x_4; + ret x_9 + def Exp.casesOn._override._boxed (x_1 : obj) (x_2 : obj) (x_3 : obj) (x_4 : obj) (x_5 : obj) (x_6 : obj) (x_7 : obj) (x_8 : obj) (x_9 : obj) : obj := + let x_10 : obj := Exp.casesOn._override x_1 x_2 x_3 x_4 x_5 x_6 x_7 x_8 x_9; + dec x_9; + dec x_8; + dec x_7; + dec x_6; + dec x_5; + ret x_10 +[Compiler.IR] [result] + def Exp.var._override (x_1 : u32) : obj := + let x_2 : u64 := UInt32.toUInt64 x_1; + let x_3 : obj := 1000; + let x_4 : u64 := UInt64.ofNat x_3; + let x_5 : u64 := UInt64.add x_2 x_4; + let x_6 : obj := ctor_0.0.12[Exp.var._impl]; + sset x_6[0, 0] : u64 := x_5; + sset x_6[0, 8] : u32 := x_1; + ret x_6 + def Exp.app._override (x_1 : obj) (x_2 : obj) : obj := + let x_3 : u64 := Exp.hash._override x_1; + let x_4 : u64 := Exp.hash._override x_2; + let x_5 : u64 := mixHash x_3 x_4; + let x_6 : obj := ctor_1.0.8[Exp.app._impl] x_1 x_2; + sset x_6[2, 0] : u64 := x_5; + ret x_6 + def Exp.a1._override : obj := + let x_1 : obj := ctor_2[Exp.a1._impl]; + ret x_1 + def Exp.a2._override : obj := + let x_1 : obj := ctor_3[Exp.a2._impl]; + ret x_1 + def Exp.a3._override : obj := + let x_1 : obj := ctor_4[Exp.a3._impl]; + ret x_1 + def Exp.a4._override : obj := + let x_1 : obj := ctor_5[Exp.a4._impl]; + ret x_1 + def Exp.a5._override : obj := + let x_1 : obj := ctor_6[Exp.a5._impl]; + ret x_1 + def Exp.hash._override (x_1 : @& obj) : u64 := + case x_1 : obj of + Exp.var._impl → + let x_2 : u64 := sproj[0, 0] x_1; + ret x_2 + Exp.app._impl → + let x_3 : u64 := sproj[2, 0] x_1; + ret x_3 + default → + let x_4 : obj := 42; + let x_5 : u64 := UInt64.ofNat x_4; + ret x_5 + def Exp.var._override._boxed (x_1 : obj) : obj := + let x_2 : u32 := unbox x_1; + dec x_1; + let x_3 : obj := Exp.var._override x_2; + ret x_3 + def Exp.hash._override._boxed (x_1 : obj) : obj := + let x_2 : u64 := Exp.hash._override x_1; + dec x_1; + let x_3 : obj := box x_2; + ret x_3 +[Compiler.IR] [result] + def f : u64 := + let x_1 : obj := 10; + let x_2 : u32 := UInt32.ofNat x_1; + let x_3 : obj := Exp.var._override x_2; + let x_4 : obj := ctor_5[Exp.a4._impl]; + let x_5 : obj := Exp.app._override x_3 x_4; + let x_6 : u64 := Exp.hash._override x_5; + dec x_5; + ret x_6 +[Compiler.IR] [result] + def g (x_1 : @& obj) : u8 := + let x_2 : obj := ctor_0[Bool.false]; + case x_1 : obj of + Exp.var._impl → + let x_3 : u8 := unbox x_2; + ret x_3 + Exp.app._impl → + let x_4 : u8 := unbox x_2; + ret x_4 + Exp.a3._impl → + let x_5 : obj := ctor_1[Bool.true]; + let x_6 : u8 := unbox x_5; + ret x_6 + default → + let x_7 : u8 := unbox x_2; + ret x_7 + def g._boxed (x_1 : obj) : obj := + let x_2 : u8 := g x_1; + dec x_1; + let x_3 : obj := box x_2; + ret x_3 +[Compiler.IR] [result] + def hash' (x_1 : @& obj) : obj := + case x_1 : obj of + Exp.var._impl → + let x_2 : u32 := sproj[0, 8] x_1; + let x_3 : obj := UInt32.toNat x_2; + ret x_3 + Exp.app._impl → + let x_4 : obj := proj[0] x_1; + let x_5 : obj := proj[1] x_1; + let x_6 : obj := hash' x_4; + let x_7 : obj := hash' x_5; + let x_8 : obj := Nat.add x_6 x_7; + dec x_7; + dec x_6; + ret x_8 + default → + let x_9 : obj := 42; + ret x_9 + def hash'._boxed (x_1 : obj) : obj := + let x_2 : obj := hash' x_1; + dec x_1; + ret x_2 +[Compiler.IR] [result] + def getAppFn (x_1 : @& obj) : obj := + case x_1 : obj of + Exp.app._impl → + let x_2 : obj := proj[0] x_1; + let x_3 : obj := getAppFn x_2; + ret x_3 + default → + inc x_1; + ret x_1 + def getAppFn._boxed (x_1 : obj) : obj := + let x_2 : obj := getAppFn x_1; + dec x_1; + ret x_2 +[Compiler.IR] [result] + def Exp.f (x_1 : obj) : obj := + inc x_1; + let x_2 : obj := Exp.app._override x_1 x_1; + let x_3 : obj := proj[0] x_2; + inc x_3; + dec x_2; + let x_4 : obj := getAppFn x_3; + dec x_3; + ret x_4 diff --git a/tests/lean/csimpAttr.lean.expected.out b/tests/lean/csimpAttr.lean.expected.out index a491585b8c47..11854c34409e 100644 --- a/tests/lean/csimpAttr.lean.expected.out +++ b/tests/lean/csimpAttr.lean.expected.out @@ -1,7 +1,6 @@ csimpAttr.lean:7:2-7:7: error: invalid 'csimp' theorem, only constant replacement theorems (e.g., `@f = @g`) are currently supported. - -[init] -def f (x_1 : obj) : obj := - let x_2 : obj := Nat.add x_1 x_1; - let x_3 : obj := Nat.add x_2 x_2; - ret x_3 +[Compiler.IR] [init] + def f (x_1 : obj) : obj := + let x_2 : obj := Nat.add x_1 x_1; + let x_3 : obj := Nat.add x_2 x_2; + ret x_3 diff --git a/tests/lean/csimpAttrAppend.lean.expected.out b/tests/lean/csimpAttrAppend.lean.expected.out index 2e7eac595254..ab6af56e7331 100644 --- a/tests/lean/csimpAttrAppend.lean.expected.out +++ b/tests/lean/csimpAttrAppend.lean.expected.out @@ -1,6 +1,5 @@ - -[init] -def f (x_1 : obj) (x_2 : obj) (x_3 : obj) : obj := - let x_4 : obj := List.appendTR._rarg x_1 x_2; - let x_5 : obj := List.appendTR._rarg x_4 x_3; - ret x_5 +[Compiler.IR] [init] + def f (x_1 : obj) (x_2 : obj) (x_3 : obj) : obj := + let x_4 : obj := List.appendTR._redArg x_1 x_2; + let x_5 : obj := List.appendTR._redArg x_4 x_3; + ret x_5 diff --git a/tests/lean/doubleReset.lean.expected.out b/tests/lean/doubleReset.lean.expected.out index 25937761cd2e..1ca414a74e04 100644 --- a/tests/lean/doubleReset.lean.expected.out +++ b/tests/lean/doubleReset.lean.expected.out @@ -1,37 +1,33 @@ - -[reset_reuse] -def Array.mapMUnsafe.map._at.applyProjectionRules._spec_1._rarg (x_1 : obj) (x_2 : usize) (x_3 : usize) (x_4 : obj) : obj := - let x_5 : u8 := USize.decLt x_3 x_2; - case x_5 : obj of - Bool.false → - ret x_4 - Bool.true → - let x_6 : obj := Array.uget ◾ x_4 x_3 ◾; - let x_7 : obj := 0; - let x_8 : obj := Array.uset ◾ x_4 x_3 x_7 ◾; - let x_9 : obj := proj[0] x_6; - case x_9 : obj of - Prod.mk → - case x_9 : obj of - Prod.mk → +[Compiler.IR] [reset_reuse] + def Array.mapMUnsafe.map._at_.applyProjectionRules.spec_0._redArg (x_1 : obj) (x_2 : usize) (x_3 : usize) (x_4 : obj) : obj := + let x_5 : u8 := USize.decLt x_3 x_2; + case x_5 : obj of + Bool.false → + ret x_4 + Bool.true → + let x_6 : obj := Array.uget ◾ x_4 x_3 ◾; + let x_7 : obj := ctor_0[Nat.zero]; + let x_8 : obj := Array.uset ◾ x_4 x_3 x_7 ◾; + let x_9 : obj := proj[0] x_6; let x_10 : obj := proj[0] x_9; let x_11 : obj := proj[1] x_9; - let x_18 : obj := reset[2] x_9; - let x_12 : obj := reuse x_18 in ctor_0[Prod.mk] x_10 x_11; + let x_12 : obj := ctor_0[Prod.mk] x_10 x_11; let x_13 : obj := ctor_0[Prod.mk] x_12 x_1; - let x_14 : usize := 1; - let x_15 : usize := USize.add x_3 x_14; - let x_16 : obj := Array.uset ◾ x_8 x_3 x_13 ◾; - let x_17 : obj := Array.mapMUnsafe.map._at.applyProjectionRules._spec_1._rarg x_1 x_2 x_15 x_16; - ret x_17 -def Array.mapMUnsafe.map._at.applyProjectionRules._spec_1 (x_1 : ◾) (x_2 : ◾) (x_3 : ◾) : obj := - let x_4 : obj := pap Array.mapMUnsafe.map._at.applyProjectionRules._spec_1._rarg; - ret x_4 -def applyProjectionRules._rarg (x_1 : obj) (x_2 : obj) : obj := - let x_3 : usize := Array.usize ◾ x_1; - let x_4 : usize := 0; - let x_5 : obj := Array.mapMUnsafe.map._at.applyProjectionRules._spec_1._rarg x_2 x_3 x_4 x_1; - ret x_5 -def applyProjectionRules (x_1 : ◾) (x_2 : ◾) (x_3 : ◾) : obj := - let x_4 : obj := pap applyProjectionRules._rarg; - ret x_4 + let x_14 : obj := 1; + let x_15 : usize := USize.ofNat x_14; + let x_16 : usize := USize.add x_3 x_15; + let x_17 : obj := Array.uset ◾ x_8 x_3 x_13 ◾; + let x_18 : obj := Array.mapMUnsafe.map._at_.applyProjectionRules.spec_0._redArg x_1 x_2 x_16 x_17; + ret x_18 + def Array.mapMUnsafe.map._at_.applyProjectionRules.spec_0 (x_1 : ◾) (x_2 : ◾) (x_3 : ◾) (x_4 : obj) (x_5 : usize) (x_6 : usize) (x_7 : obj) : obj := + let x_8 : obj := Array.mapMUnsafe.map._at_.applyProjectionRules.spec_0._redArg x_4 x_5 x_6 x_7; + ret x_8 + def applyProjectionRules._redArg (x_1 : obj) (x_2 : obj) : obj := + let x_3 : usize := Array.usize ◾ x_1; + let x_4 : obj := 0; + let x_5 : usize := USize.ofNat x_4; + let x_6 : obj := Array.mapMUnsafe.map._at_.applyProjectionRules.spec_0._redArg x_2 x_3 x_5 x_1; + ret x_6 + def applyProjectionRules (x_1 : ◾) (x_2 : ◾) (x_3 : ◾) (x_4 : obj) (x_5 : obj) : obj := + let x_6 : obj := applyProjectionRules._redArg x_4 x_5; + ret x_6 diff --git a/tests/lean/hidingInaccessibleNames.lean.expected.out b/tests/lean/hidingInaccessibleNames.lean.expected.out index 929513dea211..9c4e2e063ebf 100644 --- a/tests/lean/hidingInaccessibleNames.lean.expected.out +++ b/tests/lean/hidingInaccessibleNames.lean.expected.out @@ -42,7 +42,7 @@ h✝ : q hidingInaccessibleNames.lean:23:25-23:26: error: don't know how to synthesize placeholder context: p✝ : Prop -h✝ x✝ : p✝ +of_decide✝ x✝ : p✝ ⊢ decide p✝ = true hidingInaccessibleNames.lean:27:17-27:18: error: don't know how to synthesize placeholder context: diff --git a/tests/lean/new-compiler/jpCasesDiscrM.lean b/tests/lean/jpCasesDiscrM.lean similarity index 100% rename from tests/lean/new-compiler/jpCasesDiscrM.lean rename to tests/lean/jpCasesDiscrM.lean diff --git a/tests/lean/new-compiler/jpCasesDiscrM.lean.expected.out b/tests/lean/jpCasesDiscrM.lean.expected.out similarity index 100% rename from tests/lean/new-compiler/jpCasesDiscrM.lean.expected.out rename to tests/lean/jpCasesDiscrM.lean.expected.out diff --git a/tests/lean/new-compiler/jpCasesNary.lean b/tests/lean/jpCasesNary.lean similarity index 100% rename from tests/lean/new-compiler/jpCasesNary.lean rename to tests/lean/jpCasesNary.lean diff --git a/tests/lean/new-compiler/jpCasesNary.lean.expected.out b/tests/lean/jpCasesNary.lean.expected.out similarity index 100% rename from tests/lean/new-compiler/jpCasesNary.lean.expected.out rename to tests/lean/jpCasesNary.lean.expected.out diff --git a/tests/lean/new-compiler/jpClosureIssue.lean b/tests/lean/jpClosureIssue.lean similarity index 100% rename from tests/lean/new-compiler/jpClosureIssue.lean rename to tests/lean/jpClosureIssue.lean diff --git a/tests/lean/new-compiler/jpClosureIssue.lean.expected.out b/tests/lean/jpClosureIssue.lean.expected.out similarity index 100% rename from tests/lean/new-compiler/jpClosureIssue.lean.expected.out rename to tests/lean/jpClosureIssue.lean.expected.out diff --git a/tests/lean/new-compiler/lambdaLiftCache.lean b/tests/lean/lambdaLiftCache.lean similarity index 100% rename from tests/lean/new-compiler/lambdaLiftCache.lean rename to tests/lean/lambdaLiftCache.lean diff --git a/tests/lean/new-compiler/lambdaLiftCache.lean.expected.out b/tests/lean/lambdaLiftCache.lean.expected.out similarity index 51% rename from tests/lean/new-compiler/lambdaLiftCache.lean.expected.out rename to tests/lean/lambdaLiftCache.lean.expected.out index 6244068307f5..42820cb5db0a 100644 --- a/tests/lean/new-compiler/lambdaLiftCache.lean.expected.out +++ b/tests/lean/lambdaLiftCache.lean.expected.out @@ -2,13 +2,9 @@ def foo._lam_0 x x.1 : Nat := let _x.2 := Nat.add x.1 x; return _x.2 -[Compiler.saveMono] size: 2 - def foo x xs : List Nat := - let _f.1 := foo._lam_0 x; - let _x.2 := List.map._at_.map.spec_0 _f.1 xs; - return _x.2 +[Compiler.saveMono] size: 2 def foo x xs : List Nat := let _f.1 := foo._lam_0 x; let _x.2 := map _f.1 xs; return _x.2 [Compiler.saveMono] size: 2 def boo x xs : List Nat := let _f.1 := foo._lam_0 x; - let _x.2 := List.map._at_.map.spec_0 _f.1 xs; + let _x.2 := map _f.1 xs; return _x.2 diff --git a/tests/lean/lcnfTypes.lean.expected.out b/tests/lean/lcnfTypes.lean.expected.out index 23460d86686b..346b18a18231 100644 --- a/tests/lean/lcnfTypes.lean.expected.out +++ b/tests/lean/lcnfTypes.lean.expected.out @@ -14,7 +14,7 @@ Ty.denote : Ty → Type MonadControl.liftWith : {m : Type u → Type v} → {n : Type u → Type w} → [self : MonadControl m n] → {α : Type u} → (({β : Type u} → n β → m lcAny) → m α) → n α MonadControl.restoreM : {m : Type u → Type v} → {n : Type u → Type w} → [self : MonadControl m n] → {α : Type u} → m lcAny → n α -Decidable.casesOn : {p : Prop} → {motive : Decidable ◾ → Sort u} → Decidable ◾ → (◾ → motive lcAny) → (◾ → motive lcAny) → motive lcAny +Decidable.casesOn : {p : Prop} → {motive : Decidable ◾ → Sort u} → Decidable ◾ → (Bool → ◾ → motive lcAny) → motive lcAny Lean.getConstInfo : {m : Type → Type} → [Monad m] → [MonadEnv m] → [MonadError m] → Name → m ConstantInfo Lean.Meta.instMonadMetaM : Monad fun α => Context → ST.Ref PUnit State → Core.Context → ST.Ref PUnit Core.State → PUnit → EStateM.Result Exception PUnit α @@ -40,7 +40,7 @@ Member.head : lcErased → lcAny → List lcAny → Member lcAny lcAny lcAny Ty.denote : lcErased MonadControl.liftWith : lcErased → lcErased → MonadControl lcAny lcAny → lcErased → ((lcErased → lcAny → lcAny) → lcAny) → lcAny MonadControl.restoreM : lcErased → lcErased → MonadControl lcAny lcAny → lcErased → lcAny → lcAny -Decidable.casesOn : lcErased → lcErased → Bool → (lcErased → lcAny) → (lcErased → lcAny) → lcAny +Decidable.casesOn : lcErased → lcErased → Bool → (Bool → lcErased → lcAny) → lcAny Lean.getConstInfo : lcErased → Monad lcAny → MonadEnv lcAny → MonadError lcAny → Name → lcAny Lean.Meta.instMonadMetaM : Monad lcAny Lean.Meta.inferType : Expr → Context → lcAny → Core.Context → lcAny → PUnit → EStateM.Result Exception PUnit Expr diff --git a/tests/lean/listLength.lean.expected.out b/tests/lean/listLength.lean.expected.out index ae2f89a73ce5..941ecdd8bf19 100644 --- a/tests/lean/listLength.lean.expected.out +++ b/tests/lean/listLength.lean.expected.out @@ -1,8 +1,6 @@ - -[init] -def f (x_1 : obj) : obj := - let x_2 : obj := 0; - let x_3 : obj := List.lengthTRAux._rarg x_1 x_2; - let x_4 : obj := 2; - let x_5 : obj := Nat.mul x_4 x_3; - ret x_5 +[Compiler.IR] [init] + def f (x_1 : obj) : obj := + let x_2 : obj := List.lengthTR ◾ x_1; + let x_3 : obj := 1; + let x_4 : obj := Nat.shiftLeft x_2 x_3; + ret x_4 diff --git a/tests/lean/new-compiler/README.md b/tests/lean/new-compiler/README.md deleted file mode 100644 index a00ffda488e8..000000000000 --- a/tests/lean/new-compiler/README.md +++ /dev/null @@ -1,2 +0,0 @@ -This folder contains test files for the new compiler. -They are currently disabled while development on the compiler is paused and `compiler.enableNew` is deactivated. diff --git a/tests/lean/new-compiler/unhygienicCode.lean b/tests/lean/new-compiler/unhygienicCode.lean deleted file mode 100644 index e2a963fe1f1c..000000000000 --- a/tests/lean/new-compiler/unhygienicCode.lean +++ /dev/null @@ -1,10 +0,0 @@ -import Lean.Hygiene -open Lean - -set_option trace.Compiler.result true -set_option trace.compiler.ir.result true - --- The following function should not allocate any closures, --- nor any heap object that doesn't appear in the result: -def foo (n : Nat) : Syntax.Term := - Unhygienic.run `(a + $(quote n)) diff --git a/tests/lean/new-compiler/unhygienicCode.lean.expected.out b/tests/lean/new-compiler/unhygienicCode.lean.expected.out deleted file mode 100644 index 28ddf181d7ce..000000000000 --- a/tests/lean/new-compiler/unhygienicCode.lean.expected.out +++ /dev/null @@ -1,86 +0,0 @@ - -[result] -def foo._closed_1 : obj := - let x_1 : obj := ctor_0[Lean.Syntax.missing]; - let x_2 : u8 := 0; - let x_3 : obj := Lean.SourceInfo.fromRef x_1 x_2; - ret x_3 -def foo._closed_2 : obj := - let x_1 : obj := "UnhygienicMain"; - ret x_1 -def foo._closed_3 : obj := - let x_1 : obj := ctor_0[Lean.Name.anonymous._impl]; - let x_2 : obj := foo._closed_2; - let x_3 : obj := Lean.Name.str._override x_1 x_2; - ret x_3 -def foo._closed_4 : obj := - let x_1 : obj := "term_+_"; - ret x_1 -def foo._closed_5 : obj := - let x_1 : obj := ctor_0[Lean.Name.anonymous._impl]; - let x_2 : obj := foo._closed_4; - let x_3 : obj := Lean.Name.str._override x_1 x_2; - ret x_3 -def foo._closed_6 : obj := - let x_1 : obj := "a"; - ret x_1 -def foo._closed_7 : obj := - let x_1 : obj := foo._closed_6; - let x_2 : obj := String.toSubstring' x_1; - ret x_2 -def foo._closed_8 : obj := - let x_1 : obj := ctor_0[Lean.Name.anonymous._impl]; - let x_2 : obj := foo._closed_6; - let x_3 : obj := Lean.Name.str._override x_1 x_2; - ret x_3 -def foo._closed_9 : obj := - let x_1 : obj := foo._closed_3; - let x_2 : obj := foo._closed_8; - let x_3 : obj := Lean.firstFrontendMacroScope; - let x_4 : obj := Lean.addMacroScope x_1 x_2 x_3; - ret x_4 -def foo._closed_10 : obj := - let x_1 : obj := ctor_0[List.nil]; - let x_2 : obj := foo._closed_1; - let x_3 : obj := foo._closed_7; - let x_4 : obj := foo._closed_9; - let x_5 : obj := ctor_3[Lean.Syntax.ident] x_2 x_3 x_4 x_1; - ret x_5 -def foo._closed_11 : obj := - let x_1 : obj := "+"; - ret x_1 -def foo._closed_12 : obj := - let x_1 : obj := foo._closed_1; - let x_2 : obj := foo._closed_11; - let x_3 : obj := ctor_2[Lean.Syntax.atom] x_1 x_2; - ret x_3 -def foo (x_1 : obj) : obj := - let x_2 : obj := Nat.repr x_1; - let x_3 : obj := ctor_2[Lean.SourceInfo.none]; - let x_4 : obj := Lean.Syntax.mkNumLit x_2 x_3; - let x_5 : obj := foo._closed_1; - let x_6 : obj := foo._closed_5; - let x_7 : obj := foo._closed_10; - let x_8 : obj := foo._closed_12; - let x_9 : obj := Lean.Syntax.node3 x_5 x_6 x_7 x_8 x_4; - ret x_9[Compiler.result] size: 18 - def foo n : Syntax := - let _x.1 := Syntax.missing; - let _x.2 := 1; - let _x.3 := false; - let _x.4 := SourceInfo.fromRef _x.1 _x.3; - let _x.5 := "UnhygienicMain"; - let _x.6 := Name.mkStr1 _x.5; - let _x.7 := "term_+_"; - let _x.8 := Name.mkStr1 _x.7; - let _x.9 := "a"; - let _x.10 := String.toSubstring' _x.9; - let _x.11 := Name.mkStr1 _x.9; - let _x.12 := addMacroScope _x.6 _x.11 _x.2; - let _x.13 := [] _; - let _x.14 := Syntax.ident _x.4 _x.10 _x.12 _x.13; - let _x.15 := "+"; - let _x.16 := Syntax.atom _x.4 _x.15; - let _x.17 := Lean.instQuoteNatNumLitKind._elam_0 n; - let _x.18 := Syntax.node3 _x.4 _x.8 _x.14 _x.16 _x.17; - return _x.18 diff --git a/tests/lean/noncompSection.lean.expected.out b/tests/lean/noncompSection.lean.expected.out index 9bead8edfa7c..b278a5d979fe 100644 --- a/tests/lean/noncompSection.lean.expected.out +++ b/tests/lean/noncompSection.lean.expected.out @@ -1,2 +1,2 @@ 1 -noncompSection.lean:35:4-35:5: error: failed to compile definition, consider marking it as 'noncomputable' because it depends on 'Classical.choose', which is 'noncomputable' +noncompSection.lean:35:4-35:5: error: failed to compile definition, compiler IR check failed at 'i'. Error: depends on declaration 'Classical.choose', which has no executable code; consider marking definition as 'noncomputable' diff --git a/tests/lean/new-compiler/reduceArity.lean b/tests/lean/reduceArity.lean similarity index 100% rename from tests/lean/new-compiler/reduceArity.lean rename to tests/lean/reduceArity.lean diff --git a/tests/lean/new-compiler/reduceArity.lean.expected.out b/tests/lean/reduceArity.lean.expected.out similarity index 52% rename from tests/lean/new-compiler/reduceArity.lean.expected.out rename to tests/lean/reduceArity.lean.expected.out index 7fce8d6bb8f7..cdb7cfb150a4 100644 --- a/tests/lean/new-compiler/reduceArity.lean.expected.out +++ b/tests/lean/reduceArity.lean.expected.out @@ -1,14 +1,18 @@ -[Compiler.result] size: 5 - def g._redArg (n : Nat) (a : ◾) (f : ◾ → ◾) : ◾ := - cases n : ◾ - | Nat.zero => +[Compiler.result] size: 9 + def g._redArg (n : Nat) (a : lcAny) (f : lcAny → lcAny) : lcAny := + let zero := 0; + let isZero := Nat.decEq n zero; + cases isZero : lcAny + | Bool.true => return a - | Nat.succ (n.1 : Nat) => + | Bool.false => + let one := 1; + let n.1 := Nat.sub n one; let _x.2 := g._redArg n.1 a f; let _x.3 := f _x.2; return _x.3 [Compiler.result] size: 1 - def g (α : ◾) (n : Nat) (a : ◾) (b : ◾) (f : ◾ → ◾) : ◾ := + def g (α : ◾) (n : Nat) (a : lcAny) (b : lcAny) (f : lcAny → lcAny) : lcAny := let _x.1 := g._redArg n a f; return _x.1 [Compiler.result] size: 4 diff --git a/tests/lean/run/1785.lean b/tests/lean/run/1785.lean index c7a3a0dbf95f..ce198da7872b 100644 --- a/tests/lean/run/1785.lean +++ b/tests/lean/run/1785.lean @@ -11,8 +11,8 @@ noncomputable def char (R : Type) [∀ n, OfNat R n] : Nat := 0 /-- -error: failed to compile definition, consider marking it as 'noncomputable' -because it depends on 'char', which is 'noncomputable' +error: failed to compile definition, compiler IR check failed at 'bug._redArg'. +Error: depends on declaration 'char', which has no executable code; consider marking definition as 'noncomputable' -/ #guard_msgs in def bug (R : Type) [∀ n, OfNat R n] : R := diff --git a/tests/lean/run/2161.lean b/tests/lean/run/2161.lean index 25659ff1dd69..b2d59e109cbf 100644 --- a/tests/lean/run/2161.lean +++ b/tests/lean/run/2161.lean @@ -1,3 +1,5 @@ +#exit -- disabled due to `decide` changes + structure Foo where num : Nat deriving DecidableEq namespace Foo diff --git a/tests/lean/run/2291.lean b/tests/lean/run/2291.lean index 1f6827c1d58a..8097c5726034 100644 --- a/tests/lean/run/2291.lean +++ b/tests/lean/run/2291.lean @@ -6,107 +6,195 @@ Issue #2291 The following example would cause the pretty printer to panic. -/ -set_option trace.compiler.simp true in +set_option trace.Compiler.simp true in /-- info: [0] --- -info: [compiler.simp] >> _eval -let _x_21 := `Nat; -let _x_22 := []; -let _x_23 := Lean.Expr.const _x_21 _x_22; -let _x_24 := `List.nil; -let _x_25 := Lean.Level.zero :: _x_22; -let _x_26 := Lean.Expr.const _x_24 _x_25; -let _x_27 := _x_26.app _x_23; -let _x_28 := `List.cons; -let _x_29 := Lean.Expr.const _x_28 _x_25; -let _x_30 := _x_29.app _x_23; -let _x_31 := []; -let _x_32 := 0 :: _x_31; -let _x_33 := Lean.List.toExprAux✝ _x_27 _x_30 _x_32; -Lean.MessageData.ofExpr _x_33 -[compiler.simp] >> _private.Lean.ToExpr.0.Lean.List.toExprAux._at._eval._spec_1 -fun nilFn consFn x => - List.casesOn fun head tail => - let _x_1 := Lean.mkNatLit head; - let _x_2 := Lean.List.toExprAux._at._eval._spec_1✝ nilFn consFn tail; - Lean.mkAppB consFn _x_1 _x_2 ->> _eval -let _x_14 := Lean.Name.str._override Lean.Name.anonymous._impl "Nat"; -let _x_15 := List.nil _neutral; -let _x_16 := Lean.Expr.const._override _x_14 _x_15; -let _x_17 := `List.nil; -let _x_18 := List.cons _neutral Lean.Level.zero._impl _x_15; -let _x_19 := Lean.Expr.const._override _x_17 _x_18; -let _x_20 := Lean.Expr.app._override _x_19 _x_16; -let _x_21 := `List.cons; -let _x_22 := Lean.Expr.const._override _x_21 _x_18; -let _x_23 := Lean.Expr.app._override _x_22 _x_16; -let _x_24 := List.cons _neutral 0 _x_15; -let _x_25 := Lean.List.toExprAux._at._eval._spec_1✝ _x_20 _x_23 _x_24; -Lean.MessageData.ofExpr _x_25 -[compiler.simp] >> _private.Lean.ToExpr.0.Lean.List.toExprAux._at._eval._spec_1 -fun nilFn consFn x => - List.casesOn fun head tail => - let _x_1 := Lean.mkNatLit head; - let _x_2 := Lean.List.toExprAux._at._eval._spec_1✝ nilFn consFn tail; - Lean.mkAppB consFn _x_1 _x_2 ->> _eval -let _x_1 := Lean.Name.str._override Lean.Name.anonymous._impl "Nat"; -let _x_2 := List.nil _neutral; -let _x_3 := Lean.Expr.const._override _x_1 _x_2; -let _x_4 := `List.nil; -let _x_5 := List.cons _neutral Lean.Level.zero._impl _x_2; -let _x_6 := Lean.Expr.const._override _x_4 _x_5; -let _x_7 := Lean.Expr.app._override _x_6 _x_3; -let _x_8 := `List.cons; -let _x_9 := Lean.Expr.const._override _x_8 _x_5; -let _x_10 := Lean.Expr.app._override _x_9 _x_3; -let _x_11 := List.cons _neutral 0 _x_2; -let _x_12 := Lean.List.toExprAux._at._eval._spec_1✝ _x_7 _x_10 _x_11; -Lean.MessageData.ofExpr _x_12 -[compiler.simp] >> _private.Lean.ToExpr.0.Lean.List.toExprAux._at._eval._spec_1 -fun nilFn consFn x => - List.casesOn fun head tail => - let _x_1 := Lean.mkNatLit head; - let _x_2 := Lean.List.toExprAux._at._eval._spec_1✝ nilFn consFn tail; - Lean.mkAppB consFn _x_1 _x_2 ->> _eval._closed_1 -"Nat" ->> _eval._closed_2 -Lean.Name.str._override Lean.Name.anonymous._impl _eval._closed_1 ->> _eval._closed_3 -let _x_1 := List.nil _neutral; -Lean.Expr.const._override _eval._closed_2 _x_1 ->> _eval._closed_4 -"List" ->> _eval._closed_5 -"nil" ->> _eval._closed_6 -Lean.Name.mkStr2 _eval._closed_4 _eval._closed_5 ->> _eval._closed_7 -let _x_1 := List.nil _neutral; -List.cons _neutral Lean.Level.zero._impl _x_1 ->> _eval._closed_8 -Lean.Expr.const._override _eval._closed_6 _eval._closed_7 ->> _eval._closed_9 -Lean.Expr.app._override _eval._closed_8 _eval._closed_3 ->> _eval._closed_10 -"cons" ->> _eval._closed_11 -Lean.Name.mkStr2 _eval._closed_4 _eval._closed_10 ->> _eval._closed_12 -Lean.Expr.const._override _eval._closed_11 _eval._closed_7 ->> _eval._closed_13 -Lean.Expr.app._override _eval._closed_12 _eval._closed_3 ->> _eval._closed_14 -let _x_1 := List.nil _neutral; -List.cons _neutral 0 _x_1 ->> _eval -let _x_1 := - Lean.List.toExprAux._at._eval._spec_1✝ _eval._closed_9 _eval._closed_13 - _eval._closed_14; -Lean.MessageData.ofExpr _x_1 +trace: [Compiler.simp] size: 22 + def _eval : Lean.MessageData := + let _x.1 := Lean.instToExprNat; + let _x.2 := "Nat"; + let _x.3 := Lean.Name.mkStr1 _x.2; + let _x.4 := @List.nil _; + let type := Lean.Expr.const._override _x.3 _x.4; + let _x.5 := "List"; + let _x.6 := "nil"; + let _x.7 := Lean.Name.mkStr2 _x.5 _x.6; + let _x.8 := Lean.Level.zero; + let _x.9 := @List.nil _; + let _x.10 := @List.cons _ _x.8 _x.9; + let _x.11 := Lean.Expr.const._override _x.7 _x.10; + let nil := Lean.Expr.app._override _x.11 type; + let _x.12 := "cons"; + let _x.13 := Lean.Name.mkStr2 _x.5 _x.12; + let _x.14 := Lean.Expr.const._override _x.13 _x.10; + let cons := Lean.Expr.app._override _x.14 type; + let _x.15 := 0; + let _x.16 := @List.nil _; + let _x.17 := @List.cons _ _x.15 _x.16; + let _x.18 := @Lean.List.toExprAux.0 _ _x.1 nil cons _x.17; + let _x.19 := Lean.MessageData.ofExpr _x.18; + return _x.19 +[Compiler.simp] size: 23 + def _eval : Lean.MessageData := + let _x.1 := Lean.instToExprNat; + let _x.2 := "Nat"; + let _x.3 := Lean.Name.mkStr1 _x.2; + let _x.4 := @List.nil _; + let type := Lean.Expr.const._override _x.3 _x.4; + let _x.5 := "List"; + let _x.6 := "nil"; + let _x.7 := Lean.Name.mkStr2 _x.5 _x.6; + let _x.8 := Lean.Level.zero._impl; + let _x.9 := @unsafeCast _ _ _x.8; + let _x.10 := @List.nil _; + let _x.11 := @List.cons _ _x.9 _x.10; + let _x.12 := Lean.Expr.const._override _x.7 _x.11; + let nil := Lean.Expr.app._override _x.12 type; + let _x.13 := "cons"; + let _x.14 := Lean.Name.mkStr2 _x.5 _x.13; + let _x.15 := Lean.Expr.const._override _x.14 _x.11; + let cons := Lean.Expr.app._override _x.15 type; + let _x.16 := 0; + let _x.17 := @List.nil _; + let _x.18 := @List.cons _ _x.16 _x.17; + let _x.19 := @Lean.List.toExprAux.0 _ _x.1 nil cons _x.18; + let _x.20 := Lean.MessageData.ofExpr _x.19; + return _x.20 +[Compiler.simp] size: 6 + def _private.Lean.ToExpr.0.Lean.List.toExprAux._at_._eval.spec_0 nilFn consFn x.1 : Lean.Expr := + cases x.1 : Lean.Expr + | List.nil => + return nilFn + | List.cons head.2 tail.3 => + let _x.4 := Lean.mkNatLit head.2; + let _x.5 := Lean.List.toExprAux._at_._eval.spec_0.0 nilFn consFn tail.3; + let _x.6 := Lean.mkAppB consFn _x.4 _x.5; + return _x.6 +[Compiler.simp] size: 22 + def _eval : Lean.MessageData := + let _x.1 := "Nat"; + let _x.2 := Lean.Name.mkStr1 _x.1; + let _x.3 := @List.nil _; + let type := Lean.Expr.const._override _x.2 _x.3; + let _x.4 := "List"; + let _x.5 := "nil"; + let _x.6 := Lean.Name.mkStr2 _x.4 _x.5; + let _x.7 := Lean.Level.zero._impl; + let _x.8 := @unsafeCast _ _ _x.7; + let _x.9 := @List.nil _; + let _x.10 := @List.cons _ _x.8 _x.9; + let _x.11 := Lean.Expr.const._override _x.6 _x.10; + let nil := Lean.Expr.app._override _x.11 type; + let _x.12 := "cons"; + let _x.13 := Lean.Name.mkStr2 _x.4 _x.12; + let _x.14 := Lean.Expr.const._override _x.13 _x.10; + let cons := Lean.Expr.app._override _x.14 type; + let _x.15 := 0; + let _x.16 := @List.nil _; + let _x.17 := @List.cons _ _x.15 _x.16; + let _x.18 := Lean.List.toExprAux._at_._eval.spec_0.0 nil cons _x.17; + let _x.19 := Lean.MessageData.ofExpr _x.18; + return _x.19 +[Compiler.simp] size: 6 + def _private.Lean.ToExpr.0.Lean.List.toExprAux._at_._eval.spec_0 nilFn consFn x.1 : Lean.Expr := + cases x.1 : Lean.Expr + | List.nil => + return nilFn + | List.cons head.2 tail.3 => + let _x.4 := Lean.mkNatLit head.2; + let _x.5 := Lean.List.toExprAux._at_._eval.spec_0.0 nilFn consFn tail.3; + let _x.6 := Lean.mkAppB consFn _x.4 _x.5; + return _x.6 +[Compiler.simp] size: 20 + def _eval : Lean.MessageData := + let _x.1 := "Nat"; + let _x.2 := Lean.Name.mkStr1 _x.1; + let _x.3 := [] _; + let type := Lean.Expr.const._override _x.2 _x.3; + let _x.4 := "List"; + let _x.5 := "nil"; + let _x.6 := Lean.Name.mkStr2 _x.4 _x.5; + let _x.7 := Lean.Level.zero._impl; + let _x.8 := List.cons _ _x.7 _x.3; + let _x.9 := Lean.Expr.const._override _x.6 _x.8; + let nil := Lean.Expr.app._override _x.9 type; + let _x.10 := "cons"; + let _x.11 := Lean.Name.mkStr2 _x.4 _x.10; + let _x.12 := Lean.Expr.const._override _x.11 _x.8; + let cons := Lean.Expr.app._override _x.12 type; + let _x.13 := 0; + let _x.14 := [] _; + let _x.15 := List.cons _ _x.13 _x.14; + let _x.16 := Lean.List.toExprAux._at_._eval.spec_0.0 nil cons _x.15; + let _x.17 := Lean.MessageData.ofExpr _x.16; + return _x.17 +[Compiler.simp] size: 6 + def _private.Lean.ToExpr.0.Lean.List.toExprAux._at_._eval.spec_0 nilFn consFn x.1 : Lean.Expr := + cases x.1 : Lean.Expr + | List.nil => + return nilFn + | List.cons head.2 tail.3 => + let _x.4 := Lean.mkNatLit head.2; + let _x.5 := Lean.List.toExprAux._at_._eval.spec_0.0 nilFn consFn tail.3; + let _x.6 := Lean.mkAppB consFn _x.4 _x.5; + return _x.6 +[Compiler.simp] size: 20 + def _eval : Lean.MessageData := + let _x.1 := "Nat"; + let _x.2 := Lean.Name.mkStr1 _x.1; + let _x.3 := [] _; + let type := Lean.Expr.const._override _x.2 _x.3; + let _x.4 := "List"; + let _x.5 := "nil"; + let _x.6 := Lean.Name.mkStr2 _x.4 _x.5; + let _x.7 := Lean.Level.zero._impl; + let _x.8 := List.cons _ _x.7 _x.3; + let _x.9 := Lean.Expr.const._override _x.6 _x.8; + let nil := Lean.Expr.app._override _x.9 type; + let _x.10 := "cons"; + let _x.11 := Lean.Name.mkStr2 _x.4 _x.10; + let _x.12 := Lean.Expr.const._override _x.11 _x.8; + let cons := Lean.Expr.app._override _x.12 type; + let _x.13 := 0; + let _x.14 := [] _; + let _x.15 := List.cons _ _x.13 _x.14; + let _x.16 := Lean.List.toExprAux._at_._eval.spec_0.0 nil cons _x.15; + let _x.17 := Lean.MessageData.ofExpr _x.16; + return _x.17 +[Compiler.simp] size: 6 + def _private.Lean.ToExpr.0.Lean.List.toExprAux._at_._eval.spec_0 nilFn consFn x.1 : Lean.Expr := + cases x.1 : Lean.Expr + | List.nil => + return nilFn + | List.cons head.2 tail.3 => + let _x.4 := Lean.mkNatLit head.2; + let _x.5 := Lean.List.toExprAux._at_._eval.spec_0.0 nilFn consFn tail.3; + let _x.6 := Lean.mkAppB consFn _x.4 _x.5; + return _x.6 +[Compiler.simp] size: 20 + def _eval : Lean.MessageData := + let _x.1 := "Nat"; + let _x.2 := Lean.Name.mkStr1 _x.1; + let _x.3 := [] _; + let type := Lean.Expr.const._override _x.2 _x.3; + let _x.4 := "List"; + let _x.5 := "nil"; + let _x.6 := Lean.Name.mkStr2 _x.4 _x.5; + let _x.7 := Lean.Level.zero._impl; + let _x.8 := List.cons _ _x.7 _x.3; + let _x.9 := Lean.Expr.const._override _x.6 _x.8; + let nil := Lean.Expr.app._override _x.9 type; + let _x.10 := "cons"; + let _x.11 := Lean.Name.mkStr2 _x.4 _x.10; + let _x.12 := Lean.Expr.const._override _x.11 _x.8; + let cons := Lean.Expr.app._override _x.12 type; + let _x.13 := 0; + let _x.14 := [] _; + let _x.15 := List.cons _ _x.13 _x.14; + let _x.16 := Lean.List.toExprAux._at_._eval.spec_0.0 nil cons _x.15; + let _x.17 := Lean.MessageData.ofExpr _x.16; + return _x.17 -/ #guard_msgs in #eval [0] diff --git a/tests/lean/run/2552.lean b/tests/lean/run/2552.lean index bfb3f05fd31e..a8c2571dd7ff 100644 --- a/tests/lean/run/2552.lean +++ b/tests/lean/run/2552.lean @@ -19,7 +19,7 @@ local instance decidableBallLT : | isFalse p => isFalse (p <| · _ _) | isTrue p => isTrue fun _ h' => (Nat.le_of_lt_succ h').lt_or_eq_dec.elim (h _) (· ▸ p) -set_option maxHeartbeats 5000 +set_option maxHeartbeats 10000 example : ∀ a, a < 9 → ∀ b, b < 9 → ∀ c, c < 9 → a ^ 2 + b ^ 2 + c ^ 2 ≠ 7 := by decide end succeeds_using_match @@ -38,7 +38,7 @@ local instance decidableBallLT' : | isFalse p => isFalse (p <| · _ _) | isTrue p => isTrue fun _ h' => (Nat.le_of_lt_succ h').lt_or_eq_dec.elim (h _) (· ▸ p) -set_option maxHeartbeats 5000 +set_option maxHeartbeats 10000 example : ∀ a, a < 9 → ∀ b, b < 9 → ∀ c, c < 9 → a ^ 2 + b ^ 2 + c ^ 2 ≠ 7 := by decide end fails_with_timeout diff --git a/tests/lean/run/3807.lean b/tests/lean/run/3807.lean index 24415fab7517..18f6ac0308e7 100644 --- a/tests/lean/run/3807.lean +++ b/tests/lean/run/3807.lean @@ -2491,7 +2491,7 @@ variable {L : Type _} [Field L] [Algebra F L] [Algebra L E] [IsScalarTower F L E (f : L →ₐ[F] K) -- This only required 16,000 heartbeats prior to #3807, and now takes ~210,000. -set_option maxHeartbeats 16000 +set_option maxHeartbeats 32000 theorem exists_algHom_adjoin_of_splits''' : ∃ φ : adjoin L S →ₐ[F] K, φ.comp (IsScalarTower.toAlgHom F L _) = f := by let L' := (IsScalarTower.toAlgHom F L E).fieldRange diff --git a/tests/lean/run/4413.lean b/tests/lean/run/4413.lean index 8690181126db..3ee84e85b1b4 100644 --- a/tests/lean/run/4413.lean +++ b/tests/lean/run/4413.lean @@ -24,7 +24,7 @@ but is expected to have type -/ #guard_msgs in set_option maxRecDepth 100 in -set_option maxHeartbeats 200 in +set_option maxHeartbeats 400 in example (n : UInt64) : n = n - 1 := rfl diff --git a/tests/lean/run/4465.lean b/tests/lean/run/4465.lean index 081b59964237..e44e5fd470cc 100644 --- a/tests/lean/run/4465.lean +++ b/tests/lean/run/4465.lean @@ -3,8 +3,8 @@ #reduce Char.ofNat (nat_lit 0) /-- -info: { val := { toBitVec := { toFin := ⟨0, isValidChar_UInt32✝ (Or.inl (Nat.le_of_ble_eq_true rfl))⟩ } }, - valid := Or.inl (Nat.le_of_ble_eq_true rfl) } +info: { val := { toBitVec := { toFin := ⟨0, isValidChar_UInt32✝ (Decidable.of_decide (Nat.isValidChar 0))⟩ } }, + valid := Decidable.of_decide (Nat.isValidChar 0) } -/ #guard_msgs in set_option pp.proofs true in diff --git a/tests/lean/run/4644.lean b/tests/lean/run/4644.lean index f2e3dfdcbb65..067393d9ec20 100644 --- a/tests/lean/run/4644.lean +++ b/tests/lean/run/4644.lean @@ -23,6 +23,8 @@ is not definitionally equal to the right-hand side example: check_sorted #[0, 3, 3, 5, 8, 10, 10, 10] = true := by rfl -- fails because `rfl` uses `.default` transparency, and `sorted_from_var` is marked as irreducible +-- disable because of `decide` changes +/- /-- error: tactic 'decide' failed for proposition check_sorted #[0, 3, 3, 5, 8, 10, 10, 10] = true @@ -40,6 +42,7 @@ After unfolding the instances 'instDecidableEqBool' and 'Bool.decEq', reduction #guard_msgs in example: check_sorted #[0, 3, 3, 5, 8, 10, 10, 10] := by decide -- fails because `decide` uses `.default` transparency, and `sorted_from_var` is marked as irreducible +-/ unseal sorted_from_var in example: check_sorted #[0, 3, 3, 5, 8, 10, 10, 10] := by diff --git a/tests/lean/run/4861.lean b/tests/lean/run/4861.lean index 9799a6cdc029..609f6deed40a 100644 --- a/tests/lean/run/4861.lean +++ b/tests/lean/run/4861.lean @@ -1,4 +1,4 @@ -set_option maxHeartbeats 210000 in +set_option maxHeartbeats 420000 in theorem foo (x y z p q : Int) : False := have : (1 * x ^ 1 + z ^ 1 * p) * (1 / 1 * p ^ 1 * x ^ 1 + 1 * q * p ^ 1 * x * z + 1 * q ^ 1 * p ^ 1 * x ^ 1 + diff --git a/tests/lean/run/CompilerCSE.lean b/tests/lean/run/CompilerCSE.lean deleted file mode 100644 index dff84f3c81fe..000000000000 --- a/tests/lean/run/CompilerCSE.lean +++ /dev/null @@ -1,32 +0,0 @@ -import Lean.Compiler.Main -import Lean.Compiler.LCNF.Testing -import Lean.Elab.Do -import Lean.Elab.Command - -open Lean -open Lean.Compiler.LCNF - -@[cpass] -def cseFixTest : PassInstaller := Testing.assertIsAtFixPoint |>.install `cse `cseFix - -@[cpass] -def cseSizeTest : PassInstaller := - Testing.assertReducesOrPreservesSize "CSE increased size of declaration" |>.install `cse `cseSizeLeq - -set_option trace.Compiler.test true in -/-- -trace: [Compiler.test] Starting wrapper test cseSizeLeq for cse occurrence 0 -[Compiler.test] Wrapper test cseSizeLeq for cse occurrence 0 successful -[Compiler.test] Starting post condition test cseFix for cse occurrence 0 -[Compiler.test] Post condition test cseFix for cse occurrence 0 successful -[Compiler.test] Starting wrapper test cseSizeLeq for cse occurrence 1 -[Compiler.test] Wrapper test cseSizeLeq for cse occurrence 1 successful -[Compiler.test] Starting post condition test cseFix for cse occurrence 1 -[Compiler.test] Post condition test cseFix for cse occurrence 1 successful -[Compiler.test] Starting wrapper test cseSizeLeq for cse occurrence 2 -[Compiler.test] Wrapper test cseSizeLeq for cse occurrence 2 successful -[Compiler.test] Starting post condition test cseFix for cse occurrence 2 -[Compiler.test] Post condition test cseFix for cse occurrence 2 successful --/ -#guard_msgs in -run_meta Compiler.compile #[``Lean.Meta.synthInstance, ``Lean.Elab.Term.Do.elabDo] diff --git a/tests/lean/run/CompilerFindJoinPoints.lean b/tests/lean/run/CompilerFindJoinPoints.lean deleted file mode 100644 index 2b60804f60b1..000000000000 --- a/tests/lean/run/CompilerFindJoinPoints.lean +++ /dev/null @@ -1,24 +0,0 @@ -import Lean.Compiler.Main -import Lean.Compiler.LCNF.Testing -import Lean.Elab.Do -import Lean.Elab.Command - -open Lean -open Lean.Compiler.LCNF - -@[cpass] -def findJoinPointFixTest : PassInstaller := Testing.assertIsAtFixPoint |>.install `findJoinPoints `findJoinPointsFix - -@[cpass] -def cseSizeTest : PassInstaller := - Testing.assertReducesOrPreservesSize "findJoinPoints increased size of declaration" |>.install `findJoinPoints `findJoinPointsSizeLeq - -set_option trace.Compiler.test true in -/-- -trace: [Compiler.test] Starting wrapper test findJoinPointsSizeLeq for findJoinPoints occurrence 0 -[Compiler.test] Wrapper test findJoinPointsSizeLeq for findJoinPoints occurrence 0 successful -[Compiler.test] Starting post condition test findJoinPointsFix for findJoinPoints occurrence 0 -[Compiler.test] Post condition test findJoinPointsFix for findJoinPoints occurrence 0 successful --/ -#guard_msgs in -run_meta Compiler.compile #[``Lean.Meta.synthInstance, ``Lean.Elab.Term.Do.elabDo, ``Lean.MetavarContext.MkBinding.collectForwardDeps] diff --git a/tests/lean/run/CompilerFloatLetIn.lean b/tests/lean/run/CompilerFloatLetIn.lean deleted file mode 100644 index 86f6f2022047..000000000000 --- a/tests/lean/run/CompilerFloatLetIn.lean +++ /dev/null @@ -1,27 +0,0 @@ -import Lean.Compiler.Main -import Lean.Compiler.LCNF.Testing -import Lean.Elab.Do -import Lean.Elab.Command - -open Lean -open Lean.Compiler.LCNF - --- #eval fails if we uncomment this pass after I added a `floatLetIn` pass at the mono phase --- @[cpass] --- def floatLetInFixTest : PassInstaller := Testing.assertIsAtFixPoint |>.install `floatLetIn `floatLetInFix - -@[cpass] -def floatLetInSizeTest : PassInstaller := - Testing.assertReducesOrPreservesSize "FloatLetIn increased size of declaration" |>.install `floatLetIn `floatLetInSizeEq - -set_option trace.Compiler.test true in -/-- -trace: [Compiler.test] Starting wrapper test floatLetInSizeEq for floatLetIn occurrence 0 -[Compiler.test] Wrapper test floatLetInSizeEq for floatLetIn occurrence 0 successful -[Compiler.test] Starting wrapper test floatLetInSizeEq for floatLetIn occurrence 1 -[Compiler.test] Wrapper test floatLetInSizeEq for floatLetIn occurrence 1 successful -[Compiler.test] Starting wrapper test floatLetInSizeEq for floatLetIn occurrence 2 -[Compiler.test] Wrapper test floatLetInSizeEq for floatLetIn occurrence 2 successful --/ -#guard_msgs in -run_meta Compiler.compile #[``Lean.Meta.synthInstance, ``Lean.Elab.Term.Do.elabDo] diff --git a/tests/lean/run/CompilerProbe.lean b/tests/lean/run/CompilerProbe.lean deleted file mode 100644 index 5fc496b68832..000000000000 --- a/tests/lean/run/CompilerProbe.lean +++ /dev/null @@ -1,49 +0,0 @@ -import Lean -import Lean.Compiler.LCNF.Probing -open Lean.Compiler.LCNF - --- Note: 2024-05-15: At the time of adding #guard_msgs here, the tests seem to all be failing. - --- Find functions that have jps which take a lambda -/-- info: #[] -/ -#guard_msgs in -#eval - Probe.runGlobally (phase := .mono) <| - Probe.filterByJp (·.params.anyM (fun param => return param.type.isForall)) >=> - Probe.declNames >=> - Probe.toString - --- Count lambda lifted functions -def lambdaCounter : Probe Decl Nat := - Probe.filter (fun decl => - if let .str _ val := decl.name then - return val.startsWith "_lam" - else - return false) >=> - Probe.declNames >=> - Probe.count - --- Run everywhere -/-- info: #[0] -/ -#guard_msgs in -#eval - Probe.runGlobally (phase := .mono) <| - lambdaCounter - --- Run limited -/-- info: #[0] -/ -#guard_msgs in -#eval - Probe.runOnModule `Lean.Compiler.LCNF.JoinPoints (phase := .mono) <| - lambdaCounter - --- Find most commonly used function with threshold -/-- info: #[] -/ -#guard_msgs in -#eval - Probe.runOnModule `Lean.Compiler.LCNF.JoinPoints (phase := .mono) <| - Probe.getLetValues >=> - Probe.filter (fun e => return e matches .const ..) >=> - Probe.map (fun | .const declName .. => return s!"{declName}" | _ => unreachable!) >=> - Probe.countUniqueSorted >=> - Probe.filter (fun (_, count) => return count > 100) diff --git a/tests/lean/run/CompilerPullInstances.lean b/tests/lean/run/CompilerPullInstances.lean deleted file mode 100644 index 0582ce399142..000000000000 --- a/tests/lean/run/CompilerPullInstances.lean +++ /dev/null @@ -1,24 +0,0 @@ -import Lean.Compiler.Main -import Lean.Compiler.LCNF.Testing -import Lean.Elab.Do -import Lean.Elab.Command - -open Lean -open Lean.Compiler.LCNF - -@[cpass] -def pullInstancesFixTest : PassInstaller := Testing.assertIsAtFixPoint |>.install `pullInstances `pullInstancesFix - -@[cpass] -def pullInstancesSizeTest : PassInstaller := - Testing.assertPreservesSize "Pulling instances changed size" |>.install `pullInstances `pullInstancesSizeEq - -set_option trace.Compiler.test true in -/-- -trace: [Compiler.test] Starting wrapper test pullInstancesSizeEq for pullInstances occurrence 0 -[Compiler.test] Wrapper test pullInstancesSizeEq for pullInstances occurrence 0 successful -[Compiler.test] Starting post condition test pullInstancesFix for pullInstances occurrence 0 -[Compiler.test] Post condition test pullInstancesFix for pullInstances occurrence 0 successful --/ -#guard_msgs in -run_meta Compiler.compile #[``Lean.Meta.synthInstance, ``Lean.Elab.Term.Do.elabDo] diff --git a/tests/lean/run/CompilerSimp.lean b/tests/lean/run/CompilerSimp.lean deleted file mode 100644 index bea661e18433..000000000000 --- a/tests/lean/run/CompilerSimp.lean +++ /dev/null @@ -1,44 +0,0 @@ -import Lean.Compiler.Main -import Lean.Compiler.LCNF.Testing -import Lean.Elab.Do -import Lean.Elab.Command - -open Lean -open Lean.Compiler.LCNF - -@[cpass] -def simpFixTest : PassInstaller := Testing.assertIsAtFixPoint |>.install `simp `simpFix - -@[cpass] -def simpReaderTest : PassInstaller := - Testing.assertDoesNotContainConstAfter `ReaderT.bind "simp did not inline ReaderT.bind" |>.install `simp `simpInlinesBinds - -set_option trace.Compiler.test true in -/-- -trace: [Compiler.test] Starting post condition test simpInlinesBinds for simp occurrence 0 -[Compiler.test] Post condition test simpInlinesBinds for simp occurrence 0 successful -[Compiler.test] Starting post condition test simpFix for simp occurrence 0 -[Compiler.test] Post condition test simpFix for simp occurrence 0 successful -[Compiler.test] Starting post condition test simpInlinesBinds for simp occurrence 1 -[Compiler.test] Post condition test simpInlinesBinds for simp occurrence 1 successful -[Compiler.test] Starting post condition test simpFix for simp occurrence 1 -[Compiler.test] Post condition test simpFix for simp occurrence 1 successful -[Compiler.test] Starting post condition test simpInlinesBinds for simp occurrence 2 -[Compiler.test] Post condition test simpInlinesBinds for simp occurrence 2 successful -[Compiler.test] Starting post condition test simpFix for simp occurrence 2 -[Compiler.test] Post condition test simpFix for simp occurrence 2 successful -[Compiler.test] Starting post condition test simpInlinesBinds for simp occurrence 3 -[Compiler.test] Post condition test simpInlinesBinds for simp occurrence 3 successful -[Compiler.test] Starting post condition test simpFix for simp occurrence 3 -[Compiler.test] Post condition test simpFix for simp occurrence 3 successful -[Compiler.test] Starting post condition test simpInlinesBinds for simp occurrence 4 -[Compiler.test] Post condition test simpInlinesBinds for simp occurrence 4 successful -[Compiler.test] Starting post condition test simpFix for simp occurrence 4 -[Compiler.test] Post condition test simpFix for simp occurrence 4 successful -[Compiler.test] Starting post condition test simpInlinesBinds for simp occurrence 5 -[Compiler.test] Post condition test simpInlinesBinds for simp occurrence 5 successful -[Compiler.test] Starting post condition test simpFix for simp occurrence 5 -[Compiler.test] Post condition test simpFix for simp occurrence 5 successful --/ -#guard_msgs in -run_meta Compiler.compile #[``Lean.Meta.synthInstance, ``Lean.Elab.Term.Do.elabDo] diff --git a/tests/lean/run/Decidable-decide-erasure.lean b/tests/lean/run/Decidable-decide-erasure.lean index 1b72c6fff954..ad1301e8db1c 100644 --- a/tests/lean/run/Decidable-decide-erasure.lean +++ b/tests/lean/run/Decidable-decide-erasure.lean @@ -6,9 +6,6 @@ open Lean.Compiler.LCNF def f (a : Nat) : Bool := decide (a = 1) --- This is only required until the new code generator is enabled. -run_meta Lean.Compiler.compile #[``f] - def countCalls : Probe Decl Nat := Probe.getLetValues >=> Probe.filter (fun e => return e matches .const `Decidable.decide ..) >=> @@ -16,7 +13,7 @@ def countCalls : Probe Decl Nat := #eval do let numCalls <- Probe.runOnDeclsNamed #[`f] (phase := .base) <| countCalls - assert! numCalls == #[1] + assert! numCalls == #[0] #eval do let numCalls <- Probe.runOnDeclsNamed #[`f] (phase := .mono) <| countCalls diff --git a/tests/lean/run/bv_bitblast_stress.lean b/tests/lean/run/bv_bitblast_stress.lean index 59147dbdb69b..d0c50bc932f2 100644 --- a/tests/lean/run/bv_bitblast_stress.lean +++ b/tests/lean/run/bv_bitblast_stress.lean @@ -4,6 +4,8 @@ open BitVec set_option exponentiation.threshold 4096 +set_option maxHeartbeats 4000000 + theorem t1 {x y : BitVec 64} (h : x = y) : (~~~x) &&& y = (~~~y) &&& x := by bv_decide diff --git a/tests/lean/run/congrThm.lean b/tests/lean/run/congrThm.lean index 24ec39c16358..bd3a78725f95 100644 --- a/tests/lean/run/congrThm.lean +++ b/tests/lean/run/congrThm.lean @@ -11,7 +11,7 @@ def test (f : Expr) : MetaM Unit := do IO.println (← Meta.ppExpr thm.type) /-- -info: ∀ (p p_1 : Prop), p = p_1 → ∀ {h : Decidable p} [h_1 : Decidable p_1], decide p = decide p_1 +info: ∀ (p p_1 : Prop), p = p_1 → ∀ {self : Decidable p} [self_1 : Decidable p_1], decide p = decide p_1 -/ #guard_msgs in #eval test (mkConst ``decide) diff --git a/tests/lean/run/decideTactic.lean b/tests/lean/run/decideTactic.lean index 210cba140bf8..516caacaee36 100644 --- a/tests/lean/run/decideTactic.lean +++ b/tests/lean/run/decideTactic.lean @@ -2,6 +2,8 @@ # Tests of the `decide` tactic -/ +-- some tests temporarily disabled due to changes to `decide` + /-! Success -/ @@ -27,6 +29,7 @@ Irreducible decidable instance -/ opaque unknownProp : Prop +/- /-- error: tactic 'decide' failed for proposition unknownProp @@ -46,6 +49,7 @@ command, which enables the instance 'Classical.propDecidable'. #guard_msgs in open scoped Classical in example : unknownProp := by decide +-/ /-! @@ -63,6 +67,7 @@ def baz (n : Nat) : Decidable (Nice n) := by instance : Decidable (Nice n) := baz n +/- /-- error: tactic 'decide' failed for proposition Nice 102 @@ -79,12 +84,14 @@ defined using tactics such as 'rw' or 'simp'. To avoid tactics, make use of func -/ #guard_msgs in example : Nice 102 := by decide +-/ /-! Following `Decidable.rec` to give better messages -/ +/- /-- error: tactic 'decide' failed for proposition ¬Nice 102 @@ -102,6 +109,7 @@ defined using tactics such as 'rw' or 'simp'. To avoid tactics, make use of func -/ #guard_msgs in example : ¬ Nice 102 := by decide +-/ /-! diff --git a/tests/lean/run/decideTacticKernel.lean b/tests/lean/run/decideTacticKernel.lean index 49bfd190a5b6..25c08a61285e 100644 --- a/tests/lean/run/decideTacticKernel.lean +++ b/tests/lean/run/decideTacticKernel.lean @@ -33,6 +33,7 @@ The kernel sees through irreducible definitions -/ @[irreducible] def irred {α : Type} (x : α) : α := x +/- /-- error: tactic 'decide' failed for proposition irred 3 = 3 @@ -46,6 +47,7 @@ After unfolding the instances 'instDecidableEqNat' and 'Nat.decEq', reduction go | false => isFalse ⋯ -/ #guard_msgs in theorem gcd_eq1 : irred 3 = 3 := by decide +-/ theorem gcd_eq2 : irred 3 = 3 := by decide +kernel diff --git a/tests/lean/run/emptyLcnf.lean b/tests/lean/run/emptyLcnf.lean index 2c5acd8dc13a..6bae0ca03abc 100644 --- a/tests/lean/run/emptyLcnf.lean +++ b/tests/lean/run/emptyLcnf.lean @@ -10,6 +10,27 @@ set_option trace.Compiler.result true trace: [Compiler.result] size: 0 def f x : Nat := ⊥ +--- +trace: [Compiler.result] size: 5 + def _eval._lam_0 _x.1 _x.2 _y.3 _y.4 _y.5 _y.6 _y.7 _y.8 _y.9 : EStateM.Result Lean.Exception PUnit PUnit := + let _x.10 := Lean.Compiler.compile _x.1 _y.7 _y.8 _y.9; + cases _x.10 : EStateM.Result Lean.Exception PUnit PUnit + | EStateM.Result.ok a.11 a.12 => + let _x.13 := EStateM.Result.ok _ _ _ _x.2 a.12; + return _x.13 + | EStateM.Result.error a.14 a.15 => + return _x.10 +[Compiler.result] size: 8 + def _eval a.1 a.2 a.3 : EStateM.Result Lean.Exception PUnit PUnit := + let _x.4 := "f"; + let _x.5 := Lean.Name.mkStr1 _x.4; + let _x.6 := 1; + let _x.7 := Array.mkEmpty ◾ _x.6; + let _x.8 := Array.push ◾ _x.7 _x.5; + let _x.9 := PUnit.unit; + let _f.10 := _eval._lam_0 _x.8 _x.9; + let _x.11 := Lean.Elab.Command.liftTermElabM._redArg _f.10 a.1 a.2 a.3; + return _x.11 -/ #guard_msgs in run_meta Lean.Compiler.compile #[``f] diff --git a/tests/lean/run/erased.lean b/tests/lean/run/erased.lean index 3262dc052686..eb9809fdfde9 100644 --- a/tests/lean/run/erased.lean +++ b/tests/lean/run/erased.lean @@ -24,6 +24,36 @@ trace: [Compiler.result] size: 1 def Erased.mk (α : lcErased) (a : lcAny) : PSigma lcErased lcAny := let _x.1 : PSigma lcErased lcAny := PSigma.mk lcErased ◾ ◾ ◾; return _x.1 +--- +trace: [Compiler.result] size: 5 + def _eval._lam_0 (_x.1 : Array + Lean.Name) (_x.2 : PUnit) (_y.3 : Lean.Elab.Term.Context) (_y.4 : lcAny) (_y.5 : Lean.Meta.Context) (_y.6 : lcAny) (_y.7 : Lean.Core.Context) (_y.8 : lcAny) (_y.9 : PUnit) : EStateM.Result + Lean.Exception PUnit PUnit := + let _x.10 : EStateM.Result Lean.Exception PUnit PUnit := compile _x.1 _y.7 _y.8 _y.9; + cases _x.10 : EStateM.Result Lean.Exception PUnit PUnit + | EStateM.Result.ok (a.11 : PUnit) (a.12 : PUnit) => + let _x.13 : EStateM.Result Lean.Exception PUnit PUnit := EStateM.Result.ok Lean.Exception PUnit PUnit _x.2 a.12; + return _x.13 + | EStateM.Result.error (a.14 : Lean.Exception) (a.15 : PUnit) => + return _x.10 +[Compiler.result] size: 9 + def _eval (a.1 : Lean.Elab.Command.Context) (a.2 : lcAny) (a.3 : PUnit) : EStateM.Result Lean.Exception PUnit + PUnit := + let _x.4 : String := "Erased"; + let _x.5 : String := "mk"; + let _x.6 : Lean.Name := Lean.Name.mkStr2 _x.4 _x.5; + let _x.7 : Nat := nat_lit 1; + let _x.8 : Array Lean.Name := Array.mkEmpty ◾ _x.7; + let _x.9 : Array Lean.Name := Array.push ◾ _x.8 _x.6; + let _x.10 : PUnit := PUnit.unit; + let _f.11 : Lean.Elab.Term.Context → + lcAny → + Lean.Meta.Context → + lcAny → + Lean.Core.Context → lcAny → PUnit → EStateM.Result Lean.Exception PUnit PUnit := _eval._lam_0 _x.9 _x.10; + let _x.12 : EStateM.Result Lean.Exception PUnit + lcAny := Lean.Elab.Command.liftTermElabM._redArg _f.11 a.1 a.2 a.3; + return _x.12 -/ #guard_msgs in run_meta Lean.Compiler.compile #[``Erased.mk] diff --git a/tests/lean/run/extract_lets.lean b/tests/lean/run/extract_lets.lean index 1a815edf9e5f..d92fb9049aba 100644 --- a/tests/lean/run/extract_lets.lean +++ b/tests/lean/run/extract_lets.lean @@ -599,7 +599,7 @@ example : ∀ n : Nat, let x := n; let y := 2; x + y = x + y := by /-! Testing lots of `let`s -/ -set_option maxHeartbeats 300 in +set_option maxHeartbeats 600 in example : let x := 2 let x := let x := x; let x := x; let x := x; let x := x; let x := x; let x := x; let x := x; x diff --git a/tests/lean/new-compiler/run/inlineApp.lean b/tests/lean/run/inlineApp.lean similarity index 100% rename from tests/lean/new-compiler/run/inlineApp.lean rename to tests/lean/run/inlineApp.lean diff --git a/tests/lean/run/isDefEqProjIssue.lean b/tests/lean/run/isDefEqProjIssue.lean index 6c2347f1f0bd..e2e4312815c9 100644 --- a/tests/lean/run/isDefEqProjIssue.lean +++ b/tests/lean/run/isDefEqProjIssue.lean @@ -83,7 +83,7 @@ instance g (x : Nat) : Foo := { x, y := ack 10 11 } open Lean Meta -set_option maxHeartbeats 400 in +set_option maxHeartbeats 800 in run_meta do withLocalDeclD `x (mkConst ``Nat) fun x => do let lhs := Expr.proj ``Foo 0 <| mkApp (mkConst ``f) x diff --git a/tests/lean/run/kernelErrorFollowup.lean b/tests/lean/run/kernelErrorFollowup.lean index 91817536e9a6..4668f126c344 100644 --- a/tests/lean/run/kernelErrorFollowup.lean +++ b/tests/lean/run/kernelErrorFollowup.lean @@ -4,7 +4,11 @@ import Lean.Elab.Tactic.Basic Kernel errors should not lead to follow-up errors but should be detectable using `#print axioms`. -/ -/-- error: (kernel) declaration has metavariables 'bad' -/ +/-- +error: (kernel) declaration has metavariables 'bad' +--- +error: unknown metavariable '?_uniq.3' +-/ #guard_msgs in def bad : Empty := by run_tac do Lean.Elab.Tactic.popMainGoal diff --git a/tests/lean/run/kernel_maxheartbeats.lean b/tests/lean/run/kernel_maxheartbeats.lean index 3bccc519db02..a59c4e3e0c01 100644 --- a/tests/lean/run/kernel_maxheartbeats.lean +++ b/tests/lean/run/kernel_maxheartbeats.lean @@ -6,7 +6,7 @@ def ack : Nat → Nat → Nat | x+1, 0 => ack x 1 | x+1, y+1 => ack x (ack (x+1) y) -set_option maxHeartbeats 500 +set_option maxHeartbeats 1000 open Lean Meta /-- error: (kernel) deterministic timeout -/ diff --git a/tests/lean/run/lcnfErasure.lean b/tests/lean/run/lcnfErasure.lean index 369facb47f9a..f34570fe0efa 100644 --- a/tests/lean/run/lcnfErasure.lean +++ b/tests/lean/run/lcnfErasure.lean @@ -25,7 +25,7 @@ def checkMonoType! (type₁ type₂ : Expr) : MetaM Unit := do -- Decidable #eval checkMonoType! - (.const ``Decidable []) + (mkApp (.const ``Decidable []) (.const ``True [])) (.const ``Bool []) -- Prop diff --git a/tests/lean/run/match_expr_perf.lean b/tests/lean/run/match_expr_perf.lean index 08c2379e6604..98bb2acd9fd2 100644 --- a/tests/lean/run/match_expr_perf.lean +++ b/tests/lean/run/match_expr_perf.lean @@ -11,7 +11,7 @@ import Lean.Elab.Tactic.Config open Lean Meta Omega -set_option maxHeartbeats 5000 +set_option maxHeartbeats 10000 def pushNot (h P : Expr) : MetaM (Option Expr) := do let P ← whnfR P trace[omega] "pushing negation: {P}" diff --git a/tests/lean/run/noncomputable_decide.lean b/tests/lean/run/noncomputable_decide.lean index 3d7a06543cea..a03ae835ace6 100644 --- a/tests/lean/run/noncomputable_decide.lean +++ b/tests/lean/run/noncomputable_decide.lean @@ -1,11 +1,11 @@ open scoped Classical /-- -error: tactic 'native_decide' failed, could not evaluate decidable instance. Error: (interpreter) -unknown declaration 'ohno._nativeDecide_1' +error: tactic 'native_decide' failed, could not evaluate decidable instance. +Error: (interpreter) unknown declaration 'ohno._nativeDecide_1' --- -error: failed to compile definition, consider marking it as 'noncomputable' because it depends on -'Classical.propDecidable', which is 'noncomputable' +error: failed to compile definition, compiler IR check failed at 'ohno._nativeDecide_1._lam_0'. +Error: depends on declaration 'Classical.propDecidable', which has no executable code; consider marking definition as 'noncomputable' -/ #guard_msgs in theorem ohno : False := by diff --git a/tests/lean/run/structWithAlgTCSynth.lean b/tests/lean/run/structWithAlgTCSynth.lean index 55e289abac55..fcd39ba3a3e0 100644 --- a/tests/lean/run/structWithAlgTCSynth.lean +++ b/tests/lean/run/structWithAlgTCSynth.lean @@ -1301,6 +1301,6 @@ Typeclass synthesis should remain fast when multiple `with` patterns are nested Prior to #2478, this requires over 30000 heartbeats. -/ -set_option synthInstance.maxHeartbeats 400 in +set_option synthInstance.maxHeartbeats 800 in instance instAlgebra' (R M : Type _) [CommRing R] (I : Ideal (Quot_r R M)) : Algebra R ((Quot_r R M) ⧸ I) := inferInstance diff --git a/tests/lean/run/timeTzifParse.lean b/tests/lean/run/timeTzifParse.lean index a82603691cba..b013d1cae39e 100644 --- a/tests/lean/run/timeTzifParse.lean +++ b/tests/lean/run/timeTzifParse.lean @@ -1,6 +1,7 @@ import Std.Time open Std.Time +set_option maxHeartbeats 2000000 def file := ByteArray.mk <| #[84, 90, 105, 102, 50, 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, 91, 0] ++ #[0, 0, 3, 0, 0, 0, 12, 150, 170, 114, 180, 184, 15, 73, 224, 184, 253, 64, 160, 185, 241, 52, 48, 186, 222, 116, 32, 218, 56, 174, 48, 218, 235, 250, 48, 220, 25] diff --git a/tests/lean/new-compiler/seqToCodeIssue.lean b/tests/lean/seqToCodeIssue.lean similarity index 100% rename from tests/lean/new-compiler/seqToCodeIssue.lean rename to tests/lean/seqToCodeIssue.lean diff --git a/tests/lean/new-compiler/seqToCodeIssue.lean.expected.out b/tests/lean/seqToCodeIssue.lean.expected.out similarity index 100% rename from tests/lean/new-compiler/seqToCodeIssue.lean.expected.out rename to tests/lean/seqToCodeIssue.lean.expected.out diff --git a/tests/lean/sint_basic.lean.expected.out b/tests/lean/sint_basic.lean.expected.out index 40f67f40be26..f369e8e5275a 100644 --- a/tests/lean/sint_basic.lean.expected.out +++ b/tests/lean/sint_basic.lean.expected.out @@ -72,16 +72,15 @@ true true true true - -[result] -def myId8 (x_1 : u8) : u8 := - ret x_1 -def myId8._boxed (x_1 : obj) : obj := - let x_2 : u8 := unbox x_1; - dec x_1; - let x_3 : u8 := myId8 x_2; - let x_4 : obj := box x_3; - ret x_4 +[Compiler.IR] [result] + def myId8 (x_1 : u8) : u8 := + ret x_1 + def myId8._boxed (x_1 : obj) : obj := + let x_2 : u8 := unbox x_1; + dec x_1; + let x_3 : u8 := myId8 x_2; + let x_4 : obj := box x_3; + ret x_4 Int16 : Type 20 -20 @@ -156,16 +155,15 @@ true true true true - -[result] -def myId16 (x_1 : u16) : u16 := - ret x_1 -def myId16._boxed (x_1 : obj) : obj := - let x_2 : u16 := unbox x_1; - dec x_1; - let x_3 : u16 := myId16 x_2; - let x_4 : obj := box x_3; - ret x_4 +[Compiler.IR] [result] + def myId16 (x_1 : u16) : u16 := + ret x_1 + def myId16._boxed (x_1 : obj) : obj := + let x_2 : u16 := unbox x_1; + dec x_1; + let x_3 : u16 := myId16 x_2; + let x_4 : obj := box x_3; + ret x_4 Int32 : Type 20 -20 @@ -240,16 +238,15 @@ true true true true - -[result] -def myId32 (x_1 : u32) : u32 := - ret x_1 -def myId32._boxed (x_1 : obj) : obj := - let x_2 : u32 := unbox x_1; - dec x_1; - let x_3 : u32 := myId32 x_2; - let x_4 : obj := box x_3; - ret x_4 +[Compiler.IR] [result] + def myId32 (x_1 : u32) : u32 := + ret x_1 + def myId32._boxed (x_1 : obj) : obj := + let x_2 : u32 := unbox x_1; + dec x_1; + let x_3 : u32 := myId32 x_2; + let x_4 : obj := box x_3; + ret x_4 Int64 : Type 20 -20 @@ -324,16 +321,15 @@ true true true true - -[result] -def myId64 (x_1 : u64) : u64 := - ret x_1 -def myId64._boxed (x_1 : obj) : obj := - let x_2 : u64 := unbox x_1; - dec x_1; - let x_3 : u64 := myId64 x_2; - let x_4 : obj := box x_3; - ret x_4 +[Compiler.IR] [result] + def myId64 (x_1 : u64) : u64 := + ret x_1 + def myId64._boxed (x_1 : obj) : obj := + let x_2 : u64 := unbox x_1; + dec x_1; + let x_3 : u64 := myId64 x_2; + let x_4 : obj := box x_3; + ret x_4 ISize : Type 20 -20 @@ -408,13 +404,12 @@ true true true true - -[result] -def myIdSize (x_1 : usize) : usize := - ret x_1 -def myIdSize._boxed (x_1 : obj) : obj := - let x_2 : usize := unbox x_1; - dec x_1; - let x_3 : usize := myIdSize x_2; - let x_4 : obj := box x_3; - ret x_4 +[Compiler.IR] [result] + def myIdSize (x_1 : usize) : usize := + ret x_1 + def myIdSize._boxed (x_1 : obj) : obj := + let x_2 : usize := unbox x_1; + dec x_1; + let x_3 : usize := myIdSize x_2; + let x_4 : obj := box x_3; + ret x_4 diff --git a/tests/lean/unboxStruct.lean.expected.out b/tests/lean/unboxStruct.lean.expected.out index feffd3756382..ecb35b1491cd 100644 --- a/tests/lean/unboxStruct.lean.expected.out +++ b/tests/lean/unboxStruct.lean.expected.out @@ -1,10 +1,9 @@ - -[result] -def test2 (x_1 : u32) (x_2 : obj) : obj := - let x_3 : obj := foo x_1 x_2; - ret x_3 -def test2._boxed (x_1 : obj) (x_2 : obj) : obj := - let x_3 : u32 := unbox x_1; - dec x_1; - let x_4 : obj := test2 x_3 x_2; - ret x_4 +[Compiler.IR] [result] + def test2 (x_1 : u32) (x_2 : obj) : obj := + let x_3 : obj := foo x_1 x_2; + ret x_3 + def test2._boxed (x_1 : obj) (x_2 : obj) : obj := + let x_3 : u32 := unbox x_1; + dec x_1; + let x_4 : obj := test2 x_3 x_2; + ret x_4 diff --git a/tests/lean/unhygienicCode.lean.expected.out b/tests/lean/unhygienicCode.lean.expected.out index e69de29bb2d1..94696e118c24 100644 --- a/tests/lean/unhygienicCode.lean.expected.out +++ b/tests/lean/unhygienicCode.lean.expected.out @@ -0,0 +1,23 @@ +[Compiler.result] size: 20 + def foo n : Syntax := + let _x.1 := Syntax.missing; + let _x.2 := 1; + let _x.3 := false; + let _x.4 := @SourceInfo.fromRef _x.1 _x.3; + let _x.5 := "UnhygienicMain"; + let _x.6 := Name.mkStr1 _x.5; + let _x.7 := "term_+_"; + let _x.8 := Name.mkStr1 _x.7; + let _x.9 := "a"; + let _x.10 := String.toSubstring' _x.9; + let _x.11 := Name.mkStr1 _x.9; + let _x.12 := addMacroScope _x.6 _x.11 _x.2; + let _x.13 := [] _; + let _x.14 := Syntax.ident _x.4 _x.10 _x.12 _x.13; + let _x.15 := "+"; + let _x.16 := Syntax.atom _x.4 _x.15; + let _x.17 := Nat.reprFast.0 n; + let _x.18 := SourceInfo.none; + let _x.19 := @Syntax.mkNumLit _x.17 _x.18; + let _x.20 := Syntax.node3 _x.4 _x.8 _x.14 _x.16 _x.19; + return _x.20 diff --git a/tests/lean/updateExprIssue.lean.expected.out b/tests/lean/updateExprIssue.lean.expected.out index 10537d1e321f..8fae34f06947 100644 --- a/tests/lean/updateExprIssue.lean.expected.out +++ b/tests/lean/updateExprIssue.lean.expected.out @@ -1,45 +1,26 @@ - -[init] -def sefFn (x_1 : obj) (x_2 : obj) : obj := - case x_1 : obj of - Lean.Expr.bvar._impl → - ret x_1 - Lean.Expr.fvar._impl → - ret x_1 - Lean.Expr.mvar._impl → - ret x_1 - Lean.Expr.sort._impl → - ret x_1 - Lean.Expr.const._impl → - ret x_1 - Lean.Expr.app._impl → - let x_3 : obj := proj[0] x_1; - let x_4 : obj := proj[1] x_1; - let x_5 : usize := ptrAddrUnsafe ◾ x_3; - let x_6 : usize := ptrAddrUnsafe ◾ x_2; - let x_7 : u8 := USize.decEq x_5 x_6; - case x_7 : obj of - Bool.false → - let x_8 : obj := Lean.Expr.app._override x_2 x_4; - ret x_8 - Bool.true → - let x_9 : usize := ptrAddrUnsafe ◾ x_4; - let x_10 : u8 := USize.decEq x_9 x_9; - case x_10 : obj of - Bool.false → - let x_11 : obj := Lean.Expr.app._override x_2 x_4; - ret x_11 - Bool.true → +[Compiler.IR] [init] + def sefFn (x_1 : obj) (x_2 : obj) : obj := + case x_1 : obj of + Lean.Expr.app._impl → + let x_3 : u64 := sproj[2, 0] x_1; + let x_4 : obj := proj[0] x_1; + let x_5 : obj := proj[1] x_1; + block_6 (x_7 : u8) := + case x_7 : obj of + Bool.false → + let x_8 : obj := Lean.Expr.app._override x_2 x_5; + ret x_8 + Bool.true → + ret x_1; + let x_9 : usize := ptrAddrUnsafe ◾ x_4; + let x_10 : usize := ptrAddrUnsafe ◾ x_2; + let x_11 : u8 := USize.decEq x_9 x_10; + case x_11 : obj of + Bool.false → + jmp block_6 x_11 + Bool.true → + let x_12 : usize := ptrAddrUnsafe ◾ x_5; + let x_13 : u8 := USize.decEq x_12 x_12; + jmp block_6 x_13 + default → ret x_1 - Lean.Expr.lam._impl → - ret x_1 - Lean.Expr.forallE._impl → - ret x_1 - Lean.Expr.letE._impl → - ret x_1 - Lean.Expr.lit._impl → - ret x_1 - Lean.Expr.mdata._impl → - ret x_1 - Lean.Expr.proj._impl → - ret x_1