diff --git a/src/Init/Prelude.lean b/src/Init/Prelude.lean index e177fd10a8f2..18b1765ae89f 100644 --- a/src/Init/Prelude.lean +++ b/src/Init/Prelude.lean @@ -535,6 +535,8 @@ structure PProd (α : Sort u) (β : Sort v) where /-- The second element of a pair. -/ snd : β +attribute [unbox] PProd + /-- A product type in which both `α` and `β` are in the same universe. @@ -546,6 +548,8 @@ structure MProd (α β : Type u) where /-- The second element of a pair. -/ snd : β +attribute [unbox] MProd + /-- `And a b`, or `a ∧ b`, is the conjunction of propositions. It can be constructed and destructed like a pair: if `ha : a` and `hb : b` then diff --git a/src/Init/System/ST.lean b/src/Init/System/ST.lean index a35df194af79..358f62b9292b 100644 --- a/src/Init/System/ST.lean +++ b/src/Init/System/ST.lean @@ -65,6 +65,7 @@ instance {σ : Type} : MonadAttach (ST σ) where CanReturn x a := ∃ s s', x s = ⟨a, s'⟩ attach x s := match h : x s with | ⟨a, s'⟩ => ⟨⟨a, s, s', h⟩, s'⟩ +@[unbox] inductive EST.Out (ε : Type) (σ : Type) (α : Type) where | ok : α → Void σ → EST.Out ε σ α | error : ε → Void σ → EST.Out ε σ α diff --git a/src/Lean/Compiler/IR.lean b/src/Lean/Compiler/IR.lean index 5dd568b1407f..22651120eb0b 100644 --- a/src/Lean/Compiler/IR.lean +++ b/src/Lean/Compiler/IR.lean @@ -48,18 +48,22 @@ def compile (decls : Array Decl) : CompilerM (Array Decl) := do let mut decls := decls decls := decls.map Decl.pushProj logDecls `push_proj decls + decls := prepareBoxParams (← getEnv) decls + logDecls `box_params decls + decls ← explicitBoxing decls + logDecls `boxing decls if compiler.reuse.get (← getOptions) then decls := decls.map (Decl.insertResetReuse (← getEnv)) logDecls `reset_reuse decls - decls := decls.map Decl.elimDead - logDecls `elim_dead decls decls := decls.map Decl.simpCase logDecls `simp_case decls decls := decls.map Decl.normalizeIds + decls := decls.map Decl.elimDead + logDecls `elim_dead decls decls ← inferBorrow decls logDecls `borrow decls - decls ← explicitBoxing decls - logDecls `boxing decls + decls := ExplicitBoxing.addBoxedVersions (← getEnv) decls + logDecls `boxed_versions decls decls ← explicitRC decls logDecls `rc decls if compiler.reuse.get (← getOptions) then @@ -70,6 +74,10 @@ def compile (decls : Array Decl) : CompilerM (Array Decl) := do decls ← updateSorryDep decls logDecls `result decls checkDecls decls + if (← getOptions).getBool (tracePrefixOptionName ++ `struct_rc) || + (← getOptions).getBool tracePrefixOptionName then + let decls2 := decls.map StructRC.visitDecl + log (LogEntry.step `struct_rc decls2) decls ← toposortDecls decls addDecls decls inferMeta decls diff --git a/src/Lean/Compiler/IR/Basic.lean b/src/Lean/Compiler/IR/Basic.lean index 3deace0e7709..0b11ab604c82 100644 --- a/src/Lean/Compiler/IR/Basic.lean +++ b/src/Lean/Compiler/IR/Basic.lean @@ -54,30 +54,41 @@ instance : ToString JoinPointId := ⟨fun a => "block_" ++ toString a.idx⟩ - `void` is used to identify uses of the state token from `BaseIO` which do no longer need to be passed around etc. at this point in the pipeline. - - `struct` and `union` are used to return small values (e.g., `Option`, `Prod`, `Except`) - on the stack. + - `struct` and `union` are used for unboxed values that should be stored on the stack + (e.g., `Option`, `Prod`, `Except`). Note that `struct` and `union` are stored boxed + (i.e. as `lean_ctor_object`s) in the interpreter. Remark: the RC operations for `tobject` are slightly more expensive because we first need to test whether the `tobject` is really a pointer or not. -Remark: the Lean runtime assumes that sizeof(void*) == sizeof(sizeT). -Lean cannot be compiled on old platforms where this is not True. - -Since values of type `struct` and `union` are only used to return values, -We assume they must be used/consumed "linearly". We use the term "linear" here -to mean "exactly once" in each execution. That is, given `x : S`, where `S` is a struct, -then one of the following must hold in each (execution) branch. -1- `x` occurs only at a single `ret x` instruction. That is, it is being consumed by being returned. -2- `x` occurs only at a single `ctor`. That is, it is being "consumed" by being stored into another `struct/union`. -3- We extract (aka project) every single field of `x` exactly once. That is, we are consuming `x` by consuming each - of one of its components. Minor refinement: we don't need to consume scalar fields or struct/union - fields that do not contain object fields. +Remark: the Lean runtime assumes that `sizeof(void*) == sizeof(size_t)`. +Lean cannot be compiled on old platforms where this is not true. -/ inductive IRType where | float | uint8 | uint16 | uint32 | uint64 | usize | erased | object | tobject | float32 - | struct (leanTypeName : Option Name) (types : Array IRType) : IRType + /-- + Unboxed structures representing a particular constructor of an inductive type. + The type `objects[i]` is the type of `proj[0, i] x` where `x` is a value of this type, or more + generally `proj[c, i]` when this type is the `c`-th constructor of a `union` type. Thus, the + amount of `objects` should be exactly `CtorInfo.size` for the corresponding constructor. + + `usize` and `ssize` each are the number of `size_t` values and the amount of bytes spent by other + scalars respectively and should be equivalent to the values in `CtorInfo.usize` and + `CtorInfo.ssize` for the corresponding constructor. + + There are two different models of `struct` types: the interpreter model, which stores `struct`s + as `lean_ctor_object`s that have their own reference count, and the compiled model, in which + `struct`s are stored as corresponding C `struct`s and RC operations on a `struct` correspond to + RC operations on every field. + -/ + | struct (leanTypeName : Option Name) + (objects : Array IRType) (usize ssize : Nat) : IRType + /-- + Unboxed tagged unions of multiple IR types. Each type in `types` should be a `struct` + corresponding to a constructor of the inductive type `leanTypeName`. + -/ | union (leanTypeName : Name) (types : Array IRType) : IRType -- TODO: Move this upwards after a stage0 update. | tagged @@ -87,28 +98,35 @@ inductive IRType where namespace IRType def isScalar : IRType → Bool - | float => true - | float32 => true - | uint8 => true - | uint16 => true - | uint32 => true - | uint64 => true - | usize => true - | _ => false + | float | float32 | uint8 | uint16 | uint32 | uint64 | usize => true + | _ => false def isObj : IRType → Bool - | object => true - | tagged => true - | tobject => true - | void => true - | _ => false + | object | tagged | tobject | void => true + | _ => false + +def isStruct : IRType → Bool + | struct _ _ _ _ | union _ _ => true + | _ => false + +def isObjOrStruct : IRType → Bool + | object | tagged | tobject | void | struct .. | union .. => true + | _ => false + +def isScalarOrStruct : IRType → Bool + | float | float32 | uint8 | uint16 | uint32 | uint64 | usize + | struct _ _ _ _ | union _ _ => true + | _ => false def isPossibleRef : IRType → Bool - | object | tobject => true + | object | tobject | struct .. | union .. => true | _ => false def isDefiniteRef : IRType → Bool | object => true + | union _ tys => tys.all (!· matches struct _ #[] 0 0) + | struct _ #[] 0 0 => false + | struct .. => true | _ => false def isErased : IRType → Bool @@ -119,10 +137,53 @@ def isVoid : IRType → Bool | void => true | _ => false +/-- +Returns `tagged`, `object` or `tobject` depending on whether the provided type is a definite +reference and/or a possible reference in boxed form. For details on boxing, see `Expr.box`. +-/ def boxed : IRType → IRType | object | float | float32 => object - | void | tagged | uint8 | uint16 => tagged - | _ => tobject + | void | erased | tagged | uint8 | uint16 => tagged + | union _ tys => if tys.any (· matches struct _ #[] 0 0) then tobject else object + | struct _ #[] 0 0 => tagged + | struct .. => object + | tobject | uint32 | uint64 | usize => tobject + +/-- +Normalize the object parts of the IR type, i.e. convert it into a type that is `compatibleWith` the +provided type but with all occurrences of `object` and `tagged` replaced by `tobject` and all +occurrences of `void` replaced by `erased`. +-/ +def normalizeObject : IRType → IRType + | object | tagged => tobject + | struct nm tys us ss => struct nm (tys.map normalizeObject) us ss + | union nm tys => union nm (tys.map normalizeObject) + | void => erased + | ty => ty + +/-- +Returns whether the two types have compatible representation, that is they are equal modulo +different reference types. If `strict` is `false` (default), object types (`object`, `tobject`, +`tagged`) are considered to be compatible with erased types (`erased`, `void`). +-/ +partial def compatibleWith (a b : IRType) (strict : Bool := false) : Bool := + match a, b with + | float, float => true + | uint8, uint8 => true + | uint16, uint16 => true + | uint32, uint32 => true + | uint64, uint64 => true + | usize, usize => true + | float32, float32 => true + | erased, t | void, t => + (t matches erased | void) || (!strict && t matches object | tobject | tagged) + | object, t | tobject, t | tagged, t => + (t matches object | tobject | tagged) || (!strict && t matches erased | void) + | struct _ tys us ss, struct _ tys' us' ss' => + us == us' && ss == ss' && tys.isEqv tys' (compatibleWith (strict := true)) + | union _ tys, union _ tys' => + tys.isEqv tys' (compatibleWith (strict := false)) + | _, _ => false end IRType @@ -174,30 +235,56 @@ def CtorInfo.type (info : CtorInfo) : IRType := if info.isRef then .object else .tagged inductive Expr where - /-- We use `ctor` mainly for constructing Lean object/tobject values `lean_ctor_object` in the runtime. - This instruction is also used to creat `struct` and `union` return values. - For `union`, only `i.cidx` is relevant. For `struct`, `i` is irrelevant. -/ + /-- + `let x : obj := ctor[i] ys...` allocates a `lean_ctor_object` with the tag `i.cidx` and a scalar + capacity of `i.usize * sizeof(size_t) + i.ssize`. If `i.isScalar` is true, instead returns + `lean_box(i.cidx)`. All arguments should be boxed values. + + This instruction is also used to create `struct` and `union` values where `i.cidx` is used for + the tag and `ys` should have types corresponding to the `objects` parameter of `IRType.struct`. + -/ | ctor (i : CtorInfo) (ys : Array Arg) | reset (n : Nat) (x : VarId) /-- `reuse x in ctor_i ys` instruction in the paper. -/ | reuse (x : VarId) (i : CtorInfo) (updtHeader : Bool) (ys : Array Arg) - /-- Extract the `tobject` value at Position `sizeof(void*)*i` from `x`. - We also use `proj` for extracting fields from `struct` return values, and casting `union` return values. -/ - | proj (i : Nat) (x : VarId) - /-- Extract the `Usize` value at Position `sizeof(void*)*i` from `x`. -/ - | uproj (i : Nat) (x : VarId) + /-- Extract the `tobject` value at Position `sizeof(void*)*i` from `x`. When `x` is represented + as a `union`, `cidx` is used to determine which part of the union to access. -/ + | proj (cidx : Nat) (i : Nat) (x : VarId) + /-- Extract the `usize` value at Position `sizeof(void*)*i` from `x`. -/ + | uproj (cidx : Nat) (i : Nat) (x : VarId) /-- Extract the scalar value at Position `sizeof(void*)*n + offset` from `x`. -/ - | sproj (n : Nat) (offset : Nat) (x : VarId) + | sproj (cidx : Nat) (n : Nat) (offset : Nat) (x : VarId) /-- Full application. -/ | fap (c : FunId) (ys : Array Arg) /-- Partial application that creates a `pap` value (aka closure in our nonstandard terminology). -/ | pap (c : FunId) (ys : Array Arg) /-- Application. `x` must be a `pap` value. -/ | ap (x : VarId) (ys : Array Arg) - /-- Given `x : ty` where `ty` is a scalar type, this operation returns a value of Type `tobject`. - For small scalar values, the Result is a tagged pointer, and no memory allocation is performed. -/ + /-- + Converts an owned value into another owned value. The `box` instruction comes in three variants: + 1. Given `x : ty` where `ty` is a scalar type, return a value of type `tobject`. + For small scalar values, the result is a tagged pointer, and no memory allocation is performed. + 2. Given `x : ty` where `ty` is a struct/union type, return a value of type `object` or `tobject` + (depending on whether `ty` has scalar constructors). This operation is the identity function + in the interpreter model and is an allocation is the compiled model. + 3. Given `x : ty` where `ty` is a struct/union type, return a value of another struct/union type + where different parts are boxed/unboxed. This variant is called a *reshape* and can contain + both boxing and unboxing. For example, `let y : {{tobj, tobj}, obj} := box x` with + `x : {obj, {tobj, tobj}}` unboxes the first component of `x` and boxes the second component + of `x` into a `lean_ctor_object`. This is also the identity function in the interpreter. + + The last two variants are distinguished by the type of the variable declaration. + -/ | box (ty : IRType) (x : VarId) - /-- Given `x : [t]object`, obtain the scalar value. -/ + /-- + Unboxed a boxed (`tagged` / `object` / `tobject`) value into a scalar or `struct` / `union` value. + The input value is taken borrowed and the result value is borrowed from the input value. + That is, no reference counting operations are performed and the RC of the result value may need + to be manually incremented and the RC of the input value may need to be decremented. + + In the interpreter, unboxing into a `struct` or `union` is the identity function since `struct` + and `union` values are stored boxed in the interpreter. + -/ | unbox (x : VarId) | lit (v : LitVal) /-- Return `1 : uint8` Iff `RC(x) > 1` -/ @@ -225,16 +312,22 @@ inductive FnBody where This operation is not part of λPure is only used during optimization. -/ | set (x : VarId) (i : Nat) (y : Arg) (b : FnBody) | setTag (x : VarId) (cidx : Nat) (b : FnBody) - /-- Store `y : Usize` at Position `sizeof(void*)*i` in `x`. `x` must be a Constructor object and `RC(x)` must be 1. -/ - | uset (x : VarId) (i : Nat) (y : VarId) (b : FnBody) - /-- Store `y : ty` at Position `sizeof(void*)*i + offset` in `x`. `x` must be a Constructor object and `RC(x)` must be 1. - `ty` must not be `object`, `tobject`, `erased` nor `Usize`. -/ - | sset (x : VarId) (i : Nat) (offset : Nat) (y : VarId) (ty : IRType) (b : FnBody) - /-- RC increment for `object`. If c == `true`, then `inc` must check whether `x` is a tagged pointer or not. - If `persistent == true` then `x` is statically known to be a persistent object. -/ + /-- Store `y : usize` at Position `sizeof(void*)*i` in `x`. `x` must be a Constructor object or + `struct`/`union` with tag `cidx` and `RC(x)` must be 1. -/ + | uset (x : VarId) (cidx : Nat) (i : Nat) (y : VarId) (b : FnBody) + /-- Store `y : ty` at Position `sizeof(void*)*i + offset` in `x`. `x` must be a Constructor object + or `struct`/`union` with tag `cidx` and `RC(x)` must be 1. `ty` must be a scalar type but not + `usize`. -/ + | sset (x : VarId) (cidx : Nat) (i : Nat) (offset : Nat) (y : VarId) (ty : IRType) (b : FnBody) + /-- RC increment for `object`. If c is `true`, then `inc` must check whether `x` is a tagged + pointer or not. If `persistent` is `true` then `x` is statically known to be a persistent object + and this operation does not need to be performed in compiled code. In compiled code, if the type + of `x` is a `struct` or `union` type, then instead all fields are incremented. -/ | inc (x : VarId) (n : Nat) (c : Bool) (persistent : Bool) (b : FnBody) - /-- RC decrement for `object`. If c == `true`, then `inc` must check whether `x` is a tagged pointer or not. - If `persistent == true` then `x` is statically known to be a persistent object. -/ + /-- RC decrement for `object`. If c is `true`, then `inc` must check whether `x` is a tagged + pointer or not. If `persistent` is `true` then `x` is statically known to be a persistent object + and this operation does not need to be performed in compiled code. In compiled code, if the type + of `x` is a `struct` or `union` type, then instead all fields are decremented. -/ | dec (x : VarId) (n : Nat) (c : Bool) (persistent : Bool) (b : FnBody) | del (x : VarId) (b : FnBody) | case (tid : Name) (x : VarId) (xType : IRType) (cs : Array Alt) @@ -258,28 +351,28 @@ def FnBody.isTerminal : FnBody → Bool | _ => false def FnBody.body : FnBody → FnBody - | FnBody.vdecl _ _ _ b => b - | FnBody.jdecl _ _ _ b => b - | FnBody.set _ _ _ b => b - | FnBody.uset _ _ _ b => b - | FnBody.sset _ _ _ _ _ b => b - | FnBody.setTag _ _ b => b - | FnBody.inc _ _ _ _ b => b - | FnBody.dec _ _ _ _ b => b - | FnBody.del _ b => b - | other => other + | FnBody.vdecl _ _ _ b => b + | FnBody.jdecl _ _ _ b => b + | FnBody.set _ _ _ b => b + | FnBody.uset _ _ _ _ b => b + | FnBody.sset _ _ _ _ _ _ b => b + | FnBody.setTag _ _ b => b + | FnBody.inc _ _ _ _ b => b + | FnBody.dec _ _ _ _ b => b + | FnBody.del _ b => b + | other => other def FnBody.setBody : FnBody → FnBody → FnBody - | FnBody.vdecl x t v _, b => FnBody.vdecl x t v b - | FnBody.jdecl j xs v _, b => FnBody.jdecl j xs v b - | FnBody.set x i y _, b => FnBody.set x i y b - | FnBody.uset x i y _, b => FnBody.uset x i y b - | FnBody.sset x i o y t _, b => FnBody.sset x i o y t b - | FnBody.setTag x i _, b => FnBody.setTag x i b - | FnBody.inc x n c p _, b => FnBody.inc x n c p b - | FnBody.dec x n c p _, b => FnBody.dec x n c p b - | FnBody.del x _, b => FnBody.del x b - | other, _ => other + | FnBody.vdecl x t v _, b => FnBody.vdecl x t v b + | FnBody.jdecl j xs v _, b => FnBody.jdecl j xs v b + | FnBody.set x i y _, b => FnBody.set x i y b + | FnBody.uset x c i y _, b => FnBody.uset x c i y b + | FnBody.sset x c i o y t _, b => FnBody.sset x c i o y t b + | FnBody.setTag x i _, b => FnBody.setTag x i b + | FnBody.inc x n c p _, b => FnBody.inc x n c p b + | FnBody.dec x n c p _, b => FnBody.dec x n c p b + | FnBody.del x _, b => FnBody.del x b + | other, _ => other @[inline] def FnBody.resetBody (b : FnBody) : FnBody := b.setBody FnBody.nil @@ -483,9 +576,9 @@ def Expr.alphaEqv (ρ : IndexRenaming) : Expr → Expr → Bool | Expr.ctor i₁ ys₁, Expr.ctor i₂ ys₂ => i₁ == i₂ && aeqv ρ ys₁ ys₂ | Expr.reset n₁ x₁, Expr.reset n₂ x₂ => n₁ == n₂ && aeqv ρ x₁ x₂ | Expr.reuse x₁ i₁ u₁ ys₁, Expr.reuse x₂ i₂ u₂ ys₂ => aeqv ρ x₁ x₂ && i₁ == i₂ && u₁ == u₂ && aeqv ρ ys₁ ys₂ - | Expr.proj i₁ x₁, Expr.proj i₂ x₂ => i₁ == i₂ && aeqv ρ x₁ x₂ - | Expr.uproj i₁ x₁, Expr.uproj i₂ x₂ => i₁ == i₂ && aeqv ρ x₁ x₂ - | Expr.sproj n₁ o₁ x₁, Expr.sproj n₂ o₂ x₂ => n₁ == n₂ && o₁ == o₂ && aeqv ρ x₁ x₂ + | Expr.proj _ i₁ x₁, Expr.proj _ i₂ x₂ => i₁ == i₂ && aeqv ρ x₁ x₂ + | Expr.uproj _ i₁ x₁, Expr.uproj _ i₂ x₂ => i₁ == i₂ && aeqv ρ x₁ x₂ + | Expr.sproj _ n₁ o₁ x₁, Expr.sproj _ n₂ o₂ x₂ => n₁ == n₂ && o₁ == o₂ && aeqv ρ x₁ x₂ | Expr.fap c₁ ys₁, Expr.fap c₂ ys₂ => c₁ == c₂ && aeqv ρ ys₁ ys₂ | Expr.pap c₁ ys₁, Expr.pap c₂ ys₂ => c₁ == c₂ && aeqv ρ ys₁ ys₂ | Expr.ap x₁ ys₁, Expr.ap x₂ ys₂ => aeqv ρ x₁ x₂ && aeqv ρ ys₁ ys₂ @@ -516,27 +609,40 @@ def addParamsRename (ρ : IndexRenaming) (ps₁ ps₂ : Array Param) : Option In pure ρ partial def FnBody.alphaEqv : IndexRenaming → FnBody → FnBody → Bool - | ρ, FnBody.vdecl x₁ t₁ v₁ b₁, FnBody.vdecl x₂ t₂ v₂ b₂ => t₁ == t₂ && aeqv ρ v₁ v₂ && alphaEqv (addVarRename ρ x₁.idx x₂.idx) b₁ b₂ - | ρ, FnBody.jdecl j₁ ys₁ v₁ b₁, FnBody.jdecl j₂ ys₂ v₂ b₂ => match addParamsRename ρ ys₁ ys₂ with + | ρ, FnBody.vdecl x₁ t₁ v₁ b₁, FnBody.vdecl x₂ t₂ v₂ b₂ => + t₁ == t₂ && aeqv ρ v₁ v₂ && alphaEqv (addVarRename ρ x₁.idx x₂.idx) b₁ b₂ + | ρ, FnBody.jdecl j₁ ys₁ v₁ b₁, FnBody.jdecl j₂ ys₂ v₂ b₂ => + match addParamsRename ρ ys₁ ys₂ with | some ρ' => alphaEqv ρ' v₁ v₂ && alphaEqv (addVarRename ρ j₁.idx j₂.idx) b₁ b₂ | none => false - | ρ, FnBody.set x₁ i₁ y₁ b₁, FnBody.set x₂ i₂ y₂ b₂ => aeqv ρ x₁ x₂ && i₁ == i₂ && aeqv ρ y₁ y₂ && alphaEqv ρ b₁ b₂ - | ρ, FnBody.uset x₁ i₁ y₁ b₁, FnBody.uset x₂ i₂ y₂ b₂ => aeqv ρ x₁ x₂ && i₁ == i₂ && aeqv ρ y₁ y₂ && alphaEqv ρ b₁ b₂ - | ρ, FnBody.sset x₁ i₁ o₁ y₁ t₁ b₁, FnBody.sset x₂ i₂ o₂ y₂ t₂ b₂ => + | ρ, FnBody.set x₁ i₁ y₁ b₁, FnBody.set x₂ i₂ y₂ b₂ => + aeqv ρ x₁ x₂ && i₁ == i₂ && aeqv ρ y₁ y₂ && alphaEqv ρ b₁ b₂ + | ρ, FnBody.uset x₁ _ i₁ y₁ b₁, FnBody.uset x₂ _ i₂ y₂ b₂ => + aeqv ρ x₁ x₂ && i₁ == i₂ && aeqv ρ y₁ y₂ && alphaEqv ρ b₁ b₂ + | ρ, FnBody.sset x₁ _ i₁ o₁ y₁ t₁ b₁, FnBody.sset x₂ _ i₂ o₂ y₂ t₂ b₂ => aeqv ρ x₁ x₂ && i₁ = i₂ && o₁ = o₂ && aeqv ρ y₁ y₂ && t₁ == t₂ && alphaEqv ρ b₁ b₂ - | ρ, FnBody.setTag x₁ i₁ b₁, FnBody.setTag x₂ i₂ b₂ => aeqv ρ x₁ x₂ && i₁ == i₂ && alphaEqv ρ b₁ b₂ - | ρ, FnBody.inc x₁ n₁ c₁ p₁ b₁, FnBody.inc x₂ n₂ c₂ p₂ b₂ => aeqv ρ x₁ x₂ && n₁ == n₂ && c₁ == c₂ && p₁ == p₂ && alphaEqv ρ b₁ b₂ - | ρ, FnBody.dec x₁ n₁ c₁ p₁ b₁, FnBody.dec x₂ n₂ c₂ p₂ b₂ => aeqv ρ x₁ x₂ && n₁ == n₂ && c₁ == c₂ && p₁ == p₂ && alphaEqv ρ b₁ b₂ - | ρ, FnBody.del x₁ b₁, FnBody.del x₂ b₂ => aeqv ρ x₁ x₂ && alphaEqv ρ b₁ b₂ - | ρ, FnBody.case n₁ x₁ _ alts₁, FnBody.case n₂ x₂ _ alts₂ => n₁ == n₂ && aeqv ρ x₁ x₂ && Array.isEqv alts₁ alts₂ (fun alt₁ alt₂ => + | ρ, FnBody.setTag x₁ i₁ b₁, FnBody.setTag x₂ i₂ b₂ => + aeqv ρ x₁ x₂ && i₁ == i₂ && alphaEqv ρ b₁ b₂ + | ρ, FnBody.inc x₁ n₁ c₁ p₁ b₁, FnBody.inc x₂ n₂ c₂ p₂ b₂ => + aeqv ρ x₁ x₂ && n₁ == n₂ && c₁ == c₂ && p₁ == p₂ && alphaEqv ρ b₁ b₂ + | ρ, FnBody.dec x₁ n₁ c₁ p₁ b₁, FnBody.dec x₂ n₂ c₂ p₂ b₂ => + aeqv ρ x₁ x₂ && n₁ == n₂ && c₁ == c₂ && p₁ == p₂ && alphaEqv ρ b₁ b₂ + | ρ, FnBody.del x₁ b₁, FnBody.del x₂ b₂ => + aeqv ρ x₁ x₂ && alphaEqv ρ b₁ b₂ + | ρ, FnBody.case n₁ x₁ _ alts₁, FnBody.case n₂ x₂ _ alts₂ => + n₁ == n₂ && aeqv ρ x₁ x₂ && Array.isEqv alts₁ alts₂ (fun alt₁ alt₂ => match alt₁, alt₂ with | Alt.ctor i₁ b₁, Alt.ctor i₂ b₂ => i₁ == i₂ && alphaEqv ρ b₁ b₂ | Alt.default b₁, Alt.default b₂ => alphaEqv ρ b₁ b₂ | _, _ => false) - | ρ, FnBody.jmp j₁ ys₁, FnBody.jmp j₂ ys₂ => j₁ == j₂ && aeqv ρ ys₁ ys₂ - | ρ, FnBody.ret x₁, FnBody.ret x₂ => aeqv ρ x₁ x₂ - | _, FnBody.unreachable, FnBody.unreachable => true - | _, _, _ => false + | ρ, FnBody.jmp j₁ ys₁, FnBody.jmp j₂ ys₂ => + j₁ == j₂ && aeqv ρ ys₁ ys₂ + | ρ, FnBody.ret x₁, FnBody.ret x₂ => + aeqv ρ x₁ x₂ + | _, FnBody.unreachable, FnBody.unreachable => + true + | _, _, _ => + false def FnBody.beq (b₁ b₂ : FnBody) : Bool := FnBody.alphaEqv ∅ b₁ b₂ diff --git a/src/Lean/Compiler/IR/Borrow.lean b/src/Lean/Compiler/IR/Borrow.lean index 87b4652c4a1b..b25ec90b483d 100644 --- a/src/Lean/Compiler/IR/Borrow.lean +++ b/src/Lean/Compiler/IR/Borrow.lean @@ -231,7 +231,7 @@ def ownArgsIfParam (xs : Array Arg) : M Unit := do | .var x => if ctx.paramSet.contains x.idx then ownVar x | .erased => pure () -def collectExpr (z : VarId) (e : Expr) : M Unit := do +def collectExpr (z : VarId) (t : IRType) (e : Expr) : M Unit := do match e with | .reset _ x => ownVar z @@ -243,9 +243,10 @@ def collectExpr (z : VarId) (e : Expr) : M Unit := do | .ctor _ xs => ownVar z ownArgsIfParam xs - | .proj _ x => + | .proj _ _ x | .unbox x => if (← isOwned x) then ownVar z - if (← isOwned z) then ownVar x + if (← isOwned z) then + unless t.isScalar do ownVar x | .fap g xs => let ps ← getParamInfo (.decl g) ownVar z @@ -257,7 +258,10 @@ def collectExpr (z : VarId) (e : Expr) : M Unit := do | .pap _ xs => ownVar z ownArgs xs - | _ => pure () + | .box _ x => + ownVar z + ownVar x + | .lit _ | .isShared _ | .sproj .. | .uproj .. => pure () def preserveTailCall (x : VarId) (v : Expr) (b : FnBody) : M Unit := do let ctx ← read @@ -279,9 +283,9 @@ partial def collectFnBody (b : FnBody) : M Unit := do let ctx ← read updateParamMap (.jp ctx.currFn j) collectFnBody b - | .vdecl x _ v b => + | .vdecl x t v b => collectFnBody b - collectExpr x v + collectExpr x t v preserveTailCall x v b | .jmp j ys => let ctx ← read diff --git a/src/Lean/Compiler/IR/Boxing.lean b/src/Lean/Compiler/IR/Boxing.lean index c8a1e4bd4796..a019b3fc3fb5 100644 --- a/src/Lean/Compiler/IR/Boxing.lean +++ b/src/Lean/Compiler/IR/Boxing.lean @@ -12,6 +12,7 @@ public import Lean.Compiler.IR.CompilerM public import Lean.Compiler.IR.ElimDeadVars public import Lean.Compiler.IR.ToIRType public import Lean.Data.AssocList +import Lean.Compiler.InitAttr public section @@ -33,8 +34,9 @@ private def N.mkFresh : N VarId := def requiresBoxedVersion (env : Environment) (decl : Decl) : Bool := let ps := decl.params - (ps.size > 0 && (decl.resultType.isScalar || ps.any (fun p => p.ty.isScalar || p.borrow || p.ty.isVoid) || isExtern env decl.name)) - || ps.size > closureMaxArgs + (ps.size > 0 && (decl.resultType.isScalarOrStruct || + ps.any (fun p => p.ty.isScalarOrStruct || p.borrow || p.ty.isVoid) || isExtern env decl.name)) + || ps.size > closureMaxArgs def mkBoxedVersionAux (decl : Decl) : N Decl := do let ps := decl.params @@ -42,14 +44,14 @@ def mkBoxedVersionAux (decl : Decl) : N Decl := do let (newVDecls, xs) ← qs.size.foldM (init := (#[], #[])) fun i _ (newVDecls, xs) => do let p := ps[i]! let q := qs[i] - if !p.ty.isScalar then + if !p.ty.isScalarOrStruct then pure (newVDecls, xs.push (.var q.x)) else let x ← N.mkFresh pure (newVDecls.push (.vdecl x p.ty (.unbox q.x) default), xs.push (.var x)) let r ← N.mkFresh let newVDecls := newVDecls.push (.vdecl r decl.resultType (.fap decl.name xs) default) - let body ← if !decl.resultType.isScalar then + let body ← if !decl.resultType.isScalarOrStruct then pure <| reshape newVDecls (.ret (.var r)) else let newR ← N.mkFresh @@ -60,14 +62,28 @@ def mkBoxedVersionAux (decl : Decl) : N Decl := do def mkBoxedVersion (decl : Decl) : Decl := (mkBoxedVersionAux decl).run' 1 +partial def transformBoxedPaps (decls : Array Decl) (env : Environment) (b : FnBody) : FnBody := + match b with + | .vdecl x ty (.pap f ys) b => + let decl := (findEnvDecl' env f decls).get! + let f := if requiresBoxedVersion env decl then mkBoxedName f else f + .vdecl x ty (.pap f ys) (transformBoxedPaps decls env b) + | .jdecl j xs v b => + .jdecl j xs (transformBoxedPaps decls env v) (transformBoxedPaps decls env b) + | .case t x xTy alts => + .case t x xTy (alts.map (Alt.modifyBody <| transformBoxedPaps decls env)) + | _ => if b.isTerminal then b else b.setBody (transformBoxedPaps decls env b.body) + def addBoxedVersions (env : Environment) (decls : Array Decl) : Array Decl := let boxedDecls := decls.foldl (init := #[]) fun newDecls decl => if requiresBoxedVersion env decl then newDecls.push (mkBoxedVersion decl) else newDecls + let decls := decls.map fun + | .fdecl f xs ty b i => + let b := transformBoxedPaps decls env b + .fdecl f xs ty b i + | d => d decls ++ boxedDecls -def eqvTypes (t₁ t₂ : IRType) : Bool := - (t₁.isScalar == t₂.isScalar) && (!t₁.isScalar || t₁ == t₂) - structure BoxingContext where f : FunId localCtx : LocalContext := {} @@ -130,7 +146,7 @@ def getDecl (fid : FunId) : M Decl := do withReader (fun ctx => { ctx with localCtx := ctx.localCtx.addJP j xs v }) k /-- If `x` declaration is of the form `x := Expr.lit _` or `x := Expr.fap c #[]`, - and `x`'s type is not cheap to box (e.g., it is `UInt64), then return its value. -/ + and `x`'s type is not cheap to box (e.g., it is `UInt64`), then return its value. -/ private def isExpensiveConstantValueBoxing (x : VarId) (xType : IRType) : M (Option Expr) := match xType with | .uint8 | .uint16 => return none @@ -149,11 +165,19 @@ private def isExpensiveConstantValueBoxing (x : VarId) (xType : IRType) : M (Opt It is used when the expected type does not match `xType`. If `xType` is scalar, then we need to "box" it. Otherwise, we need to "unbox" it. -/ def mkCast (x : VarId) (xType : IRType) (expectedType : IRType) : M Expr := do - if expectedType.isScalar then - return .unbox x + if expectedType.isScalarOrStruct then + if xType.isStruct then + -- Reshaping a struct is denoted by the .box instruction + return .box xType x + else + return .unbox x else match (← isExpensiveConstantValueBoxing x xType) with | some v => do + if let .fap nm #[] := v then + if expectedType.isObj && xType.isStruct then + -- we create boxed versions specifically for struct constants (see `boxedConstDecl`) + return .fap (mkBoxedName nm) #[] let ctx ← read let s ← get /- Create auxiliary FnBody @@ -170,7 +194,7 @@ def mkCast (x : VarId) (xType : IRType) (expectedType : IRType) : M Expr := do match s.auxDeclCache.find? body with | some v => pure v | none => do - let auxName := ctx.f ++ ((`_boxed_const).appendIndexAfter s.nextAuxId) + let auxName := ctx.f.str ("_boxed_const_" ++ s.nextAuxId.repr) let auxConst := .fap auxName #[] let auxDecl := Decl.fdecl auxName #[] expectedType body {} modify fun s => { s with @@ -183,7 +207,7 @@ def mkCast (x : VarId) (xType : IRType) (expectedType : IRType) : M Expr := do @[inline] def castVarIfNeeded (x : VarId) (expected : IRType) (k : VarId → M FnBody) : M FnBody := do let xType ← getVarType x - if eqvTypes xType expected then + if xType.compatibleWith expected then k x else let y ← M.mkFresh @@ -206,7 +230,7 @@ def castArgsIfNeededAux (xs : Array Arg) (typeFromIdx : Nat → IRType) : M (Arr xs' := xs'.push x | .var x => let xType ← getVarType x - if eqvTypes xType expected then + if xType.compatibleWith expected then xs' := xs'.push (.var x) else let y ← M.mkFresh @@ -228,14 +252,14 @@ def castArgsIfNeededAux (xs : Array Arg) (typeFromIdx : Nat → IRType) : M (Arr pure (reshape bs b) def unboxResultIfNeeded (x : VarId) (ty : IRType) (e : Expr) (b : FnBody) : M FnBody := do - if ty.isScalar then + if ty.isScalarOrStruct then let y ← M.mkFresh return .vdecl y .tobject e (.vdecl x ty (.unbox y) b) else return .vdecl x ty e b def castResultIfNeeded (x : VarId) (ty : IRType) (e : Expr) (eType : IRType) (b : FnBody) : M FnBody := do - if eqvTypes ty eType then + if ty.compatibleWith eType then return .vdecl x ty e b else let y ← M.mkFresh @@ -243,28 +267,35 @@ def castResultIfNeeded (x : VarId) (ty : IRType) (e : Expr) (eType : IRType) (b let v ← mkCast y boxedTy ty return .vdecl y boxedTy e (.vdecl x ty v b) -def visitVDeclExpr (x : VarId) (ty : IRType) (e : Expr) (b : FnBody) : M FnBody := +def visitCtorExpr (x : VarId) (ty : IRType) (c : CtorInfo) (ys : Array Arg) (b : FnBody) : + M FnBody := do + if c.isScalar && ty.isScalar then + return .vdecl x ty (.lit (.num c.cidx)) b + else if let .struct _ tys _ _ := ty then + let (ys, bs) ← castArgsIfNeededAux ys (fun i => tys[i]!) + return reshape bs (.vdecl x ty (.ctor c ys) b) + else if let .union _ tys := ty then + let .struct _ tys _ _ := tys[c.cidx]! | unreachable! + let (ys, bs) ← castArgsIfNeededAux ys (fun i => tys[i]!) + return reshape bs (.vdecl x ty (.ctor c ys) b) + else + boxArgsIfNeeded ys fun ys => return .vdecl x ty (.ctor c ys) b + +def visitVDeclExpr (x : VarId) (ty : IRType) (e : Expr) (b : FnBody) : M FnBody := do match e with - | .ctor c ys => - if c.isScalar && ty.isScalar then - return .vdecl x ty (.lit (.num c.cidx)) b - else - boxArgsIfNeeded ys fun ys => return .vdecl x ty (.ctor c ys) b + | .ctor c ys => visitCtorExpr x ty c ys b | .reuse w c u ys => boxArgsIfNeeded ys fun ys => return .vdecl x ty (.reuse w c u ys) b - | .fap f ys => do + | .fap f ys => let decl ← getDecl f castArgsIfNeeded ys decl.params fun ys => castResultIfNeeded x ty (.fap f ys) decl.resultType b - | .pap f ys => do - let env ← getEnv - let decl ← getDecl f - let f := if requiresBoxedVersion env decl then mkBoxedName f else f + | .pap f ys => boxArgsIfNeeded ys fun ys => return .vdecl x ty (.pap f ys) b | .ap f ys => boxArgsIfNeeded ys fun ys => unboxResultIfNeeded x ty (.ap f ys) b - | _ => + | _ => return .vdecl x ty e b /-- @@ -276,14 +307,23 @@ let y : obj := f x where `f : obj -> uint8`. It is the job of the boxing pass to enforce a stricter obj/scalar separation and as such it needs to correct situations like this. -/ -def tryCorrectVDeclType (ty : IRType) (e : Expr) : M IRType := +def tryCorrectVDeclType (ty : IRType) (e : Expr) : M IRType := do match e with - | .fap f _ => do + | .fap f _ => let decl ← getDecl f return decl.resultType | .pap .. => return .object | .uproj .. => return .usize - | .ctor .. | .reuse .. | .ap .. | .lit .. | .sproj .. | .proj .. | .reset .. => + | .proj c i x => + let type ← getVarType x + match type with + | .struct _ tys _ _ => + return tys[i]! + | .union _ tys => + let .struct _ tys _ _ := tys[c]! | unreachable! + return tys[i]! + | _ => return ty.boxed + | .ctor .. | .reuse .. | .ap .. | .lit .. | .sproj .. | .reset .. => return ty | .unbox .. | .box .. | .isShared .. => unreachable! @@ -296,16 +336,18 @@ partial def visitFnBody : FnBody → M FnBody let v ← withParams xs (visitFnBody v) let b ← withJDecl j xs v (visitFnBody b) return .jdecl j xs v b - | .uset x i y b => do + | .uset x c i y b => do let b ← visitFnBody b castVarIfNeeded y IRType.usize fun y => - return .uset x i y b - | .sset x i o y ty b => do + return .uset x c i y b + | .sset x c i o y ty b => do let b ← visitFnBody b castVarIfNeeded y ty fun y => - return .sset x i o y ty b + return .sset x c i o y ty b | .case tid x xType alts => do let alts ← alts.mapM fun alt => alt.modifyBodyM visitFnBody + if xType.isObj then + return .case tid x (← getVarType x) alts castVarIfNeeded x xType fun x => do return .case tid x xType alts | .ret x => do @@ -317,25 +359,156 @@ partial def visitFnBody : FnBody → M FnBody | other => pure other +namespace BoxParams + +def isExpensiveScalar (ty : IRType) : Bool := + ty.isScalarOrStruct && !ty matches .uint8 | .uint16 + +abbrev M := ReaderT BoxingContext (StateM VarIdSet) + +def checkArgsBoxed (args : Array Arg) (paramsBoxed : Nat → Bool) : VarIdSet → VarIdSet := + go 0 +where + go (i : Nat) (set : VarIdSet) : VarIdSet := + if h : i < args.size then + if let .var v := args[i] then + if paramsBoxed i then + go (i + 1) (set.insert v) + else + go (i + 1) set + else + go (i + 1) set + else set + termination_by args.size - i + +def visitVDeclExpr (x : VarId) (ty : IRType) (e : Expr) (b : FnBody) : M FnBody := do + match e with + | .ctor c ys => + if ty.isObj then + modify (checkArgsBoxed ys fun _ => true) + else if (← get).contains x then + modify (checkArgsBoxed ys fun _ => true) + return .vdecl x ty.boxed e b + else if let .struct _ tys _ _ := ty then + modify (checkArgsBoxed ys fun i => !tys[i]!.isScalarOrStruct) + else if let .union _ tys := ty then + let .struct _ tys _ _ := tys[c.cidx]! | unreachable! + modify (checkArgsBoxed ys fun i => !tys[i]!.isScalarOrStruct) + return .vdecl x ty e b + | .reuse _ _ _ ys => + modify (checkArgsBoxed ys fun _ => true) + return .vdecl x ty e b + | .fap f ys => do + let ctx ← read + let decl := (findEnvDecl' ctx.env f ctx.decls).get! + modify (checkArgsBoxed ys fun i => !decl.params[i]!.ty.isScalarOrStruct) + return .vdecl x ty e b + | .pap _ ys => do + modify (checkArgsBoxed ys fun _ => true) + return .vdecl x ty e b + | .ap _ ys => + modify (checkArgsBoxed ys fun _ => true) + return .vdecl x ty e b + | _ => return .vdecl x ty e b + +partial def boxParams : FnBody → M FnBody + | .vdecl x t v b => do + let b ← boxParams b + visitVDeclExpr x t v b + | .jdecl j xs v b => do + let v ← boxParams v + let boxed ← get + let xs := xs.mapMono fun param => + if isExpensiveScalar param.ty ∧ boxed.contains param.x then + { param with ty := param.ty.boxed } + else + param + withReader (fun ctx => { ctx with localCtx := ctx.localCtx.addJP j xs v }) do + return .jdecl j xs v (← boxParams b) + | .case tid x xType alts => do + let alts ← alts.mapM fun alt => alt.modifyBodyM boxParams + return .case tid x xType alts + | .ret x => do + if let .var v := x then + let expected := (← read).resultType + if !expected.isScalarOrStruct then + modify (·.insert v) + return .ret x + | .jmp j ys => do + let ps := ((← read).localCtx.getJPParams j).get! + modify (checkArgsBoxed ys fun i => !ps[i]!.ty.isScalarOrStruct) + return .jmp j ys + | .uset x c i y b => do + return .uset x c i y (← boxParams b) + | .sset x c i o y ty b => do + return .sset x c i o y ty (← boxParams b) + | b => pure b + +end BoxParams + +def boxedConstDecl (f : FunId) (ty : IRType) : Decl := + /- + def f._boxed : tobj := + let x_1 : ty := f; + let x_2 : tobj := box x_1; + ret x_2 + -/ + .fdecl (mkBoxedName f) #[] ty.boxed (info := {}) <| + .vdecl { idx := 1 } ty (.fap f #[]) <| + .vdecl { idx := 2 } ty.boxed (.box ty { idx := 1 }) <| + .ret (.var { idx := 2 }) + def run (env : Environment) (decls : Array Decl) : Array Decl := - let decls := decls.foldl (init := #[]) fun newDecls decl => + decls.foldl (init := #[]) fun newDecls decl => match decl with | .fdecl f xs resultType b _ => let nextIdx := decl.maxIndex + 1 let (b, s) := withParams xs (visitFnBody b) { f, resultType, decls, env } |>.run { nextIdx } let newDecls := newDecls ++ s.auxDecls + let newDecls := + if resultType.isStruct && xs.isEmpty then + newDecls.push (boxedConstDecl f resultType) + else + newDecls let newDecl := decl.updateBody! b let newDecl := newDecl.elimDead newDecls.push newDecl | d => newDecls.push d - addBoxedVersions env decls end ExplicitBoxing +open ExplicitBoxing BoxParams in +def prepareBoxParams (env : Environment) (decls : Array Decl) : Array Decl := + decls.map fun decl => + match decl with + | .fdecl f xs resultType b info => + let exported := isExport env f + let resultType := if exported && resultType.isStruct then resultType.boxed else resultType + let (b, boxed) := boxParams b { f, resultType, decls, env } + |>.run {} + if isExport env f then + -- exports shouldn't use unboxed structs + let xs := xs.map (fun x => { x with ty := if x.ty.isStruct then x.ty.boxed else x.ty }) + .fdecl f xs resultType b info + else + let xs := xs.mapMono fun param => + if isExpensiveScalar param.ty ∧ boxed.contains param.x then + { param with ty := param.ty.boxed } + else + param + .fdecl f xs resultType b info + | .extern f xs resultType e => + -- extern declarations should also not use unboxed structs + let xs := xs.map (fun x => { x with ty := if x.ty.isStruct then x.ty.boxed else x.ty }) + let resultType := if resultType.isStruct then resultType.boxed else resultType + .extern f xs resultType e + def explicitBoxing (decls : Array Decl) : CompilerM (Array Decl) := do let env ← getEnv return ExplicitBoxing.run env decls +builtin_initialize registerTraceClass `compiler.ir.box_params (inherited := true) builtin_initialize registerTraceClass `compiler.ir.boxing (inherited := true) +builtin_initialize registerTraceClass `compiler.ir.boxed_versions (inherited := true) end Lean.IR diff --git a/src/Lean/Compiler/IR/Checker.lean b/src/Lean/Compiler/IR/Checker.lean index 6bfe759f4c1a..9f181bf68dc2 100644 --- a/src/Lean/Compiler/IR/Checker.lean +++ b/src/Lean/Compiler/IR/Checker.lean @@ -80,7 +80,7 @@ def checkArgs (as : Array Arg) : M Unit := @[inline] def checkEqTypes (ty₁ ty₂ : IRType) : M Unit := do unless ty₁ == ty₂ do - throwCheckerError "unexpected type '{ty₁}' != '{ty₂}'" + throwCheckerError s!"unexpected type '{ty₁}' != '{ty₂}'" @[inline] def checkType (ty : IRType) (p : IRType → Bool) (suffix? : Option String := none): M Unit := do unless p ty do @@ -91,8 +91,14 @@ def checkArgs (as : Array Arg) : M Unit := def checkObjType (ty : IRType) : M Unit := checkType ty IRType.isObj "object expected" +def checkObjOrStructType (ty : IRType) : M Unit := + checkType ty IRType.isObjOrStruct "object or struct expected" + def checkScalarType (ty : IRType) : M Unit := checkType ty IRType.isScalar "scalar expected" +def checkScalarOrStructType (ty : IRType) : M Unit := + checkType ty IRType.isScalarOrStruct "scalar expected" + def getType (x : VarId) : M IRType := do let ctx ← read match ctx.localCtx.getType x with @@ -105,8 +111,11 @@ def getType (x : VarId) : M IRType := do def checkObjVar (x : VarId) : M Unit := checkVarType x IRType.isObj "object expected" -def checkScalarVar (x : VarId) : M Unit := - checkVarType x IRType.isScalar "scalar expected" +def checkObjOrStructVar (x : VarId) : M Unit := + checkVarType x IRType.isObjOrStruct "object or struct expected" + +def checkScalarOrStructVar (x : VarId) : M Unit := + checkVarType x IRType.isScalarOrStruct "scalar expected" def checkFullApp (c : FunId) (ys : Array Arg) : M Unit := do let decl ← getDecl c @@ -141,7 +150,7 @@ def checkExpr (ty : IRType) (e : Expr) : M Unit := do if c.ssize + c.usize * usizeSize > maxCtorScalarsSize then throwCheckerError s!"constructor '{c.name}' has too many scalar fields" if c.isRef then - checkObjType ty + checkObjOrStructType ty checkArgs ys | .reset _ x => checkObjVar x @@ -151,13 +160,13 @@ def checkExpr (ty : IRType) (e : Expr) : M Unit := do checkArgs ys checkObjType ty | .box xty x => - checkObjType ty - checkScalarVar x + checkObjOrStructType ty + checkScalarOrStructVar x checkVarType x (· == xty) | .unbox x => - checkScalarType ty + checkScalarOrStructType ty checkObjVar x - | .proj i x => + | .proj c i x => let xType ← getType x; /- Projections are a valid operation on `tobject`. Thus they should also @@ -166,19 +175,28 @@ def checkExpr (ty : IRType) (e : Expr) : M Unit := do -/ match xType with | .object | .tobject => - checkObjType ty - | .struct _ tys | .union _ tys => - if h : i < tys.size then - checkEqTypes (tys[i]) ty - else + checkObjOrStructType ty + | .struct _ tys _ _ => + if c ≠ 0 then + throwCheckerError "expected constructor index 0 for struct" + if h : i ≥ tys.size then throwCheckerError "invalid proj index" + --else checkEqTypes (tys[i]) ty + | .union _ tys => + if h : c < tys.size then + let .struct _ tys _ _ := tys[c] | throwCheckerError "expected struct inside union" + if h : i ≥ tys.size then + throwCheckerError "invalid proj index" + --else checkEqTypes (tys[i]) ty + else + throwCheckerError "invalid proj constructor index" | .tagged => pure () | _ => throwCheckerError s!"unexpected IR type '{xType}'" - | .uproj _ x => - checkObjVar x + | .uproj _ _ x => + checkObjOrStructVar x checkType ty (· == .usize) - | .sproj _ _ x => - checkObjVar x + | .sproj _ _ _ x => + checkObjOrStructVar x checkScalarType ty | .isShared x => checkObjVar x @@ -208,11 +226,11 @@ partial def checkFnBody (fnBody : FnBody) : M Unit := do checkVar x checkArg y checkFnBody b - | .uset x _ y b => + | .uset x _ _ y b => checkVar x checkVar y checkFnBody b - | .sset x _ _ y _ b => + | .sset x _ _ _ y _ b => checkVar x checkVar y checkFnBody b diff --git a/src/Lean/Compiler/IR/ElimDeadVars.lean b/src/Lean/Compiler/IR/ElimDeadVars.lean index de077004c0f8..6bebb1c57ac0 100644 --- a/src/Lean/Compiler/IR/ElimDeadVars.lean +++ b/src/Lean/Compiler/IR/ElimDeadVars.lean @@ -46,7 +46,6 @@ partial def reshapeWithoutDead (bs : Array FnBody) (term : FnBody) : FnBody := reshape bs b used match curr with | FnBody.vdecl x _ e _ => keepIfUsedLet x.idx e - -- TODO: we should keep all struct/union projections because they are used to ensure struct/union values are fully consumed. | FnBody.jdecl j _ _ _ => keepIfUsedJp j.idx | _ => keep () reshape bs term term.freeIndices diff --git a/src/Lean/Compiler/IR/EmitC.lean b/src/Lean/Compiler/IR/EmitC.lean index 97f9d9738d31..80522e407de2 100644 --- a/src/Lean/Compiler/IR/EmitC.lean +++ b/src/Lean/Compiler/IR/EmitC.lean @@ -11,6 +11,7 @@ public import Lean.Compiler.IR.EmitUtil public import Lean.Compiler.IR.NormIds public import Lean.Compiler.IR.SimpCase public import Lean.Compiler.IR.Boxing +public import Lean.Compiler.IR.StructRC public import Lean.Compiler.ModPkgExt public section @@ -26,6 +27,8 @@ structure Context where jpMap : JPParamsMap := {} mainFn : FunId := default mainParams : Array Param := #[] + varTypes : VarTypeMap := {} + structs : StructTypeLookup := {} abbrev M := ReaderT Context (EStateM String String) @@ -60,21 +63,28 @@ def argToCString (x : Arg) : String := def emitArg (x : Arg) : M Unit := emit (argToCString x) -def toCType : IRType → String - | IRType.float => "double" - | IRType.float32 => "float" - | IRType.uint8 => "uint8_t" - | IRType.uint16 => "uint16_t" - | IRType.uint32 => "uint32_t" - | IRType.uint64 => "uint64_t" - | IRType.usize => "size_t" - | IRType.object => "lean_object*" - | IRType.tagged => "lean_object*" - | IRType.tobject => "lean_object*" - | IRType.erased => "lean_object*" - | IRType.void => "lean_object*" - | IRType.struct _ _ => panic! "not implemented yet" - | IRType.union _ _ => panic! "not implemented yet" +@[inline] +def lookupStruct (ty : IRType) : M Nat := do + return (← read).structs[IRTypeApprox.mk ty]! + +def structType (id : Nat) : String := + "struct l_s" ++ id.repr + +def toCType : IRType → M String + | .float => return "double" + | .float32 => return "float" + | .uint8 => return "uint8_t" + | .uint16 => return "uint16_t" + | .uint32 => return "uint32_t" + | .uint64 => return "uint64_t" + | .usize => return "size_t" + | .object => return "lean_object*" + | .tagged => return "lean_object*" + | .tobject => return "lean_object*" + | .erased => return "lean_object*" + | .void => return "lean_object*" + | ty@(.struct ..) | ty@(.union ..) => + return structType (← lookupStruct ty) def throwInvalidExportName {α : Type} (n : Name) : M α := throw s!"invalid export name '{n}'" @@ -101,6 +111,83 @@ def toCInitName (n : Name) : M String := do def emitCInitName (n : Name) : M Unit := toCInitName n >>= emit +def emitSpreadArg (ty : IRType) (name : Option String) (first : Bool) : M Bool := do + if let .struct _ tys 0 0 := ty then + let mut first := first + for h : i in *...tys.size do + let ty := tys[i] + if ty matches .erased then + continue + first ← emitSpreadArg ty (name.map fun nm => nm ++ "_" ++ i.repr) first + return first + if ty matches .void then + return first + unless first do + emit ", " + emit (← toCType ty) + if let some nm := name then + emit " "; emit nm + return false + +def emitSpreadArgs (ps : Array Param) (emitNames : Bool) : M Unit := do + let mut first := true + for p in ps do + first ← emitSpreadArg p.ty (if emitNames then some (toString p.x) else none) first + +def emitSpreadValue (ty : IRType) (name : String) : M Unit := do + if let .struct _ tys 0 0 := ty then + emit "{" + let mut first := true + for h : i in *...tys.size do + let ty := tys[i] + if ty matches .erased | .void then + continue + unless first do + emit ", " + emit ".i"; emit i; emit " = " + emitSpreadValue ty (name ++ "_" ++ i.repr) + first := false + emit "}" + return + emit name + +def emitSpreads (ps : Array Param) : M Unit := do + for p in ps do + if let .struct _ _ 0 0 := p.ty then + emit (← toCType p.ty); emit " "; emit p.x; emit " = ("; emit (← toCType p.ty); emit ")" + emitSpreadValue p.ty (toString p.x) + emitLn ";" + +def emitFullAppArg (ty : IRType) (nm : String) (first : Bool) : M Bool := do + if let .struct _ tys 0 0 := ty then + let mut first := first + for h : i in *...tys.size do + let ty := tys[i] + if ty matches .erased | .void then + continue + first ← emitFullAppArg ty (nm ++ ".i" ++ i.repr) first + return first + if ty matches .void then + return first + unless first do emit ", " + emit nm + return false + +def emitFullAppArgs (ps : Array Param) (args : Array Arg) : M Unit := do + let mut first := true + for h : i in *...args.size do + let ty := ps[i]!.ty + let arg := args[i] + if ty matches .void then + continue + match arg with + | .erased => + unless first do emit ", " + emit "lean_box(0)" + first := false + | .var v => + first ← emitFullAppArg ty (toString v) first + def emitFnDeclAux (decl : Decl) (cppBaseName : String) (isExternal : Bool) : M Unit := do let ps := decl.params let env ← getEnv @@ -110,7 +197,7 @@ def emitFnDeclAux (decl : Decl) (cppBaseName : String) (isExternal : Bool) : M U else emit "LEAN_EXPORT " else if !isExternal then emit "LEAN_EXPORT " - emit (toCType decl.resultType ++ " " ++ cppBaseName) + emit ((← toCType decl.resultType) ++ " " ++ cppBaseName) unless ps.isEmpty do emit "(" -- We omit void parameters, note that they are guaranteed not to occur in boxed functions @@ -120,9 +207,7 @@ def emitFnDeclAux (decl : Decl) (cppBaseName : String) (isExternal : Bool) : M U if ps.size > closureMaxArgs && isBoxedName decl.name then emit "lean_object**" else - ps.size.forM fun i _ => do - if i > 0 then emit ", " - emit (toCType ps[i].ty) + emitSpreadArgs ps false emit ")" emitLn ";" @@ -263,7 +348,7 @@ def getJPParams (j : JoinPointId) : M (Array Param) := do | none => throw "unknown join point" def declareVar (x : VarId) (t : IRType) : M Unit := do - emit (toCType t); emit " "; emit x; emit "; " + emit (← toCType t); emit " "; emit x; emit "; " def declareParams (ps : Array Param) : M Unit := ps.forM fun p => declareVar p.x p.ty @@ -278,9 +363,32 @@ partial def declareVars : FnBody → Bool → M Bool | FnBody.jdecl _ xs _ b, d => do declareParams xs; declareVars b (d || xs.size > 0) | e, d => if e.isTerminal then pure d else declareVars e.body d +partial def optionLikePath (ty : IRType) : Option (List Nat) := Id.run do + match ty with + | .struct _ tys _ _ => + for h : i in *...tys.size do + if let some l := optionLikePath tys[i] then + return i :: l + return none + | .tagged | .object | .tobject => return some [] + | _ => return none + +def optionLike? (ty : IRType) : Option (List Nat) := + match ty with + | .union _ #[.struct _ #[] 0 0, ty] => optionLikePath ty + | _ => none + +def emitPath (path : List Nat) : M Unit := do + emit ".cs.c1"; path.forM (emit s!".i{·}") + def emitTag (x : VarId) (xType : IRType) : M Unit := do if xType.isObj then do emit "lean_obj_tag("; emit x; emit ")" + else if xType.isStruct then + if let some a := optionLike? xType then + emit x; emitPath a; emit " != 0" + else + emit x; emit ".tag" else emit x @@ -290,19 +398,54 @@ def isIf (alts : Array Alt) : Option (Nat × FnBody × FnBody) := | Alt.ctor c b => some (c.cidx, b, alts[1].body) | _ => none -def emitInc (x : VarId) (n : Nat) (checkRef : Bool) : M Unit := do - emit $ - if checkRef then (if n == 1 then "lean_inc" else "lean_inc_n") - else (if n == 1 then "lean_inc_ref" else "lean_inc_ref_n") - emit "("; emit x - if n != 1 then emit ", "; emit n - emitLn ");" +partial def needsRC (ty : IRType) : Bool := + match ty with + | .object | .tobject => true + | .union _ tys => tys.any needsRC + | .struct _ tys _ _ => tys.any needsRC + | _ => false + +def structIncFnPrefix := "_l_struct_inc_" +def structDecFnPrefix := "_l_struct_dec_" +def structReshapeFnPrefix := "_l_struct_reshape_" +def structBoxFnPrefix := "_l_struct_box_" +def structUnboxFnPrefix := "_l_struct_unbox_" + +def emitIncOfType (x : String) (ty : IRType) (n : Nat) (checkRef : Bool) (nstr : String) : + M Unit := do + if ty.isStruct then + unless needsRC ty do + return + let id ← lookupStruct ty + emit structIncFnPrefix; emit id; emit "(" + emit x; emit ", "; emit nstr + emitLn ");" + else + emit $ + if checkRef then (if n == 1 then "lean_inc" else "lean_inc_n") + else (if n == 1 then "lean_inc_ref" else "lean_inc_ref_n") + emit "("; emit x + if n != 1 then emit ", "; emit nstr + emitLn ");" + +def emitInc (x : VarId) (n : Nat) (checkRef : Bool) : M Unit := do + let ty := (← read).varTypes[x]! + emitIncOfType (toString x) ty n checkRef n.repr + +def emitDecOfType (x : String) (ty : IRType) (n : Nat) (checkRef : Bool) : M Unit := do + if ty.isStruct then + unless needsRC ty do + return + let id ← lookupStruct ty + emit structDecFnPrefix; emit id; emit "("; emit x; emitLn ");" + else + n.forM fun _ _ => do + emit (if checkRef then "lean_dec" else "lean_dec_ref"); + emit "("; emit x; emitLn ");" def emitDec (x : VarId) (n : Nat) (checkRef : Bool) : M Unit := do - emit (if checkRef then "lean_dec" else "lean_dec_ref"); - emit "("; emit x; - if n != 1 then emit ", "; emit n - emitLn ");" + let ty := (← read).varTypes[x]! + emitDecOfType (toString x) ty n checkRef def emitDel (x : VarId) : M Unit := do emit "lean_free_object("; emit x; emitLn ");" @@ -320,10 +463,27 @@ def emitOffset (n : Nat) (offset : Nat) : M Unit := do else emit offset -def emitUSet (x : VarId) (n : Nat) (y : VarId) : M Unit := do - emit "lean_ctor_set_usize("; emit x; emit ", "; emit n; emit ", "; emit y; emitLn ");" -def emitSSet (x : VarId) (n : Nat) (offset : Nat) (y : VarId) (t : IRType) : M Unit := do +def emitUSet (x : VarId) (cidx : Nat) (n : Nat) (y : VarId) : M Unit := do + let ty := (← read).varTypes[x]! + if let .union _ tys := ty then + let .struct _ tys _ _ := tys[cidx]! | unreachable! + emit x; emit ".cs.c"; emit cidx; emit ".u["; emit (n - tys.size); emit "] = "; emit y; emitLn ";" + else if let .struct _ tys _ _ := ty then + emit x; emit ".u["; emit (n - tys.size); emit "] = "; emit y; emitLn ";" + else + emit "lean_ctor_set_usize("; emit x; emit ", "; emit n; emit ", "; emit y; emitLn ");" + +def emitSSet (x : VarId) (cidx : Nat) (n : Nat) (offset : Nat) (y : VarId) (t : IRType) : M Unit := do + let ty := (← read).varTypes[x]! + if ty matches .union .. then + emit "*(("; emit (← toCType t); emit "*)(" + emit x; emit ".cs.c"; emit cidx; emit ".s+"; emit offset; emit ")) = "; emit y; emitLn ";" + return + else if ty matches .struct .. then + emit "*(("; emit (← toCType t); emit "*)(" + emit x; emit ".s+"; emit offset; emit ")) = "; emit y; emitLn ";" + return match t with | IRType.float => emit "lean_ctor_set_float" | IRType.float32 => emit "lean_ctor_set_float32" @@ -366,8 +526,28 @@ def emitCtorSetArgs (z : VarId) (ys : Array Arg) : M Unit := ys.size.forM fun i _ => do emit "lean_ctor_set("; emit z; emit ", "; emit i; emit ", "; emitArg ys[i]; emitLn ");" -def emitCtor (z : VarId) (c : CtorInfo) (ys : Array Arg) : M Unit := do - emitLhs z; +def emitCtor (z : VarId) (t : IRType) (c : CtorInfo) (ys : Array Arg) : M Unit := do + if let .union _ tys := t then + if let some path := optionLike? t then + if c.cidx = 0 then + emit z; emitPath path; emitLn " = 0;" + return + else + emit z; emit ".tag = "; emit c.cidx; emitLn ";" + let .struct _ tys _ _ := tys[c.cidx]! | unreachable! + for h : i in *...ys.size do + if tys[i]! matches .erased | .void then + continue + emit z; emit ".cs.c"; emit c.cidx; emit ".i"; emit i + emit " = "; emitArg ys[i]; emitLn ";" + return + else if let .struct _ tys _ _ := t then + for h : i in *...ys.size do + if tys[i]! matches .erased | .void then + continue + emit z; emit ".i"; emit i; emit " = "; emitArg ys[i]; emitLn ";" + return + emitLhs z if c.size == 0 && c.usize == 0 && c.ssize == 0 then do emit "lean_box("; emit c.cidx; emitLn ");" else do @@ -392,22 +572,46 @@ def emitReuse (z : VarId) (x : VarId) (c : CtorInfo) (updtHeader : Bool) (ys : A emitLn "}"; emitCtorSetArgs z ys -def emitProj (z : VarId) (i : Nat) (x : VarId) : M Unit := do - emitLhs z; emit "lean_ctor_get("; emit x; emit ", "; emit i; emitLn ");" +def emitProj (z : VarId) (c : Nat) (i : Nat) (x : VarId) : M Unit := do + emitLhs z + let ty := (← read).varTypes[x]! + if ty matches .struct .. then + emit x; emit ".i"; emit i; emitLn ";" + else if ty matches .union .. then + emit x; emit ".cs.c"; emit c; emit ".i"; emit i; emitLn ";" + else + emit "lean_ctor_get("; emit x; emit ", "; emit i; emitLn ");" -def emitUProj (z : VarId) (i : Nat) (x : VarId) : M Unit := do - emitLhs z; emit "lean_ctor_get_usize("; emit x; emit ", "; emit i; emitLn ");" +def emitUProj (z : VarId) (cidx : Nat) (i : Nat) (x : VarId) : M Unit := do + emitLhs z + let ty := (← read).varTypes[x]! + if let .union _ tys := ty then + let .struct _ tys _ _ := tys[cidx]! | unreachable! + emit x; emit ".cs.c"; emit cidx; emit ".u["; emit (i - tys.size); emitLn "];" + else if let .struct _ tys _ _ := ty then + emit x; emit ".u["; emit (i - tys.size); emitLn "];" + else + emit "lean_ctor_get_usize("; emit x; emit ", "; emit i; emitLn ");" -def emitSProj (z : VarId) (t : IRType) (n offset : Nat) (x : VarId) : M Unit := do - emitLhs z; +def emitSProj (z : VarId) (t : IRType) (cidx : Nat) (n offset : Nat) (x : VarId) : M Unit := do + emitLhs z + let ty := (← read).varTypes[x]! + if ty matches .union .. then + emit "*(("; emit (← toCType t); emit "*)(" + emit x; emit ".cs.c"; emit cidx; emit ".s+"; emit offset; emitLn "));" + return + else if ty matches .struct .. then + emit "*(("; emit (← toCType t); emit "*)(" + emit x; emit ".s+"; emit offset; emitLn "));" + return match t with - | IRType.float => emit "lean_ctor_get_float" - | IRType.float32 => emit "lean_ctor_get_float32" - | IRType.uint8 => emit "lean_ctor_get_uint8" - | IRType.uint16 => emit "lean_ctor_get_uint16" - | IRType.uint32 => emit "lean_ctor_get_uint32" - | IRType.uint64 => emit "lean_ctor_get_uint64" - | _ => throw "invalid instruction" + | IRType.float => emit "lean_ctor_get_float" + | IRType.float32 => emit "lean_ctor_get_float32" + | IRType.uint8 => emit "lean_ctor_get_uint8" + | IRType.uint16 => emit "lean_ctor_get_uint16" + | IRType.uint32 => emit "lean_ctor_get_uint32" + | IRType.uint64 => emit "lean_ctor_get_uint64" + | _ => throw "invalid instruction" emit "("; emit x; emit ", "; emitOffset n offset; emitLn ");" def toStringArgs (ys : Array Arg) : List String := @@ -442,8 +646,7 @@ def emitFullApp (z : VarId) (f : FunId) (ys : Array Arg) : M Unit := do | .fdecl (xs := ps) .. | .extern (xs := ps) (ext := { entries := [.opaque], .. }) .. => emitCName f if ys.size > 0 then - let (ys, _) := ys.zip ps |>.filter (fun (_, p) => !p.ty.isVoid) |>.unzip - emit "("; emitArgs ys; emit ")" + emit "("; emitFullAppArgs ps ys; emit ")" emitLn ";" | Decl.extern _ ps _ extData => emitExternCall f ps extData ys @@ -462,22 +665,37 @@ def emitApp (z : VarId) (f : VarId) (ys : Array Arg) : M Unit := else do emitLhs z; emit "lean_apply_"; emit ys.size; emit "("; emit f; emit ", "; emitArgs ys; emitLn ");" -def emitBoxFn (xType : IRType) : M Unit := +def emitBoxFn (xType tgt : IRType) : M Unit := do match xType with - | IRType.usize => emit "lean_box_usize" - | IRType.uint32 => emit "lean_box_uint32" - | IRType.uint64 => emit "lean_box_uint64" - | IRType.float => emit "lean_box_float" - | IRType.float32 => emit "lean_box_float32" - | _ => emit "lean_box" + | .usize => emit "lean_box_usize" + | .uint32 => emit "lean_box_uint32" + | .uint64 => emit "lean_box_uint64" + | .float => emit "lean_box_float" + | .float32 => emit "lean_box_float32" + | ty@(.struct ..) | ty@(.union ..) => + let id ← lookupStruct ty + if tgt.isStruct then + emit structReshapeFnPrefix + emit id + emit "_" + emit (← lookupStruct tgt) + else + emit structBoxFnPrefix + emit id + | _ => emit "lean_box" -def emitBox (z : VarId) (x : VarId) (xType : IRType) : M Unit := do - emitLhs z; emitBoxFn xType; emit "("; emit x; emitLn ");" +def emitBox (z : VarId) (x : VarId) (xType tgt : IRType) : M Unit := do + emitLhs z; emitBoxFn xType tgt; emit "("; emit x; emitLn ");" + +def emitUnboxFn (t : IRType) : M Unit := do + if t.isStruct then + emit structUnboxFnPrefix + emit (← lookupStruct t) + else + emit (getUnboxOpName t) def emitUnbox (z : VarId) (t : IRType) (x : VarId) : M Unit := do - emitLhs z - emit (getUnboxOpName t) - emit "("; emit x; emitLn ");" + emitLhs z; emitUnboxFn t; emit "("; emit x; emitLn ");" def emitIsShared (z : VarId) (x : VarId) : M Unit := do emitLhs z; emit "!lean_is_exclusive("; emit x; emitLn ");" @@ -531,16 +749,16 @@ def emitLit (z : VarId) (t : IRType) (v : LitVal) : M Unit := do def emitVDecl (z : VarId) (t : IRType) (v : Expr) : M Unit := match v with - | Expr.ctor c ys => emitCtor z c ys + | Expr.ctor c ys => emitCtor z t c ys | Expr.reset n x => emitReset z n x | Expr.reuse x c u ys => emitReuse z x c u ys - | Expr.proj i x => emitProj z i x - | Expr.uproj i x => emitUProj z i x - | Expr.sproj n o x => emitSProj z t n o x + | Expr.proj c i x => emitProj z c i x + | Expr.uproj c i x => emitUProj z c i x + | Expr.sproj c n o x => emitSProj z t c n o x | Expr.fap c ys => emitFullApp z c ys | Expr.pap c ys => emitPartialApp z c ys | Expr.ap x ys => emitApp z x ys - | Expr.box t x => emitBox z x t + | Expr.box t' x => emitBox z x t' t | Expr.unbox x => emitUnbox z t x | Expr.isShared x => emitIsShared z x | Expr.lit v => emitLit z t v @@ -590,7 +808,7 @@ def emitTailCall (v : Expr) : M Unit := let p := ps[i] let y := ys[i]! unless paramEqArg p y do - emit (toCType p.ty); emit " _tmp_"; emit i; emit " = "; emitArg y; emitLn ";" + emit (← toCType p.ty); emit " _tmp_"; emit i; emit " = "; emitArg y; emitLn ";" ps.size.forM fun i _ => do let p := ps[i] let y := ys[i]! @@ -645,8 +863,8 @@ partial def emitBlock (b : FnBody) : M Unit := do | FnBody.del x b => emitDel x; emitBlock b | FnBody.setTag x i b => emitSetTag x i; emitBlock b | FnBody.set x i y b => emitSet x i y; emitBlock b - | FnBody.uset x i y b => emitUSet x i y; emitBlock b - | FnBody.sset x i o y t b => emitSSet x i o y t; emitBlock b + | FnBody.uset x c i y b => emitUSet x c i y; emitBlock b + | FnBody.sset x c i o y t b => emitSSet x c i o y t; emitBlock b | FnBody.ret x => emit "return "; emitArg x; emitLn ";" | FnBody.case _ x xType alts => emitCase x xType alts | FnBody.jmp j xs => emitJmp j xs @@ -666,10 +884,290 @@ partial def emitFnBody (b : FnBody) : M Unit := do end +partial def emitStructDefn (ty : IRType) (nm : String) : M Unit := do + emitLn (nm ++ " {") + match ty with + | .union _ tys => + emitLn "union {" + for h : i in *...tys.size do + emit (← toCType tys[i]); emit " c"; emit i; emitLn ";" + assert! tys.size ≤ 256 + emitLn "} cs;" + if (optionLike? ty).isNone then + emitLn "uint8_t tag;" + | .struct _ tys us ss => + for h : i in *...tys.size do + if tys[i] matches .erased | .void then + continue + emit (← toCType tys[i]); emit " i"; emit i; emitLn ";" + if us ≠ 0 ∨ ss ≠ 0 then + -- Note: we keep a `size_t[0]` to ensure alignment of the scalars + emitLn s!"size_t u[{us}];" + emitLn s!"uint8_t s[{ss}];" + | _ => unreachable! + emitLn "};" + +def emitUnionSwitch (n : Nat) (x : String) (branch : (i : Nat) → i < n → M Unit) : M Unit := do + if h : n = 1 then + branch 0 (by simp_all) + else if h : n = 2 then + emit "if ("; emit x; emitLn " == 0) {" + branch 0 (by simp_all) + emitLn "} else {" + branch 1 (by simp_all) + emitLn "}" + else + emit "switch ("; emit x; emitLn ") {" + for h : i in *...n do + emit "case "; emit i; emitLn ":" + branch i (by get_elem_tactic) + emitLn "break;" + emitLn "}" + +def emitUnionSwitchWithImpossible (n : Nat) (x : String) + (possible : Nat → Bool) + (branch : (i : Nat) → i < n → M Unit) : M Unit := do + let branches : Array (Fin n) := (Array.ofFn id).filter fun i => possible i + if h : branches.size = 1 then + branch branches[0].1 branches[0].2 + else if h : branches.size = 2 then + emit "if ("; emit x; emitLn " == 0) {" + branch branches[0].1 branches[0].2 + emitLn "} else {" + branch branches[1].1 branches[1].2 + emitLn "}" + else + emit "switch ("; emit x; emitLn ") {" + for i in branches do + emit "case "; emit i; emitLn ":" + branch i.1 i.2 + emitLn "break;" + emitLn "}" + +def emitStructIncDecFn (ty : IRType) (id : Nat) (isInc : Bool) : M Unit := do + let prfx := if isInc then structIncFnPrefix else structDecFnPrefix + let call (tgt : String) (ty : IRType) := do + if ty matches .void | .erased || ty.isScalar then + return + if isInc then + emitIncOfType tgt ty 0 true "n" + else + emitDecOfType tgt ty 1 true + emit "static void "; emit prfx; emit id + emit "("; emit (structType id); + if isInc then emitLn " x, size_t n) {" else emitLn " x) {" + + if let .union _ tys := ty then + if let some path := optionLike? ty then + emit "if (x"; emitPath path; emitLn " != 0) {" + call s!"x.cs.c1" tys[1]! + emitLn "}" + else + emitUnionSwitch tys.size "x.tag" fun i _ => call s!"x.cs.c{i}" tys[i] + else if let .struct _ tys _ _ := ty then + for h : i in *...tys.size do + call s!"x.i{i}" tys[i] + emitLn "}" + +def emitReshapeFn (origin target : IRType) : M Unit := do + if origin.isObj then + emitUnboxFn target + else + emitBoxFn origin target + +partial def compatibleReshape (origin target : IRType) : Bool := + match origin, target with + | .struct _ tys _ _, .struct _ tys' _ _ => + tys.isEqv tys' compatibleReshape + | .union .., .struct .. => false + | .struct .., .union .. => false + | _, _ => (!origin.isScalar || !target.isStruct) && (!origin.isStruct || !target.isScalar) + +def emitStructReshapeFn (origin target : IRType) (id1 id2 : Nat) : M Unit := do + if id1 == id2 then + return + unless compatibleReshape origin target do + return + emit "static "; emit (structType id2); emit " " + emit structReshapeFnPrefix; emit id1; emit "_"; emit id2 + emit "("; emit (structType id1); emitLn " x) {" + emit (structType id2); emitLn " y;" + + match origin, target with + | .union _ tys, .union _ tys' => + let reshapeUnion (i : Nat) : M Unit := do + emit "y.cs.c"; emit i; emit " = " + if tys[i]!.compatibleWith tys'[i]! then + emit "x.cs.c"; emit i; emitLn ";" + else + emitReshapeFn tys[i]! tys'[i]! + emit "(x.cs.c"; emit i; emitLn ");" + if let some path := optionLike? origin then + let compatible := compatibleReshape tys[1]! tys'[1]! + if !compatible then + if let some path2 := optionLike? target then + emit "y"; emitPath path2; emitLn " = 0;" + else + emitLn "y.tag = 0;" + else + emit "if (x"; emitPath path; emitLn "== 0) {" + if let some path2 := optionLike? target then + emit "y"; emitPath path2; emitLn " = 0;" + else + emitLn "y.tag = 0;" + emitLn "} else {" + if (optionLike? target).isNone then + emitLn "y.tag = 1;" + reshapeUnion 1 + emitLn "}" + else if let some path := optionLike? target then + let compatible := compatibleReshape tys[1]! tys'[1]! + if !compatible then emit "y"; emitPath path; emitLn " = 0;" else + emitLn "if (x.tag == 0) {" + emit "y"; emitPath path; emitLn " = 0;" + emitLn "} else {" + reshapeUnion 1 + emitLn "}" + else + emitLn "y.tag = x.tag;" + -- Note: through cse in the mono stage we can get "bad" reshapes + -- where e.g. `let a : Option UInt8 := none` and `let a : Option (Option UInt8) := none` + -- merged into one variable. In that case, ignore branches with incompatibilities. + emitUnionSwitchWithImpossible tys.size "x.tag" + (fun i => compatibleReshape tys[i]! tys'[i]!) + fun i _ => reshapeUnion i + | .struct _ tys us ss, .struct _ tys' _ _ => + if us ≠ 0 ∨ ss ≠ 0 then + emit "memcpy(y.u, x.u, sizeof(size_t)*"; emit us; emit "+"; emit ss; emitLn ");" + for h : i in *...tys.size do + if tys'[i]! matches .erased | .void then + continue + if tys[i] matches .erased | .void then + emit "y.i"; emit i; emitLn " = lean_box(0);" + else if tys[i].compatibleWith tys'[i]! then + emit "y.i"; emit i; emit " = x.i"; emit i; emitLn ";" + else + emit "y.i"; emit i; emit " = " + emitReshapeFn tys[i] tys'[i]! + emit "(x.i"; emit i; emitLn ");" + if tys[i].isObj then + -- note: for unboxing functions the calling conventions don't match up + -- (has @& -> @&, expected owned -> owned) so we need to compensate + -- TODO: avoid these reference counting instructions when possible + if needsRC tys'[i]! then + emitIncOfType s!"y.i{i}" tys'[i]! 1 true "1" + emit "lean_dec(x.i"; emit i; emitLn ");" + | _, _ => pure () + emitLn "return y;" + emitLn "}" + +def emitStructBox (ty : IRType) (cidx : Nat) (x : String) : M Unit := do + let .struct _ tys us ss := ty | unreachable! + if tys.isEmpty && us == 0 && ss == 0 then + emit "y = lean_box("; emit cidx; emitLn ");" + return + emit "y = lean_alloc_ctor("; + emit cidx; emit ", "; emit tys.size; + emit ", sizeof(size_t)*"; emit us; emit "+"; emit ss + emitLn ");" + if us ≠ 0 ∨ ss ≠ 0 then + emit "memcpy(lean_ctor_scalar_cptr(y), "; emit x; emit ".u, sizeof(size_t)*" + emit us; emit "+"; emit ss; emitLn ");" + for h : i in *...tys.size do + emit "lean_ctor_set(y, "; emit i; emit ", " + if tys[i] matches .erased | .void then + emitLn "lean_box(0));" + else if tys[i] matches .object | .tobject | .tagged then + emit x; emit ".i"; emit i; emitLn ");" + else + emitBoxFn tys[i] .tobject + emit "("; emit x; emit ".i"; emit i; emitLn "));" + +def emitStructBoxFn (ty : IRType) (id : Nat) : M Unit := do + emit "static lean_object* " + emit structBoxFnPrefix; emit id + emit "("; emit (structType id); emitLn " x) {" + emitLn "lean_object* y;" + if let .union _ tys := ty then + if let some path := optionLike? ty then + emitLn "if (x"; emitPath path; emit " == 0) {" + emitLn "y = lean_box(0);" + emitLn "} else {" + emitStructBox tys[1]! 1 s!"x.cs.c1" + emitLn "}" + else + emitUnionSwitch tys.size "x.tag" fun i _ => do + emitStructBox tys[i] i s!"x.cs.c{i}" + else + emitStructBox ty 0 "x" + emitLn "return y;" + emitLn "}" + +def emitStructUnboxFn (ty : IRType) (id : Nat) : M Unit := do + emit "static "; emit (structType id); emit " " + emit structUnboxFnPrefix; emit id + emitLn "(lean_object* x) {" + emit (structType id); emitLn " y;" + + if let .union _ tys := ty then + if let some path := optionLike? ty then + emitLn "if (lean_is_scalar(x)) {" + emit "y"; emitPath path; emitLn " = 0;" + emitLn "} else {" + emit "y.cs.c1 = " + emitUnboxFn tys[1]! + emitLn "(x);" + emitLn "}" + else + emitLn "y.tag = lean_obj_tag(x);" + emitUnionSwitch tys.size "y.tag" fun i _ => do + emit "y.cs.c"; emit i; emit " = " + emitUnboxFn tys[i] + emitLn "(x);" + else + let .struct _ tys us ss := ty | unreachable! + if us ≠ 0 ∨ ss ≠ 0 then + emit "memcpy(y.u, lean_ctor_scalar_cptr(x), sizeof(size_t)*" + emit us; emit "+"; emit ss; emitLn ");" + for h : i in *...tys.size do + if tys[i] matches .erased | .void then + continue + else if tys[i] matches .object | .tobject | .tagged then + emit "y.i"; emit i; emit " = lean_ctor_get(x, "; emit i; emitLn ");" + else + emit "y.i"; emit i; emit " = " + emitUnboxFn tys[i] + emit "(lean_ctor_get(x, "; emit i; emitLn "));" + emitLn "return y;" + emitLn "}" + +def emitStructDeclsFor (id : Nat) (info : StructTypeInfo) : M Unit := do + let ty := info.type + emitStructDefn ty (structType id) + if needsRC ty then + emitStructIncDecFn ty id true + emitStructIncDecFn ty id false + emitStructBoxFn ty id + emitStructUnboxFn ty id + +def emitStructDecls (cont : M α) : M α := do + let env ← getEnv + let decls := getDecls env + let (data, lookup) := collectStructTypes decls + withReader (fun ctx => { ctx with structs := lookup }) do + let mut emitted : Std.HashSet Nat := {} + for h : i in *...data.size do + emitStructDeclsFor i data[i] + for h : i in *...data.size do + let info := data[i] + for reshape in info.reshapes do + emitStructReshapeFn info.type data[reshape]!.type i reshape + cont + def emitDeclAux (d : Decl) : M Unit := do let env ← getEnv - let (_, jpMap) := mkVarJPMaps d - withReader (fun ctx => { ctx with jpMap := jpMap }) do + let (varTypes, jpMap) := mkVarJPMaps d + withReader (fun ctx => { ctx with jpMap, varTypes }) do unless hasInitAttr env d.name do match d with | .fdecl (f := f) (xs := xs) (type := t) (body := b) .. => @@ -678,7 +1176,7 @@ def emitDeclAux (d : Decl) : M Unit := do emit "static " else emit "LEAN_EXPORT " -- make symbol visible to the interpreter - emit (toCType t); emit " "; + emit (← toCType t); emit " "; if xs.size > 0 then let xs := xs.filter (fun p => !p.ty.isVoid) emit baseName; @@ -686,10 +1184,7 @@ def emitDeclAux (d : Decl) : M Unit := do if xs.size > closureMaxArgs && isBoxedName d.name then emit "lean_object** _args" else - xs.size.forM fun i _ => do - if i > 0 then emit ", " - let x := xs[i] - emit (toCType x.ty); emit " "; emit x.x + emitSpreadArgs xs true emit ")" else emit ("_init_" ++ baseName ++ "()") @@ -698,13 +1193,17 @@ def emitDeclAux (d : Decl) : M Unit := do xs.size.forM fun i _ => do let x := xs[i]! emit "lean_object* "; emit x.x; emit " = _args["; emit i; emitLn "];" + emitSpreads xs emitLn "_start:"; withReader (fun ctx => { ctx with mainFn := f, mainParams := xs }) (emitFnBody b); emitLn "}" | _ => pure () def emitDecl (d : Decl) : M Unit := do - let d := d.normalizeIds; -- ensure we don't have gaps in the variable indices + -- this is a bit of weird optimization step since we are not allowed to run it for + -- interpreted code, so we run it here instead + let d := d.normalizeIds -- ensure we don't have gaps in the variable indices + let d := StructRC.visitDecl d try emitDeclAux d catch err => @@ -721,13 +1220,14 @@ def emitMarkPersistent (d : Decl) (n : Name) : M Unit := do emitCName n emitLn ");" +open ExplicitBoxing in def emitDeclInit (d : Decl) : M Unit := do let env ← getEnv let n := d.name if isIOUnitInitFn env n then if isIOUnitBuiltinInitFn env n then emit "if (builtin) {" - emit "res = "; emitCName n; emitLn "();" + emit "res = "; emitCName (mkBoxedName n); emitLn "(lean_box(0));" emitLn "if (lean_io_result_is_error(res)) return res;" emitLn "lean_dec_ref(res);" if isIOUnitBuiltinInitFn env n then @@ -737,11 +1237,13 @@ def emitDeclInit (d : Decl) : M Unit := do | some initFn => if getBuiltinInitFnNameFor? env d.name |>.isSome then emit "if (builtin) {" - emit "res = "; emitCName initFn; emitLn "();" + emit "res = "; emitCName (mkBoxedName initFn); emitLn "(lean_box(0));" emitLn "if (lean_io_result_is_error(res)) return res;" emitCName n - if d.resultType.isScalar then - emitLn (" = " ++ getUnboxOpName d.resultType ++ "(lean_io_result_get_value(res));") + if d.resultType.isScalarOrStruct then + emit " = " + emitUnboxFn d.resultType + emitLn "(lean_io_result_get_value(res));" else emitLn " = lean_io_result_get_value(res);" emitMarkPersistent d n @@ -777,6 +1279,8 @@ def emitInitFn : M Unit := do def main : M Unit := do emitFileHeader + emitLn "void* memcpy(void* restrict, const void* restrict, size_t);" + emitStructDecls do emitFnDecls emitFns emitInitFn diff --git a/src/Lean/Compiler/IR/EmitLLVM.lean b/src/Lean/Compiler/IR/EmitLLVM.lean index e3533874b845..767d28ee421d 100644 --- a/src/Lean/Compiler/IR/EmitLLVM.lean +++ b/src/Lean/Compiler/IR/EmitLLVM.lean @@ -339,7 +339,7 @@ def toLLVMType (t : IRType) : M llvmctx (LLVM.LLVMType llvmctx) := do | IRType.tobject => do LLVM.pointerType (← LLVM.i8Type llvmctx) | IRType.erased => do LLVM.pointerType (← LLVM.i8Type llvmctx) | IRType.void => do LLVM.pointerType (← LLVM.i8Type llvmctx) - | IRType.struct _ _ => panic! "not implemented yet" + | IRType.struct .. => panic! "not implemented yet" | IRType.union _ _ => panic! "not implemented yet" def throwInvalidExportName {α : Type} (n : Name) : M llvmctx α := do @@ -804,7 +804,8 @@ def callLeanCtorGet (builder : LLVM.Builder llvmctx) let i ← LLVM.buildSextOrTrunc builder i (← LLVM.i32Type llvmctx) LLVM.buildCall2 builder fnty fn #[x, i] retName -def emitProj (builder : LLVM.Builder llvmctx) (z : VarId) (i : Nat) (x : VarId) : M llvmctx Unit := do +def emitProj (builder : LLVM.Builder llvmctx) (z : VarId) (_c : Nat) (i : Nat) (x : VarId) : + M llvmctx Unit := do let xval ← emitLhsVal builder x let zval ← callLeanCtorGet builder xval (← constIntUnsigned i) "" emitLhsSlotStore builder z zval @@ -981,9 +982,9 @@ def emitVDecl (builder : LLVM.Builder llvmctx) (z : VarId) (t : IRType) (v : Exp | Expr.ctor c ys => emitCtor builder z c ys | Expr.reset n x => emitReset builder z n x | Expr.reuse x c u ys => emitReuse builder z x c u ys - | Expr.proj i x => emitProj builder z i x - | Expr.uproj i x => emitUProj builder z i x - | Expr.sproj n o x => emitSProj builder z t n o x + | Expr.proj c i x => emitProj builder z c i x + | Expr.uproj _ i x => emitUProj builder z i x + | Expr.sproj _ n o x => emitSProj builder z t n o x | Expr.fap c ys => emitFullApp builder z c ys | Expr.pap c ys => emitPartialApp builder z c ys | Expr.ap x ys => emitApp builder z x ys @@ -1171,8 +1172,8 @@ partial def emitBlock (builder : LLVM.Builder llvmctx) (b : FnBody) : M llvmctx | FnBody.del x b => emitDel builder x; emitBlock builder b | FnBody.setTag x i b => emitSetTag builder x i; emitBlock builder b | FnBody.set x i y b => emitSet builder x i y; emitBlock builder b - | FnBody.uset x i y b => emitUSet builder x i y; emitBlock builder b - | FnBody.sset x i o y t b => emitSSet builder x i o y t; emitBlock builder b + | FnBody.uset x _ i y b => emitUSet builder x i y; emitBlock builder b + | FnBody.sset x _ i o y t b => emitSSet builder x i o y t; emitBlock builder b | FnBody.ret x => do let (_xty, xv) ← emitArgVal builder x "ret_val" let _ ← LLVM.buildRet builder xv diff --git a/src/Lean/Compiler/IR/EmitUtil.lean b/src/Lean/Compiler/IR/EmitUtil.lean index b26748ecb20c..47a7ae41d69d 100644 --- a/src/Lean/Compiler/IR/EmitUtil.lean +++ b/src/Lean/Compiler/IR/EmitUtil.lean @@ -8,12 +8,14 @@ module prelude public import Lean.Compiler.InitAttr public import Lean.Compiler.IR.CompilerM +import Lean.Compiler.NameMangling public section /-! # Helper functions for backend code generators -/ namespace Lean.IR + /-- Return true iff `b` is of the form `let x := g ys; ret x` -/ def isTailCallTo (g : Name) (b : FnBody) : Bool := match b with @@ -23,6 +25,102 @@ def isTailCallTo (g : Name) (b : FnBody) : Bool := def usesModuleFrom (env : Environment) (modulePrefix : Name) : Bool := env.allImportedModuleNames.toList.any fun modName => modulePrefix.isPrefixOf modName +/-- +Wrapper around `IRType` where different object types are considered equal. +-/ +structure IRTypeApprox where + type : IRType + +partial def IRType.hashApprox : IRType → UInt64 + | .float => 11 + | .uint8 => 13 + | .uint16 => 17 + | .uint32 => 19 + | .uint64 => 23 + | .usize => 29 + | .float32 => 31 + | .erased | .object | .tobject | .tagged | .void => 37 + | .struct _ tys us ss => + let : Hashable IRType := { hash := hashApprox } + mixHash (mixHash (mixHash 41 (hash tys)) (hash us)) (hash ss) + | .union _ tys => + let : Hashable IRType := { hash := hashApprox } + mixHash 43 (hash tys) + +instance : BEq IRTypeApprox := ⟨fun a b => a.type.compatibleWith b.type (strict := true)⟩ +instance : Hashable IRTypeApprox := ⟨fun a => a.type.hashApprox⟩ + +structure StructTypeInfo where + /-- The type itself after `normalizeObject`. -/ + type : IRType + /-- The ids of struct types that this type can Reshape to. -/ + reshapes : Array Nat +deriving Inhabited + +abbrev StructTypeData := Array StructTypeInfo +abbrev StructTypeLookup := Std.HashMap IRTypeApprox Nat + +namespace CollectStructTypes + +abbrev M := StateM (StructTypeData × StructTypeLookup) + +partial def registerType (ty : IRType) : M Unit := do + match ty with + | .struct _ tys _ _ => tys.forM registerType + | .union _ tys => tys.forM registerType + | _ => return + + let (arr, map) ← get + match map[IRTypeApprox.mk ty]? with + | none => + let id := arr.size + let ty := ty.normalizeObject + modify fun m => (m.1.push ⟨ty, #[]⟩, m.2.insert ⟨ty⟩ id) + | some _ => pure () + +def addReshapeEntry (origin target : IRType) : M Unit := do + let id1 := (← get).2[IRTypeApprox.mk origin]! + let id2 := (← get).2[IRTypeApprox.mk target]! + modify fun m => (m.1.modify id1 fun info => + if info.reshapes.contains id2 then info + else { info with reshapes := info.reshapes.push id2 }, m.2) + +def addReshape (origin target : IRType) : M Unit := do + match origin, target with + | .struct _ tys _ _, .struct _ tys' _ _ + | .union _ tys, .union _ tys' => + for ty in tys, ty' in tys' do + addReshape ty ty' + addReshapeEntry origin target + | _, _ => pure () -- ignore + +def collectParams (params : Array Param) : M Unit := do + for x in params do + if x.ty.isStruct then + registerType x.ty + +partial def collectFnBody : FnBody → M Unit + | .vdecl _ t v b => do + if t.isStruct then + registerType t + match v with + | .box t' _ => addReshape t' t + | _ => pure () + collectFnBody b + | .jdecl _ xs v b => + collectParams xs *> collectFnBody v *> collectFnBody b + | .case _ _ _ alts => alts.forM fun alt => collectFnBody alt.body + | e => do unless e.isTerminal do collectFnBody e.body + +def collectDecl : Decl → M Unit + | .fdecl _f xs ty b _ => collectParams xs *> registerType ty *> collectFnBody b + | .extern _f xs ty _ => collectParams xs *> registerType ty + +end CollectStructTypes + +def collectStructTypes (decls : List Decl) : StructTypeData × StructTypeLookup := + ((decls.forM CollectStructTypes.collectDecl).run ({}, {})).2 + namespace CollectUsedDecls abbrev M := ReaderT Environment (StateM NameSet) diff --git a/src/Lean/Compiler/IR/ExpandResetReuse.lean b/src/Lean/Compiler/IR/ExpandResetReuse.lean index dd8c08073810..a265f3c1daba 100644 --- a/src/Lean/Compiler/IR/ExpandResetReuse.lean +++ b/src/Lean/Compiler/IR/ExpandResetReuse.lean @@ -62,13 +62,13 @@ partial def eraseProjIncForAux (y : VarId) (bs : Array FnBody) (mask : Mask) (ke else let b := bs.back! match b with - | .vdecl _ _ (.sproj _ _ _) _ => keepInstr b - | .vdecl _ _ (.uproj _ _) _ => keepInstr b + | .vdecl _ _ (.sproj ..) _ => keepInstr b + | .vdecl _ _ (.uproj ..) _ => keepInstr b | .inc z n c p _ => if n == 0 then done () else let b' := bs[bs.size - 2] match b' with - | .vdecl w _ (.proj i x) _ => + | .vdecl w _ (.proj _ i x) _ => if w == z && y == x then /- Found ``` @@ -143,7 +143,9 @@ def releaseUnreadFields (y : VarId) (mask : Mask) (b : FnBody) : M FnBody := | some _ => pure b -- code took ownership of this field | none => do let fld ← mkFresh - pure (FnBody.vdecl fld .tobject (Expr.proj i y) (FnBody.dec fld 1 true false b)) + -- we don't need cidx information here since `reset` cannot occur + -- on a boxed value at this stage + pure (FnBody.vdecl fld .tobject (Expr.proj 0 i y) (FnBody.dec fld 1 true false b)) def setFields (y : VarId) (zs : Array Arg) (b : FnBody) : FnBody := zs.size.fold (init := b) fun i _ b => FnBody.set y i zs[i] b @@ -153,33 +155,33 @@ def isSelfSet (ctx : Context) (x : VarId) (i : Nat) (y : Arg) : Bool := match y with | .var y => match ctx.projMap[y]? with - | some (Expr.proj j w) => j == i && w == x + | some (Expr.proj _ j w) => j == i && w == x | _ => false | .erased => false /-- Given `uset x[i] := y`, return true iff `y := uproj[i] x` -/ def isSelfUSet (ctx : Context) (x : VarId) (i : Nat) (y : VarId) : Bool := match ctx.projMap[y]? with - | some (Expr.uproj j w) => j == i && w == x - | _ => false + | some (Expr.uproj _ j w) => j == i && w == x + | _ => false /-- Given `sset x[n, i] := y`, return true iff `y := sproj[n, i] x` -/ def isSelfSSet (ctx : Context) (x : VarId) (n : Nat) (i : Nat) (y : VarId) : Bool := match ctx.projMap[y]? with - | some (Expr.sproj m j w) => n == m && j == i && w == x - | _ => false + | some (Expr.sproj _ m j w) => n == m && j == i && w == x + | _ => false /-- Remove unnecessary `set/uset/sset` operations -/ partial def removeSelfSet (ctx : Context) : FnBody → FnBody | FnBody.set x i y b => if isSelfSet ctx x i y then removeSelfSet ctx b else FnBody.set x i y (removeSelfSet ctx b) - | FnBody.uset x i y b => + | FnBody.uset x c i y b => if isSelfUSet ctx x i y then removeSelfSet ctx b - else FnBody.uset x i y (removeSelfSet ctx b) - | FnBody.sset x n i y t b => + else FnBody.uset x c i y (removeSelfSet ctx b) + | FnBody.sset x c n i y t b => if isSelfSSet ctx x n i y then removeSelfSet ctx b - else FnBody.sset x n i y t (removeSelfSet ctx b) + else FnBody.sset x c n i y t (removeSelfSet ctx b) | FnBody.case tid y yType alts => let alts := alts.map fun alt => alt.modifyBody (removeSelfSet ctx) FnBody.case tid y yType alts @@ -237,16 +239,18 @@ def mkFastPath (x y : VarId) (mask : Mask) (b : FnBody) : M FnBody := do let b := reuseToSet ctx x y b releaseUnreadFields y mask b +mutual + -- Expand `bs; x := reset[n] y; b` -partial def expand (mainFn : FnBody → Array FnBody → M FnBody) +partial def expand (bs : Array FnBody) (x : VarId) (n : Nat) (y : VarId) (b : FnBody) : M FnBody := do let (bs, mask) := eraseProjIncFor n y bs /- Remark: we may be duplicating variable/JP indices. That is, `bSlow` and `bFast` may - have duplicate indices. We run `normalizeIds` to fix the ids after we have expand them. -/ - let bSlow := mkSlowPath x y mask b + have duplicate indices. We run `normalizeIds` to fix the ids after we have expanded them. -/ + let bSlow := mkSlowPath x y mask b let bFast ← mkFastPath x y mask b /- We only optimize recursively the fast. -/ - let bFast ← mainFn bFast #[] + let bFast ← searchAndExpand bFast #[] let c ← mkFresh let b := FnBody.vdecl c IRType.uint8 (Expr.isShared y) (mkIf c bSlow bFast) return reshape bs b @@ -254,7 +258,7 @@ partial def expand (mainFn : FnBody → Array FnBody → M FnBody) partial def searchAndExpand : FnBody → Array FnBody → M FnBody | d@(FnBody.vdecl x _ (Expr.reset n y) b), bs => if consumed x b then do - expand searchAndExpand bs x n y b + expand bs x n y b else searchAndExpand b (push bs d) | FnBody.jdecl j xs v b, bs => do @@ -267,6 +271,8 @@ partial def searchAndExpand : FnBody → Array FnBody → M FnBody if b.isTerminal then return reshape bs b else searchAndExpand b.body (push bs b) +end + def main (d : Decl) : Decl := match d with | .fdecl (body := b) .. => diff --git a/src/Lean/Compiler/IR/Format.lean b/src/Lean/Compiler/IR/Format.lean index ff1b04a80271..fc6ebf19dcb3 100644 --- a/src/Lean/Compiler/IR/Format.lean +++ b/src/Lean/Compiler/IR/Format.lean @@ -43,9 +43,9 @@ private def formatExpr : Expr → Format | Expr.ctor i ys => format i ++ formatArray ys | Expr.reset n x => "reset[" ++ format n ++ "] " ++ format x | Expr.reuse x i u ys => "reuse" ++ (if u then "!" else "") ++ " " ++ format x ++ " in " ++ format i ++ formatArray ys - | Expr.proj i x => "proj[" ++ format i ++ "] " ++ format x - | Expr.uproj i x => "uproj[" ++ format i ++ "] " ++ format x - | Expr.sproj n o x => "sproj[" ++ format n ++ ", " ++ format o ++ "] " ++ format x + | Expr.proj c i x => "proj[" ++ format c ++ ", " ++ format i ++ "] " ++ format x + | Expr.uproj c i x => "uproj[" ++ format c ++ ", " ++ format i ++ "] " ++ format x + | Expr.sproj c n o x => "sproj[" ++ format c ++ ", " ++ format n ++ ", " ++ format o ++ "] " ++ format x | Expr.fap c ys => format c ++ formatArray ys | Expr.pap c ys => "pap " ++ format c ++ formatArray ys | Expr.ap x ys => "app " ++ format x ++ formatArray ys @@ -70,12 +70,13 @@ private partial def formatIRType : IRType → Format | IRType.object => "obj" | IRType.tagged => "tagged" | IRType.tobject => "tobj" - | IRType.struct _ tys => + | IRType.struct _ tys nu ns => let _ : ToFormat IRType := ⟨formatIRType⟩ - "struct " ++ Format.bracket "{" (Format.joinSep tys.toList ", ") "}" + let fmt := Format.bracket "{" (Format.joinSep tys.toList ", ") "}" + if nu = 0 ∧ ns = 0 then fmt else format (nu.repr ++ ":" ++ ns.repr) ++ fmt | IRType.union _ tys => let _ : ToFormat IRType := ⟨formatIRType⟩ - "union " ++ Format.bracket "{" (Format.joinSep tys.toList ", ") "}" + Format.bracket "(" (Format.joinSep tys.toList " | ") ")" instance : ToFormat IRType := ⟨private_decl% formatIRType⟩ instance : ToString IRType := ⟨toString ∘ format⟩ @@ -96,8 +97,8 @@ def formatFnBodyHead : FnBody → Format | FnBody.vdecl x ty e _ => "let " ++ format x ++ " : " ++ format ty ++ " := " ++ format e | FnBody.jdecl j xs _ _ => format j ++ formatParams xs ++ " := ..." | FnBody.set x i y _ => "set " ++ format x ++ "[" ++ format i ++ "] := " ++ format y - | FnBody.uset x i y _ => "uset " ++ format x ++ "[" ++ format i ++ "] := " ++ format y - | FnBody.sset x i o y ty _ => "sset " ++ format x ++ "[" ++ format i ++ ", " ++ format o ++ "] : " ++ format ty ++ " := " ++ format y + | FnBody.uset x c i y _ => "uset " ++ format x ++ "[" ++ format c ++ ", " ++ format i ++ "] := " ++ format y + | FnBody.sset c x i o y ty _ => "sset " ++ format x ++ "[" ++ format c ++ ", " ++ format i ++ ", " ++ format o ++ "] : " ++ format ty ++ " := " ++ format y | FnBody.setTag x cidx _ => "setTag " ++ format x ++ " := " ++ format cidx | FnBody.inc x n _ _ _ => "inc" ++ (if n != 1 then Format.sbracket (format n) else "") ++ " " ++ format x | FnBody.dec x n _ _ _ => "dec" ++ (if n != 1 then Format.sbracket (format n) else "") ++ " " ++ format x @@ -116,8 +117,8 @@ partial def formatFnBody (fnBody : FnBody) (indent : Nat := 2) : Format := | FnBody.vdecl x ty e b => "let " ++ format x ++ " : " ++ format ty ++ " := " ++ format e ++ ";" ++ Format.line ++ loop b | FnBody.jdecl j xs v b => format j ++ formatParams xs ++ " :=" ++ Format.nest indent (Format.line ++ loop v) ++ ";" ++ Format.line ++ loop b | FnBody.set x i y b => "set " ++ format x ++ "[" ++ format i ++ "] := " ++ format y ++ ";" ++ Format.line ++ loop b - | FnBody.uset x i y b => "uset " ++ format x ++ "[" ++ format i ++ "] := " ++ format y ++ ";" ++ Format.line ++ loop b - | FnBody.sset x i o y ty b => "sset " ++ format x ++ "[" ++ format i ++ ", " ++ format o ++ "] : " ++ format ty ++ " := " ++ format y ++ ";" ++ Format.line ++ loop b + | FnBody.uset x c i y b => "uset " ++ format x ++ "[" ++ format c ++ ", " ++ format i ++ "] := " ++ format y ++ ";" ++ Format.line ++ loop b + | FnBody.sset x c i o y ty b => "sset " ++ format x ++ "[" ++ format c ++ ", " ++ format i ++ ", " ++ format o ++ "] : " ++ format ty ++ " := " ++ format y ++ ";" ++ Format.line ++ loop b | FnBody.setTag x cidx b => "setTag " ++ format x ++ " := " ++ format cidx ++ ";" ++ Format.line ++ loop b | FnBody.inc x n _ _ b => "inc" ++ (if n != 1 then Format.sbracket (format n) else "") ++ " " ++ format x ++ ";" ++ Format.line ++ loop b | FnBody.dec x n _ _ b => "dec" ++ (if n != 1 then Format.sbracket (format n) else "") ++ " " ++ format x ++ ";" ++ Format.line ++ loop b diff --git a/src/Lean/Compiler/IR/FreeVars.lean b/src/Lean/Compiler/IR/FreeVars.lean index 05fab205f123..f0f70ccfc2db 100644 --- a/src/Lean/Compiler/IR/FreeVars.lean +++ b/src/Lean/Compiler/IR/FreeVars.lean @@ -45,7 +45,7 @@ private def visitParam (p : Param) : M Unit := private def visitExpr (e : Expr) : M Unit := do match e with - | .proj _ x | .uproj _ x | .sproj _ _ x | .box _ x | .unbox x | .reset _ x | .isShared x => + | .proj _ _ x | .uproj _ _ x | .sproj _ _ _ x | .box _ x | .unbox x | .reset _ x | .isShared x => visitVar x | .ctor _ ys | .fap _ ys | .pap _ ys => ys.forM visitArg @@ -69,7 +69,7 @@ partial def visitFnBody (fnBody : FnBody) : M Unit := do visitVar x visitArg y visitFnBody b - | .uset x _ y b | .sset x _ _ y _ b => + | .uset x _ _ y b | .sset x _ _ _ y _ b => visitVar x visitVar y visitFnBody b @@ -132,7 +132,7 @@ private def visitParam (p : Param) : M Unit := private def visitExpr (e : Expr) : M Unit := do match e with - | .proj _ x | .uproj _ x | .sproj _ _ x | .box _ x | .unbox x | .reset _ x | .isShared x => + | .proj _ _ x | .uproj _ _ x | .sproj _ _ _ x | .box _ x | .unbox x | .reset _ x | .isShared x => visitVar x | .ctor _ ys | .fap _ ys | .pap _ ys => ys.forM visitArg @@ -156,7 +156,7 @@ partial def visitFnBody (fnBody : FnBody) : M Unit := do visitVar x visitArg y visitFnBody b - | .uset x _ y b | .sset x _ _ y _ b => + | .uset x _ _ y b | .sset x _ _ _ y _ b => visitVar x visitVar y visitFnBody b @@ -209,7 +209,7 @@ def visitParams (w : Index) (ps : Array Param) : Bool := ps.any (fun p => w == p.x.idx) def visitExpr (w : Index) : Expr → Bool - | .proj _ x | .uproj _ x | .sproj _ _ x | .box _ x | .unbox x | .reset _ x | .isShared x => + | .proj _ _ x | .uproj _ _ x | .sproj _ _ _ x | .box _ x | .unbox x | .reset _ x | .isShared x => visitVar w x | .ctor _ ys | .fap _ ys | .pap _ ys => visitArgs w ys @@ -224,7 +224,7 @@ partial def visitFnBody (w : Index) : FnBody → Bool visitFnBody w v || visitFnBody w b | FnBody.set x _ y b => visitVar w x || visitArg w y || visitFnBody w b - | .uset x _ y b | .sset x _ _ y _ b => + | .uset x _ _ y b | .sset x _ _ _ y _ b => visitVar w x || visitVar w y || visitFnBody w b | .setTag x _ b | .inc x _ _ _ b | .dec x _ _ _ b | .del x b => visitVar w x || visitFnBody w b diff --git a/src/Lean/Compiler/IR/LiveVars.lean b/src/Lean/Compiler/IR/LiveVars.lean index c0bbfc5ba514..10cd42fc7f36 100644 --- a/src/Lean/Compiler/IR/LiveVars.lean +++ b/src/Lean/Compiler/IR/LiveVars.lean @@ -53,7 +53,7 @@ partial def visitFnBody (w : Index) : FnBody → M Bool visitFnBody w v <||> visitFnBody w b | .set x _ y b => visitVar w x <||> visitArg w y <||> visitFnBody w b - | .uset x _ y b | .sset x _ _ y _ b => + | .uset x _ _ y b | .sset x _ _ _ y _ b => visitVar w x <||> visitVar w y <||> visitFnBody w b | .setTag x _ b | .inc x _ _ _ b | .dec x _ _ _ b | .del x b => visitVar w x <||> visitFnBody w b @@ -118,7 +118,7 @@ private def visitJP (m : JPLiveVarMap) (j : JoinPointId) : M Unit := private def useExpr (e : Expr) : M Unit := do match e with - | .proj _ x | .uproj _ x | .sproj _ _ x | .box _ x | .unbox x | .reset _ x | .isShared x => + | .proj _ _ x | .uproj _ _ x | .sproj _ _ _ x | .box _ x | .unbox x | .reset _ x | .isShared x => useVar x | .ctor _ ys | .fap _ ys | .pap _ ys => ys.forM useArg @@ -141,7 +141,7 @@ private partial def visitFnBody (fnBody : FnBody) (m : JPLiveVarMap) : M Unit := visitFnBody b m useVar x useArg y - | .uset x _ y b | .sset x _ _ y _ b => + | .uset x _ _ y b | .sset x _ _ _ y _ b => visitFnBody b m useVar x useVar y diff --git a/src/Lean/Compiler/IR/NormIds.lean b/src/Lean/Compiler/IR/NormIds.lean index c473d10d9b66..8ce888e76e84 100644 --- a/src/Lean/Compiler/IR/NormIds.lean +++ b/src/Lean/Compiler/IR/NormIds.lean @@ -64,9 +64,9 @@ def normExpr : Expr → M Expr | Expr.ctor c ys, m => Expr.ctor c (normArgs ys m) | Expr.reset n x, m => Expr.reset n (normVar x m) | Expr.reuse x c u ys, m => Expr.reuse (normVar x m) c u (normArgs ys m) - | Expr.proj i x, m => Expr.proj i (normVar x m) - | Expr.uproj i x, m => Expr.uproj i (normVar x m) - | Expr.sproj n o x, m => Expr.sproj n o (normVar x m) + | Expr.proj c i x, m => Expr.proj c i (normVar x m) + | Expr.uproj c i x, m => Expr.uproj c i (normVar x m) + | Expr.sproj c n o x, m => Expr.sproj c n o (normVar x m) | Expr.fap c ys, m => Expr.fap c (normArgs ys m) | Expr.pap c ys, m => Expr.pap c (normArgs ys m) | Expr.ap x ys, m => Expr.ap (normVar x m) (normArgs ys m) @@ -100,20 +100,20 @@ partial def normFnBody : FnBody → N FnBody | FnBody.jdecl j ys v b => do let (ys, v) ← withParams ys fun ys => do let v ← normFnBody v; pure (ys, v) withJP j fun j => return FnBody.jdecl j ys v (← normFnBody b) - | FnBody.set x i y b => return FnBody.set (← normVar x) i (← normArg y) (← normFnBody b) - | FnBody.uset x i y b => return FnBody.uset (← normVar x) i (← normVar y) (← normFnBody b) - | FnBody.sset x i o y t b => return FnBody.sset (← normVar x) i o (← normVar y) t (← normFnBody b) - | FnBody.setTag x i b => return FnBody.setTag (← normVar x) i (← normFnBody b) - | FnBody.inc x n c p b => return FnBody.inc (← normVar x) n c p (← normFnBody b) - | FnBody.dec x n c p b => return FnBody.dec (← normVar x) n c p (← normFnBody b) - | FnBody.del x b => return FnBody.del (← normVar x) (← normFnBody b) + | FnBody.set x i y b => return FnBody.set (← normVar x) i (← normArg y) (← normFnBody b) + | FnBody.uset x c i y b => return FnBody.uset (← normVar x) c i (← normVar y) (← normFnBody b) + | FnBody.sset x c i o y t b => return FnBody.sset (← normVar x) c i o (← normVar y) t (← normFnBody b) + | FnBody.setTag x i b => return FnBody.setTag (← normVar x) i (← normFnBody b) + | FnBody.inc x n c p b => return FnBody.inc (← normVar x) n c p (← normFnBody b) + | FnBody.dec x n c p b => return FnBody.dec (← normVar x) n c p (← normFnBody b) + | FnBody.del x b => return FnBody.del (← normVar x) (← normFnBody b) | FnBody.case tid x xType alts => do let x ← normVar x let alts ← alts.mapM fun alt => alt.modifyBodyM normFnBody return FnBody.case tid x xType alts - | FnBody.jmp j ys => return FnBody.jmp (← normJP j) (← normArgs ys) - | FnBody.ret x => return FnBody.ret (← normArg x) - | FnBody.unreachable => pure FnBody.unreachable + | FnBody.jmp j ys => return FnBody.jmp (← normJP j) (← normArgs ys) + | FnBody.ret x => return FnBody.ret (← normArg x) + | FnBody.unreachable => pure FnBody.unreachable def normDecl (d : Decl) : N Decl := match d with @@ -141,9 +141,9 @@ def mapExpr (f : VarId → VarId) : Expr → Expr | Expr.ctor c ys => Expr.ctor c (mapArgs f ys) | Expr.reset n x => Expr.reset n (f x) | Expr.reuse x c u ys => Expr.reuse (f x) c u (mapArgs f ys) - | Expr.proj i x => Expr.proj i (f x) - | Expr.uproj i x => Expr.uproj i (f x) - | Expr.sproj n o x => Expr.sproj n o (f x) + | Expr.proj c i x => Expr.proj c i (f x) + | Expr.uproj c i x => Expr.uproj c i (f x) + | Expr.sproj c n o x => Expr.sproj c n o (f x) | Expr.fap c ys => Expr.fap c (mapArgs f ys) | Expr.pap c ys => Expr.pap c (mapArgs f ys) | Expr.ap x ys => Expr.ap (f x) (mapArgs f ys) @@ -157,8 +157,8 @@ partial def mapFnBody (f : VarId → VarId) : FnBody → FnBody | FnBody.jdecl j ys v b => FnBody.jdecl j ys (mapFnBody f v) (mapFnBody f b) | FnBody.set x i y b => FnBody.set (f x) i (mapArg f y) (mapFnBody f b) | FnBody.setTag x i b => FnBody.setTag (f x) i (mapFnBody f b) - | FnBody.uset x i y b => FnBody.uset (f x) i (f y) (mapFnBody f b) - | FnBody.sset x i o y t b => FnBody.sset (f x) i o (f y) t (mapFnBody f b) + | FnBody.uset x c i y b => FnBody.uset (f x) c i (f y) (mapFnBody f b) + | FnBody.sset x c i o y t b => FnBody.sset (f x) c i o (f y) t (mapFnBody f b) | FnBody.inc x n c p b => FnBody.inc (f x) n c p (mapFnBody f b) | FnBody.dec x n c p b => FnBody.dec (f x) n c p (mapFnBody f b) | FnBody.del x b => FnBody.del (f x) (mapFnBody f b) diff --git a/src/Lean/Compiler/IR/PushProj.lean b/src/Lean/Compiler/IR/PushProj.lean index 9d4f7c64b187..065e4faa3d68 100644 --- a/src/Lean/Compiler/IR/PushProj.lean +++ b/src/Lean/Compiler/IR/PushProj.lean @@ -32,11 +32,11 @@ partial def pushProjs (bs : Array FnBody) (alts : Array Alt) (altsF : Array Inde match b with | FnBody.vdecl x _ v _ => match v with - | Expr.proj _ _ => push x - | Expr.uproj _ _ => push x - | Expr.sproj _ _ _ => push x - | Expr.isShared _ => skip () - | _ => done () + | Expr.proj .. => push x + | Expr.uproj .. => push x + | Expr.sproj .. => push x + | Expr.isShared _ => skip () + | _ => done () | _ => done () partial def FnBody.pushProj (b : FnBody) : FnBody := diff --git a/src/Lean/Compiler/IR/RC.lean b/src/Lean/Compiler/IR/RC.lean index 026894d25632..a7b088c77569 100644 --- a/src/Lean/Compiler/IR/RC.lean +++ b/src/Lean/Compiler/IR/RC.lean @@ -66,10 +66,11 @@ private partial def removeFromParent (child : VarId) : M Unit := do private partial def visitFnBody (b : FnBody) : M Unit := do match b with - | .vdecl x _ e b => + | .vdecl x t e b => match e with - | .proj _ parent => - addDerivedValue parent x + | .proj _ _ parent => + unless t.isScalar do + addDerivedValue parent x | .fap ``Array.getInternal args => if let .var parent := args[1]! then addDerivedValue parent x @@ -78,6 +79,9 @@ private partial def visitFnBody (b : FnBody) : M Unit := do addDerivedValue parent x | .reset _ x => removeFromParent x + | .unbox y => + if t.isStruct then + addDerivedValue y x | _ => pure () visitFnBody b | .jdecl _ ps v b => @@ -185,7 +189,7 @@ private def useArgs (ctx : Context) (args : Array Arg) (liveVars : LiveVars) : L private def useExpr (ctx : Context) (e : Expr) (liveVars : LiveVars) : LiveVars := match e with - | .proj _ x | .uproj _ x | .sproj _ _ x | .box _ x | .unbox x | .reset _ x | .isShared x => + | .proj _ _ x | .uproj _ _ x | .sproj _ _ _ x | .box _ x | .unbox x | .reset _ x | .isShared x => useVar ctx x liveVars | .ctor _ ys | .fap _ ys | .pap _ ys => useArgs ctx ys liveVars @@ -356,11 +360,11 @@ private def processVDecl (ctx : Context) (z : VarId) (t : IRType) (v : Expr) (b let b := match v with | .ctor _ ys | .reuse _ _ _ ys | .pap _ ys => addIncBeforeConsumeAll ctx ys (.vdecl z t v b) bLiveVars - | .proj _ x => + | .proj _ _ x | .unbox x => let b := addDecIfNeeded ctx x b bLiveVars - let b := if !bLiveVars.borrows.contains z then addInc ctx z b else b + let b := if t.isPossibleRef && !bLiveVars.borrows.contains z then addInc ctx z b else b .vdecl z t v b - | .uproj _ x | .sproj _ _ x | .unbox x => + | .uproj _ _ x | .sproj _ _ _ x => .vdecl z t v (addDecIfNeeded ctx x b bLiveVars) | .fap f ys => let ps := (getDecl ctx f).params @@ -376,7 +380,16 @@ private def processVDecl (ctx : Context) (z : VarId) (t : IRType) (v : Expr) (b | .ap x ys => let ysx := ys.push (.var x) -- TODO: avoid temporary array allocation addIncBeforeConsumeAll ctx ysx (.vdecl z t v b) bLiveVars - | .lit _ | .box .. | .reset .. | .isShared _ => + | .box xTy x => + if xTy.isStruct then + let b := .vdecl z t v b + if bLiveVars.vars.contains x || bLiveVars.borrows.contains x then + addInc ctx x b + else + b + else + .vdecl z t v b + | .lit _ | .reset .. | .isShared _ => .vdecl z t v b let liveVars := useExpr ctx v bLiveVars let liveVars := bindVar z liveVars @@ -406,16 +419,16 @@ partial def visitFnBody (b : FnBody) (ctx : Context) : FnBody × LiveVars := } let ⟨b, bLiveVars⟩ := visitFnBody b ctx ⟨.jdecl j xs v b, bLiveVars⟩ - | .uset x i y b => + | .uset x c i y b => let ⟨b, s⟩ := visitFnBody b ctx -- We don't need to insert `y` since we only need to track live variables that are references at runtime let s := useVar ctx x s - ⟨.uset x i y b, s⟩ - | .sset x i o y t b => + ⟨.uset x c i y b, s⟩ + | .sset x c i o y t b => let ⟨b, s⟩ := visitFnBody b ctx -- We don't need to insert `y` since we only need to track live variables that are references at runtime let s := useVar ctx x s - ⟨.sset x i o y t b, s⟩ + ⟨.sset x c i o y t b, s⟩ | .case tid x xType alts => let alts : Array (Alt × LiveVars) := alts.map fun alt => match alt with | .ctor c b => diff --git a/src/Lean/Compiler/IR/ResetReuse.lean b/src/Lean/Compiler/IR/ResetReuse.lean index e6cf1d065527..37c940ecb2d1 100644 --- a/src/Lean/Compiler/IR/ResetReuse.lean +++ b/src/Lean/Compiler/IR/ResetReuse.lean @@ -49,7 +49,7 @@ private partial def S (w : VarId) (c : CtorInfo) (relaxedReuse : Bool) (b : FnBo where go : FnBody → FnBody | .vdecl x t v@(.ctor c' ys) b => - if mayReuse c c' relaxedReuse then + if t.isObj && mayReuse c c' relaxedReuse then let updtCidx := c.cidx != c'.cidx .vdecl x t (.reuse w c' updtCidx ys) b else @@ -223,6 +223,9 @@ private def D (x : VarId) (c : CtorInfo) (b : FnBody) : M FnBody := partial def R (e : FnBody) : M FnBody := do match e with | .case tid x xType alts => + -- Unboxed struct/union values cannot be reused + if xType.isStruct then + return .case tid x xType (← alts.mapM <| Alt.modifyBodyM R) let alreadyFound := (← read).alreadyFound.contains x withReader (fun ctx => { ctx with alreadyFound := ctx.alreadyFound.insert x }) do let alts ← alts.mapM fun alt => do diff --git a/src/Lean/Compiler/IR/StructRC.lean b/src/Lean/Compiler/IR/StructRC.lean new file mode 100644 index 000000000000..333c02729032 --- /dev/null +++ b/src/Lean/Compiler/IR/StructRC.lean @@ -0,0 +1,457 @@ +/- +Copyright (c) 2026 Robin Arnez. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Robin Arnez +-/ +module + +prelude +public import Lean.Compiler.IR.FreeVars +public import Lean.Compiler.IR.Format +import Std.Data.TreeMap.AdditionalOperations + +public section + +/-! +Optimizes reference counting instructions on struct/union variables. This pass assumes that an +inc/dec on a struct/union corresponds to an inc/dec on all of its (active) non-scalar fields. +This assumption is incompatible with the model of struct/union values that the interpreter uses: +In the interpreter, struct/union values are simply passed as their boxed form and thus have their +own reference count, contradicting the model of "inc/dec all fields". + +Thus this pass is only run *after* the regular IR passes, specifically from EmitC, or for debugging +purposes when the trace class is activated. +-/ + +namespace Lean.IR.StructRC + +/-- +Describes an argument of a constructor. +-/ +inductive ProjEntry where + /-- The value is erased. -/ + | erased + /-- + The projection is contained in the variable `v` and there is an entry in the context that + describes the obligations on `v`. + -/ + | var (v : VarId) + /-- + The projection was not bound to variable and its reference count still has to be changed by `rc`. + -/ + | unbound (rc : Int) +deriving Inhabited, Repr + +/-- +Describes accumulated RC obligations of a variable. +-/ +inductive Entry where + /-- The value is persistent or has a type that doesn't require reference counting (e.g. scalars). -/ + | persistent + /-- The value is of type `IRType.union _ tys` with some unknown constructor and the reference + count of each field still has to be changed by `rc`. -/ + | unknownUnion (tys : Array IRType) (rc : Int) + /-- The value is of type `IRType.struct _ tys _ _` or `IRType.union _ utys _ _` with + `utys[cidx] = IRType.struct _ tys _ _` and `fields` describes the RC obligations of each field. -/ + | struct (cidx : Nat) (tys : Array IRType) (fields : Vector ProjEntry tys.size) + /-- + The value is of type `object` (if `checkRef` is false) or `tobject` (if `checkRef` is true) and + its reference count still has to be changed by `rc`. + -/ + | ref (checkRef : Bool) (rc : Int) +deriving Inhabited, Repr + +/-- The context of `visitFnBody`. -/ +structure Context where + /-- A map from variable id to its corresponding RC obligations. There must be an entry for every + `struct`/`union` variable. -/ + vars : Std.TreeMap VarId Entry fun a b => compare a.idx b.idx := {} + /-- All preceding instructions that need to be `reshape`d. This is used to make `visitFnBody` + tail-recursive. -/ + instrs : Array FnBody := #[] + rename : Std.TreeMap VarId Arg fun a b => compare a.idx b.idx := {} + +def Context.renameVar (ctx : Context) (a : VarId) : Arg := + ctx.rename.getD a (.var a) + +def Context.renameVar! (ctx : Context) (a : VarId) : VarId := + match ctx.rename[a]? with + | none => a + | some .erased => panic! "malformed IR" + | some (.var v) => v + +def Context.renameArg (ctx : Context) (a : Arg) : Arg := + match a with + | .erased => .erased + | .var v => ctx.rename.getD v a + +def Context.renameArgs (ctx : Context) (a : Array Arg) : Array Arg := + a.map ctx.renameArg + +def Context.insertRename (ctx : Context) (v : VarId) (a : Arg) : Context := + { ctx with rename := ctx.rename.insert v a } + +def Context.insert (var : VarId) (e : Entry) (ctx : Context) : Context := + { ctx with vars := ctx.vars.insert var e } + +def Context.insertIfNew (var : VarId) (e : Entry) (ctx : Context) : Context := + { ctx with vars := ctx.vars.insertIfNew var e } + +def Context.remove (ctx : Context) (v : VarId) : Context := + { ctx with vars := ctx.vars.erase v } + +def Context.addInstr (ctx : Context) (b : FnBody) : Context := + { ctx with instrs := ctx.instrs.push b } + +abbrev M := StateM Nat + +def mkFreshVar : M VarId := + modifyGet fun i => ({ idx := i }, i + 1) + +/-- +Returns whether RC operations need to be performed on a value of this type. This is restricted +version of `IRType.isPossibleRef` that also returns false for structs and unions that only consist +of scalars, e.g. `{uint8, tagged}`. +-/ +partial def needsRC (ty : IRType) : Bool := + match ty with + | .object | .tobject => true + | .union _ tys => tys.any needsRC + | .struct _ tys _ _ => tys.any needsRC + | _ => false + +@[inline] +def Entry.ofStruct (cidx : Nat) (ty : IRType) (rc : Int) : Entry := + match ty with + | .struct _ tys _ _ => + .struct cidx tys (Vector.replicate tys.size (.unbound rc)) + | _ => unreachable! + +def Entry.ofType (ty : IRType) (rc : Int) : Entry := + match ty with + | ty@(.struct _ tys _ _) => + if tys.any needsRC then + .ofStruct 0 ty rc + else + .persistent + | .union _ tys => + if tys.any needsRC then + .unknownUnion tys rc + else + .persistent + | .object => .ref false rc + | .tobject => .ref true rc + | _ => .persistent + +def Context.addVar (ctx : Context) (x : VarId) (ty : IRType) (rc : Int) : Context := + ctx.insertIfNew x (.ofType ty rc) + +/-- Add a variable entry if we need to (i.e. if `ty` is a struct type). -/ +def Context.addVarBasic (ctx : Context) (x : VarId) (ty : IRType) : Context := + match ty with + | ty@(.struct _ tys _ _) => + if tys.any needsRC then + ctx.insert x (.ofStruct 0 ty 0) + else + ctx.insert x .persistent + | .union _ tys => + if tys.any needsRC then + ctx.insert x (.unknownUnion tys 0) + else + ctx.insert x .persistent + | _ => ctx + +/-- +Given the entry `parentEntry` for `parent`, marks `proj : t` as the `idx`-th projection of `parent`. +-/ +@[inline] +def Context.addProjInfo (ctx : Context) (proj : VarId) (t : IRType) (idx : Nat) + (parent : VarId) (parentEntry : Entry) : Context := Id.run do + let : Inhabited Context := ⟨ctx⟩ + let .struct cidx tys fields := parentEntry | unreachable! + match fields[idx]! with + | .erased => + return (ctx.insert parent parentEntry).insertRename proj .erased + | .var v => + return (ctx.insert parent parentEntry).insertRename proj (.var v) + | .unbound rc => + let ctx := ctx.addVar proj t rc + let ctx := ctx.insert parent (.struct cidx tys (fields.set! idx (.var proj))) + ctx.addInstr (.vdecl proj t (.proj cidx idx parent) .nil) + +def Context.emitRCDiff (v : VarId) (check : Bool) (rc : Int) (ctx : Context) : Context := + if rc > 0 then + ctx.addInstr (.inc v rc.natAbs check false .nil) + else if rc < 0 then + ctx.addInstr (.dec v rc.natAbs check false .nil) + else + ctx + +partial def Context.accumulateRCDiff (v : VarId) (check : Bool) (rc : Int) (ctx : Context) : + Context := + match ctx.vars[v]? with + | none => ctx.emitRCDiff v check rc + | some .persistent => ctx + | some (.ref check' rc') => ctx.insert v (.ref (check && check') (rc + rc')) + | some (.unknownUnion tys rc') => ctx.insert v (.unknownUnion tys (rc + rc')) + | some (.struct cidx tys objs) => + let (objs, ctx) := Id.run <| StateT.run (s := ctx) <| objs.mapM fun + | .var v' => modifyGet fun ctx => (.var v', ctx.accumulateRCDiff v' true rc) + | .unbound rc' => modifyGet fun ctx => (.unbound (rc + rc'), ctx) + | .erased => return .erased + ctx.insert v (.struct cidx tys objs) + +/-- +Performs all necessary accumulated RC increments and decrements on `v`. If `ignoreInc` is true then +only decrements are performed and increments are ignored. +-/ +partial def Context.useVar (ctx : Context) (v : VarId) (ignoreInc : Bool := false) : M Context := do + match ctx.vars[v]? with + | none | some .persistent => return ctx + | some (.unknownUnion tys rc) => + if ignoreInc ∧ rc ≥ 0 then + return ctx + let ctx := ctx.emitRCDiff v (tys.any (!·.isDefiniteRef)) rc + let ctx := ctx.insert v (.unknownUnion tys 0) + return ctx + | some (.ref check rc) => + if ignoreInc ∧ rc ≥ 0 then + return ctx + let ctx := ctx.emitRCDiff v check rc + let ctx := ctx.insert v (.ref check 0) + return ctx + | some (.struct cidx tys objs) => + let mut ctx := ctx + let mut objs := objs + for h : i in *...tys.size do + let obj := objs[i] + let ty := tys[i] + match obj with + | .var v => + ctx ← ctx.useVar v ignoreInc + | .unbound rc => + if ignoreInc ∧ rc ≥ 0 then + continue + if rc ≠ 0 ∧ needsRC ty then + let var ← mkFreshVar + ctx := ctx.addInstr (.vdecl var ty (.proj cidx i v) .nil) + ctx := ctx.emitRCDiff var (!ty.isDefiniteRef) rc + ctx := ctx.addVar var ty 0 + objs := objs.set i (.var var) + | .erased => pure () + ctx := ctx.insert v (.struct cidx tys objs) + return ctx + +def Context.useArg (ctx : Context) (a : Arg) : M Context := + match a with + | .var v => ctx.useVar v + | .erased => return ctx + +def Context.addParams (params : Array Param) (ctx : Context) : Context := + params.foldl (init := ctx) fun ctx param => ctx.addVarBasic param.x param.ty + +/-- +Returns `ctx` but without any pending instructions or RC obligations. +-/ +def Context.resetRC (ctx : Context) : Context := + let vars := ctx.vars.map fun + | _, .unknownUnion tys _ => .unknownUnion tys 0 + | _, .ref c _ => .ref c 0 + | _, e => e + { vars, rename := ctx.rename } + +def Context.finish (ctx : Context) : M Context := do + let mut ctx := ctx + for (var, _) in ctx.vars do + ctx ← ctx.useVar var + return { ctx with vars := {} } + +def Context.addCtorKnowledge (ctx : Context) (v : VarId) (cidx : Nat) : Context := + match ctx.vars[v]? with + | some (.unknownUnion tys rc) => + ctx.insert v (.ofStruct cidx tys[cidx]! rc) + | _ => ctx + +def atConstructorIndex (ty : IRType) (i : Nat) : Array IRType := + match ty with + | .struct _ tys _ _ => tys + | .union _ tys => + match tys[i]! with + | .struct _ tys _ _ => tys + | _ => unreachable! + | _ => unreachable! + +def visitExpr (x : VarId) (t : IRType) (v : Expr) (ctx : Context) : M Context := do + match v with + | .proj c i y => + let y := ctx.renameVar! y + let v := .proj c i y + match ctx.vars[y]? with + | none => return ctx.addInstr (.vdecl x t v .nil) -- just an object projection, nothing to see here + | some .persistent => + return (ctx.insert x .persistent).addInstr (.vdecl x t v .nil) + | some (.unknownUnion tys rc) => + return ctx.addProjInfo x t i y (.ofStruct c tys[c]! rc) + | some e@(.struct cidx ..) => + if c ≠ cidx then + return (ctx.addInstr .unreachable).addInstr (.vdecl x t v .nil) + else + return ctx.addProjInfo x t i y e + | some (.ref _ _) => + return (ctx.addVarBasic x t).addInstr (.vdecl x t v .nil) + | .fap nm args => + let args := ctx.renameArgs args + let v := .fap nm args + if args.size = 0 then + return (ctx.insert x .persistent).addInstr (.vdecl x t v .nil) + else + return (← args.foldlM (·.useArg) (ctx.addVarBasic x t)).addInstr (.vdecl x t v .nil) + | .ap y args => + match ctx.renameVar y with + | .erased => + return ctx.insertRename x .erased + | .var y => + let args := ctx.renameArgs args + let v := .ap y args + return (← args.foldlM (·.useArg) (← ctx.useVar y)).addInstr (.vdecl x t v .nil) + | .pap nm args => + let args := ctx.renameArgs args + let v := .pap nm args + return (← args.foldlM (·.useArg) ctx).addInstr (.vdecl x t v .nil) + | .isShared y => + match ctx.renameVar y with + | .erased => return ctx.addInstr (.vdecl x t (.lit (.num 1)) .nil) + | .var y => + let v := .isShared y + return (← ctx.useVar y).addInstr (.vdecl x t v .nil) + | .reset n y => + let y := ctx.renameVar! y + let v := .reset n y + return (← ctx.useVar y).addInstr (.vdecl x t v .nil) + | .reuse y i u args => + let y := ctx.renameVar! y + let v := .reuse y i u args + return (← args.foldlM (·.useArg) (← ctx.useVar x)).addInstr (.vdecl x t v .nil) + | .box ty y => + let y := ctx.renameVar! y + let v := .box ty y + return (← (ctx.addVarBasic x t).useVar y).addInstr (.vdecl x t v .nil) + | .lit _ => return ctx.addInstr (.vdecl x t v .nil) + | .sproj c n o y => return ctx.addInstr (.vdecl x t (.sproj c n o (ctx.renameVar! y)) .nil) + | .uproj c i y => return ctx.addInstr (.vdecl x t (.uproj c i (ctx.renameVar! y)) .nil) + | .unbox y => return ctx.addInstr (.vdecl x t (.unbox (ctx.renameVar! y)) .nil) + | .ctor c args => + let args := ctx.renameArgs args + let v := .ctor c args + if t.isStruct then + let tys := atConstructorIndex t c.cidx + let e := .struct c.cidx tys <| Vector.ofFn fun ⟨i, _⟩ => + match args[i]! with + | .var v => .var v + | .erased => .erased + let ctx := tys.size.fold (init := ctx) fun i hi ctx => + match args[i]! with + | .var v => ctx.addVar v tys[i] 0 + | .erased => ctx + let ctx := ctx.insert x e + return ctx.addInstr (.vdecl x t v .nil) + else + return (← args.foldlM (·.useArg) ctx).addInstr (.vdecl x t v .nil) + +partial def visitFnBody (body : FnBody) (ctx : Context) : M FnBody := do + match body with + | .vdecl x t v b => + let ctx ← visitExpr x t v ctx + visitFnBody b ctx + | .jdecl j xs v b => + let v ← visitFnBody v (ctx.resetRC.addParams xs) + visitFnBody b (ctx.addInstr (.jdecl j xs v .nil)) + | .inc x n c p b => + if p then + -- increment on persistent value, ignore + visitFnBody b ctx + else + match ctx.renameVar x with + | .erased => visitFnBody b ctx + | .var x => + visitFnBody b (ctx.accumulateRCDiff x c n) + | .dec x n c p b => + if p then + -- decrement on persistent value, ignore + visitFnBody b ctx + else + match ctx.renameVar x with + | .erased => visitFnBody b ctx + | .var x => + let ctx := ctx.accumulateRCDiff x c (-n) + -- we can delay increments but we shouldn't delay deallocations + let ctx ← ctx.useVar x (ignoreInc := true) + visitFnBody b ctx + | .unreachable => return reshape ctx.instrs body + | .ret a => + let a := ctx.renameArg a + let ctx ← ctx.finish + return reshape ctx.instrs (.ret a) + | .jmp jp args => + let args := ctx.renameArgs args + let ctx ← ctx.finish + return reshape ctx.instrs (.jmp jp args) + | .case nm x xTy alts => + let x := ctx.renameVar! x + if let some (.struct cidx _ _) := ctx.vars[x]? then + -- because apparently this isn't guaranteed?! + let body? := alts.findSome? fun alt => + match alt with + | .ctor c b => if c.cidx = cidx then some b else none + | .default b => some b + match body? with + | none => return reshape ctx.instrs .unreachable + | some b => visitFnBody b ctx + else + let instrs := ctx.instrs + let ctx := { ctx with instrs := #[] } + let alts ← alts.mapM fun alt => do + match alt with + | .ctor c b => return .ctor c (← visitFnBody b (ctx.addCtorKnowledge x c.cidx)) + | .default b => return .default (← visitFnBody b ctx) + let body := .case nm x xTy alts + return reshape instrs body + | .del v b => + visitFnBody b (ctx.remove v |>.addInstr (.del v .nil)) + | .sset v c i o y t b => + let v := ctx.renameVar! v + let y := ctx.renameVar! y + let ctx ← ctx.useVar v + let ctx := ctx.addInstr (.sset v c i o y t .nil) + visitFnBody b ctx + | .uset v c i y b => + let v := ctx.renameVar! v + let y := ctx.renameVar! y + let ctx ← ctx.useVar v + let ctx := ctx.addInstr (.uset v c i y .nil) + visitFnBody b ctx + | .setTag v i b => + let v := ctx.renameVar! v + let ctx ← ctx.useVar v + let ctx := ctx.addInstr (.setTag v i .nil) + visitFnBody b ctx + | .set v i a b => + let v := ctx.renameVar! v + let a := ctx.renameArg a + let ctx ← ctx.useVar v + let ctx ← ctx.useArg a + let ctx := ctx.addInstr (.set v i a .nil) + visitFnBody b ctx + +def visitDecl (decl : Decl) : Decl := + match decl with + | .fdecl f xs t b i => + let b := (visitFnBody b (.addParams xs {})) |>.run' (decl.maxIndex + 1) |>.run + .fdecl f xs t b i + | .extern .. => decl + +end StructRC + +builtin_initialize registerTraceClass `compiler.ir.struct_rc (inherited := true) + +end Lean.IR diff --git a/src/Lean/Compiler/IR/ToIR.lean b/src/Lean/Compiler/IR/ToIR.lean index a3ec28a2ced6..d40c1b28e7b4 100644 --- a/src/Lean/Compiler/IR/ToIR.lean +++ b/src/Lean/Compiler/IR/ToIR.lean @@ -8,6 +8,8 @@ module prelude public import Lean.Compiler.IR.CompilerM public import Lean.Compiler.IR.ToIRType +import Lean.Compiler.IR.UnboxResult +import Lean.Compiler.IR.Boxing public section @@ -115,9 +117,9 @@ inductive TranslatedProj where def lowerProj (base : VarId) (ctorInfo : CtorInfo) (field : CtorFieldInfo) : TranslatedProj × IRType := match field with - | .object i irType => ⟨.expr (.proj i base), irType⟩ - | .usize i => ⟨.expr (.uproj i base), .usize⟩ - | .scalar _ offset irType => ⟨.expr (.sproj (ctorInfo.size + ctorInfo.usize) offset base), irType⟩ + | .object i irType => ⟨.expr (.proj ctorInfo.cidx i base), irType⟩ + | .usize i => ⟨.expr (.uproj ctorInfo.cidx i base), .usize⟩ + | .scalar _ offset irType => ⟨.expr (.sproj ctorInfo.cidx (ctorInfo.size + ctorInfo.usize) offset base), irType⟩ | .erased => ⟨.erased, .erased⟩ | .void => ⟨.erased, .void⟩ @@ -238,15 +240,20 @@ partial def lowerLet (decl : LCNF.LetDecl) (k : LCNF.Code) : M FnBody := do match fields[i]! with | .usize usizeIdx => let k ← loop (i + 1) - return .uset objVar usizeIdx varId k + return .uset objVar ctorInfo.cidx usizeIdx varId k | .scalar _ offset argType => let k ← loop (i + 1) - return .sset objVar (ctorInfo.size + ctorInfo.usize) offset varId argType k + return .sset objVar ctorInfo.cidx (ctorInfo.size + ctorInfo.usize) offset varId argType k | .object .. | .erased | .void => loop (i + 1) | some .erased => loop (i + 1) | none => lowerCode k loop 0 - return .vdecl objVar ctorInfo.type (.ctor ctorInfo objArgs) (← lowerNonObjectFields ()) + let type ← + if UnboxResult.hasUnboxAttr (← getEnv) ctorVal.induct then + toIRType decl.type + else + pure ctorInfo.type + return .vdecl objVar type (.ctor ctorInfo objArgs) (← lowerNonObjectFields ()) | some (.defnInfo ..) | some (.opaqueInfo ..) => mkFap name irArgs | some (.axiomInfo ..) | .some (.quotInfo ..) | .some (.inductInfo ..) | .some (.thmInfo ..) => @@ -360,6 +367,8 @@ def lowerDecl (d : LCNF.Decl) : M (Option Decl) := do -- TODO: This matches the behavior of the old compiler, but we should -- find a better way to handle this. addDecl (mkDummyExternDecl d.name params resultType) + if params.isEmpty && resultType.isStruct then + addDecl (ExplicitBoxing.boxedConstDecl d.name resultType) pure <| none else pure <| some <| .extern d.name params resultType externAttrData diff --git a/src/Lean/Compiler/IR/ToIRType.lean b/src/Lean/Compiler/IR/ToIRType.lean index 602c4509e17f..d8165e5a69e3 100644 --- a/src/Lean/Compiler/IR/ToIRType.lean +++ b/src/Lean/Compiler/IR/ToIRType.lean @@ -8,6 +8,7 @@ module prelude public import Lean.Compiler.IR.Format public import Lean.Compiler.LCNF.MonoTypes +import Lean.Compiler.IR.UnboxResult public section @@ -37,6 +38,35 @@ builtin_initialize trivialStructureInfoExt : LCNF.CacheExtension Name (Option LCNF.TrivialStructureInfo) ← LCNF.CacheExtension.register +inductive CtorFieldInfo where + | erased + | object (i : Nat) (type : IRType) + | usize (i : Nat) + | scalar (sz : Nat) (offset : Nat) (type : IRType) + | void + deriving Inhabited + +namespace CtorFieldInfo + +def format : CtorFieldInfo → Format + | erased => "◾" + | void => "void" + | object i type => f!"obj@{i}:{type}" + | usize i => f!"usize@{i}" + | scalar sz offset type => f!"scalar#{sz}@{offset}:{type}" + +instance : ToFormat CtorFieldInfo := ⟨format⟩ + +end CtorFieldInfo + +structure CtorLayout where + ctorInfo : CtorInfo + fieldInfo : Array CtorFieldInfo + deriving Inhabited + +builtin_initialize ctorLayoutExt : LCNF.CacheExtension Name CtorLayout ← + LCNF.CacheExtension.register + /-- The idea of this function is the same as in `ToMono`, however the notion of "irrelevancy" has changed because we now have the `void` type which can only be erased in impure context and thus at @@ -102,6 +132,8 @@ private def isAnyProducingType (type : Lean.Expr) : Bool := | .forallE _ _ b _ => isAnyProducingType b | _ => false +mutual + partial def toIRType (type : Lean.Expr) : CoreM IRType := do match type with | .const name _ => visitApp name #[] @@ -126,39 +158,47 @@ where let ctorType ← LCNF.getOtherDeclBaseType info.ctorName [] let monoType ← LCNF.toMonoType (LCNF.getParamTypes (← LCNF.instantiateForall ctorType args[*...info.numParams]))[info.fieldIdx]! toIRType monoType + else if UnboxResult.hasUnboxAttr (← getEnv) declName then + unboxedIRType declName args else nameToIRType declName -inductive CtorFieldInfo where - | erased - | object (i : Nat) (type : IRType) - | usize (i : Nat) - | scalar (sz : Nat) (offset : Nat) (type : IRType) - | void - deriving Inhabited - -namespace CtorFieldInfo - -def format : CtorFieldInfo → Format - | erased => "◾" - | void => "void" - | object i type => f!"obj@{i}:{type}" - | usize i => f!"usize@{i}" - | scalar sz offset type => f!"scalar#{sz}@{offset}:{type}" - -instance : ToFormat CtorFieldInfo := ⟨format⟩ - -end CtorFieldInfo - -structure CtorLayout where - ctorInfo : CtorInfo - fieldInfo : Array CtorFieldInfo - deriving Inhabited - -builtin_initialize ctorLayoutExt : LCNF.CacheExtension Name CtorLayout ← - LCNF.CacheExtension.register +partial def unboxedIRType (declName : Name) (args : Array Lean.Expr) : CoreM IRType := do + let induct ← getConstInfoInduct declName + if args.size < induct.numParams then + throwError "too few parameters for type {declName}: \ + expected {induct.numParams} but got only {args.size}" + let params := args.extract 0 induct.numParams + if let [] := induct.ctors then + return .erased -- really, an "impossible" type but erased is good enough + if let [ctor] := induct.ctors then + return ← handleCtor declName ctor params + let mut types : Array IRType := #[] + for ctor in induct.ctors do + types := types.push (← handleCtor none ctor params) + return .union declName types +where + /-- `induct` is specified if this is the only constructor -/ + handleCtor (induct : Option Name) (ctorName : Name) + (params : Array Lean.Expr) : CoreM IRType := do + let ctorLayout ← getCtorLayout ctorName + let type := (← getConstVal ctorName).type + Meta.MetaM.run' <| Meta.forallTelescope type fun vars _ => do + let mut objectTypes : Array IRType := Array.replicate ctorLayout.ctorInfo.size default + let paramVars := vars.extract 0 params.size + for h : i in *...ctorLayout.fieldInfo.size do + let info := ctorLayout.fieldInfo[i] + let .object j _ := info | continue + let field := vars[i + params.size]! + let fieldType ← field.fvarId!.getType + let lcnfFieldType ← LCNF.toLCNFType fieldType + let lcnfFieldType := lcnfFieldType.replaceFVars paramVars params + let monoFieldType ← LCNF.toMonoType lcnfFieldType + let irFieldType ← toIRType monoFieldType + objectTypes := objectTypes.set! j irFieldType + return .struct induct objectTypes ctorLayout.ctorInfo.usize ctorLayout.ctorInfo.ssize -def getCtorLayout (ctorName : Name) : CoreM CtorLayout := do +partial def getCtorLayout (ctorName : Name) : CoreM CtorLayout := do match (← ctorLayoutExt.find? ctorName) with | some info => return info | none => @@ -180,7 +220,7 @@ where fillCache := do let monoFieldType ← LCNF.toMonoType lcnfFieldType let irFieldType ← toIRType monoFieldType let ctorField ← match irFieldType with - | .object | .tagged | .tobject => do + | .object | .tagged | .tobject | .struct .. | .union .. => do let i := nextIdx nextIdx := nextIdx + 1 pure <| .object i irFieldType @@ -205,7 +245,6 @@ where fillCache := do | .float => has8BScalar := true .pure <| .scalar 8 0 .float - | .struct .. | .union .. => unreachable! fields := fields.push ctorField let numObjs := nextIdx ⟨fields, nextIdx⟩ := Id.run <| StateT.run (s := nextIdx) <| fields.mapM fun field => do @@ -246,5 +285,7 @@ where fillCache := do fieldInfo := fields } +end + end IR end Lean diff --git a/src/library/ir_interpreter.cpp b/src/library/ir_interpreter.cpp index 321e630f0480..a60fa5d22519 100644 --- a/src/library/ir_interpreter.cpp +++ b/src/library/ir_interpreter.cpp @@ -63,8 +63,7 @@ typedef object_ref lit_val; typedef object_ref ctor_info; type to_type(object * obj) { - if (!is_scalar(obj)) throw exception("unsupported IRType"); - else return static_cast(unbox(obj)); + return static_cast(lean_obj_tag(obj)); } type cnstr_get_type(object_ref const & o, unsigned i) { return to_type(cnstr_get(o.raw(), i)); } @@ -97,13 +96,13 @@ var_id const & expr_reuse_obj(expr const & e) { lean_assert(expr_tag(e) == expr_ ctor_info const & expr_reuse_ctor(expr const & e) { lean_assert(expr_tag(e) == expr_kind::Reuse); return cnstr_get_ref_t(e, 1); } bool expr_reuse_update_header(expr const & e) { lean_assert(expr_tag(e) == expr_kind::Reuse); return get_bool_field(e.raw(), 3); } array_ref const & expr_reuse_args(expr const & e) { lean_assert(expr_tag(e) == expr_kind::Reuse); return cnstr_get_ref_t>(e, 2); } -nat const & expr_proj_idx(expr const & e) { lean_assert(expr_tag(e) == expr_kind::Proj); return cnstr_get_ref_t(e, 0); } -var_id const & expr_proj_obj(expr const & e) { lean_assert(expr_tag(e) == expr_kind::Proj); return cnstr_get_ref_t(e, 1); } -nat const & expr_uproj_idx(expr const & e) { lean_assert(expr_tag(e) == expr_kind::UProj); return cnstr_get_ref_t(e, 0); } -var_id const & expr_uproj_obj(expr const & e) { lean_assert(expr_tag(e) == expr_kind::UProj); return cnstr_get_ref_t(e, 1); } -nat const & expr_sproj_idx(expr const & e) { lean_assert(expr_tag(e) == expr_kind::SProj); return cnstr_get_ref_t(e, 0); } -nat const & expr_sproj_offset(expr const & e) { lean_assert(expr_tag(e) == expr_kind::SProj); return cnstr_get_ref_t(e, 1); } -var_id const & expr_sproj_obj(expr const & e) { lean_assert(expr_tag(e) == expr_kind::SProj); return cnstr_get_ref_t(e, 2); } +nat const & expr_proj_idx(expr const & e) { lean_assert(expr_tag(e) == expr_kind::Proj); return cnstr_get_ref_t(e, 1); } +var_id const & expr_proj_obj(expr const & e) { lean_assert(expr_tag(e) == expr_kind::Proj); return cnstr_get_ref_t(e, 2); } +nat const & expr_uproj_idx(expr const & e) { lean_assert(expr_tag(e) == expr_kind::UProj); return cnstr_get_ref_t(e, 1); } +var_id const & expr_uproj_obj(expr const & e) { lean_assert(expr_tag(e) == expr_kind::UProj); return cnstr_get_ref_t(e, 2); } +nat const & expr_sproj_idx(expr const & e) { lean_assert(expr_tag(e) == expr_kind::SProj); return cnstr_get_ref_t(e, 1); } +nat const & expr_sproj_offset(expr const & e) { lean_assert(expr_tag(e) == expr_kind::SProj); return cnstr_get_ref_t(e, 2); } +var_id const & expr_sproj_obj(expr const & e) { lean_assert(expr_tag(e) == expr_kind::SProj); return cnstr_get_ref_t(e, 3); } fun_id const & expr_fap_fun(expr const & e) { lean_assert(expr_tag(e) == expr_kind::FAp); return cnstr_get_ref_t(e, 0); } array_ref const & expr_fap_args(expr const & e) { lean_assert(expr_tag(e) == expr_kind::FAp); return cnstr_get_ref_t>(e, 1); } fun_id const & expr_pap_fun(expr const & e) { lean_assert(expr_tag(e) == expr_kind::PAp); return cnstr_get_ref_t(e, 0); } @@ -133,6 +132,7 @@ enum class fn_body_kind { VDecl, JDecl, Set, SetTag, USet, SSet, Inc, Dec, Del, fn_body_kind fn_body_tag(fn_body const & a) { return is_scalar(a.raw()) ? static_cast(unbox(a.raw())) : static_cast(cnstr_tag(a.raw())); } var_id const & fn_body_vdecl_var(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::VDecl); return cnstr_get_ref_t(b, 0); } type fn_body_vdecl_type(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::VDecl); return cnstr_get_type(b, 1); } +object_ref const & fn_body_vdecl_type_ref(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::VDecl); return cnstr_get_ref(b, 1); } expr const & fn_body_vdecl_expr(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::VDecl); return cnstr_get_ref_t(b, 2); } fn_body const & fn_body_vdecl_cont(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::VDecl); return cnstr_get_ref_t(b, 3); } jp_id const & fn_body_jdecl_id(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::JDecl); return cnstr_get_ref_t(b, 0); } @@ -147,15 +147,15 @@ var_id const & fn_body_set_tag_var(fn_body const & b) { lean_assert(fn_body_tag( nat const & fn_body_set_tag_cidx(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::SetTag); return cnstr_get_ref_t(b, 1); } fn_body const & fn_body_set_tag_cont(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::SetTag); return cnstr_get_ref_t(b, 2); } var_id const & fn_body_uset_target(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::USet); return cnstr_get_ref_t(b, 0); } -nat const & fn_body_uset_idx(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::USet); return cnstr_get_ref_t(b, 1); } -var_id const & fn_body_uset_source(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::USet); return cnstr_get_ref_t(b, 2); } -fn_body const & fn_body_uset_cont(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::USet); return cnstr_get_ref_t(b, 3); } +nat const & fn_body_uset_idx(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::USet); return cnstr_get_ref_t(b, 2); } +var_id const & fn_body_uset_source(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::USet); return cnstr_get_ref_t(b, 3); } +fn_body const & fn_body_uset_cont(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::USet); return cnstr_get_ref_t(b, 4); } var_id const & fn_body_sset_target(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::SSet); return cnstr_get_ref_t(b, 0); } -nat const & fn_body_sset_idx(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::SSet); return cnstr_get_ref_t(b, 1); } -nat const & fn_body_sset_offset(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::SSet); return cnstr_get_ref_t(b, 2); } -var_id const & fn_body_sset_source(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::SSet); return cnstr_get_ref_t(b, 3); } -type fn_body_sset_type(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::SSet); return cnstr_get_type(b, 4); } -fn_body const & fn_body_sset_cont(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::SSet); return cnstr_get_ref_t(b, 5); } +nat const & fn_body_sset_idx(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::SSet); return cnstr_get_ref_t(b, 2); } +nat const & fn_body_sset_offset(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::SSet); return cnstr_get_ref_t(b, 3); } +var_id const & fn_body_sset_source(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::SSet); return cnstr_get_ref_t(b, 4); } +type fn_body_sset_type(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::SSet); return cnstr_get_type(b, 5); } +fn_body const & fn_body_sset_cont(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::SSet); return cnstr_get_ref_t(b, 6); } var_id const & fn_body_inc_var(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::Inc); return cnstr_get_ref_t(b, 0); } nat const & fn_body_inc_val(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::Inc); return cnstr_get_ref_t(b, 1); } bool fn_body_inc_maybe_scalar(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::Inc); return get_bool_field(b.raw(), 3); } @@ -221,8 +221,8 @@ std::string format_fn_body_head(fn_body const & b) { } static bool type_is_scalar(type t) { - return t != type::Object && t != type::Tagged && t != type::TObject && t != type::Irrelevant - && t != type::Void; + return t == type::Float || t == type::Float32 || t == type::UInt8 || t == type::UInt16 || + t == type::UInt32 || t == type::UInt64 || t == type::USize; } extern "C" object* lean_get_regular_init_fn_name_for(object* env, object* fn); @@ -277,10 +277,11 @@ object * box_t(value v, type t) { case type::TObject: case type::Irrelevant: case type::Void: - return v.m_obj; + // Note: structs and unions are not supported by the interpreter + // and thus are passed around in boxed form, i.e. unbox and box are no-ops case type::Struct: case type::Union: - throw exception("not implemented yet"); + return v.m_obj; } lean_unreachable(); } @@ -302,7 +303,9 @@ value unbox_t(object * o, type t) { break; case type::Struct: case type::Union: - throw exception("not implemented yet"); + // Note: structs and unions are not supported by the interpreter + // and thus are passed around in boxed form, i.e. unbox and box are no-ops + return value(o); } lean_unreachable(); } @@ -443,7 +446,7 @@ class interpreter { } /** \brief Allocate constructor object with given tag and arguments */ - object * alloc_ctor(ctor_info const & i, array_ref const & args) { + object * alloc_ctor(ctor_info const & i, array_ref const & args, object_ref const & tref) { size_t tag = ctor_info_tag(i).get_small_value(); // number of boxed object fields size_t size = ctor_info_size(i).get_small_value(); @@ -456,8 +459,27 @@ class interpreter { return box(tag); } else { object *o = alloc_cnstr(tag, size, usize * sizeof(void *) + ssize); - for (size_t i = 0; i < args.size(); i++) { - cnstr_set(o, i, eval_arg(args[i]).m_obj); + if (is_scalar(tref.raw())) { + for (size_t i = 0; i < args.size(); i++) { + cnstr_set(o, i, eval_arg(args[i]).m_obj); + } + } else { + // struct or union + object_ref const * t_ref = &tref; + if (to_type(tref.raw()) == type::Union) { + array_ref const & types = cnstr_get_ref_t>(tref, 1); + lean_assert(tag < types.size()); + t_ref = &types[tag]; + } + lean_assert(to_type(t_ref->raw()) == type::Struct); + array_ref const & types = cnstr_get_ref_t>(*t_ref, 1); + lean_assert(types.size() == args.size()); + for (size_t i = 0; i < args.size(); i++) { + type t = to_type(types[i].raw()); + value arg = eval_arg(args[i]); + object * val = box_t(arg, t); + cnstr_set(o, i, val); + } } return o; } @@ -476,10 +498,11 @@ class interpreter { return cls; } - value eval_expr(expr const & e, type t) { + value eval_expr(expr const & e, object_ref const & tref) { + type t = to_type(tref.raw()); switch (expr_tag(e)) { case expr_kind::Ctor: - return value { alloc_ctor(expr_ctor_info(e), expr_ctor_args(e)) }; + return value { alloc_ctor(expr_ctor_info(e), expr_ctor_args(e), tref) }; case expr_kind::Reset: { // release fields if unique reference in preparation for `Reuse` below object * o = var(expr_reset_obj(e)).m_obj; if (is_exclusive(o)) { @@ -497,7 +520,7 @@ class interpreter { // check if `Reset` above had a unique reference it consumed if (is_scalar(o)) { // fall back to regular allocation - return alloc_ctor(expr_reuse_ctor(e), expr_reuse_args(e)); + return alloc_ctor(expr_reuse_ctor(e), expr_reuse_args(e), tref); } else { // create new constructor object in-place if (expr_reuse_update_header(e)) { @@ -510,6 +533,20 @@ class interpreter { } } case expr_kind::Proj: // object field access + if (LEAN_UNLIKELY(type_is_scalar(t))) { + // possible for unboxed structs + object * field = cnstr_get(var(expr_proj_obj(e)).m_obj, expr_proj_idx(e).get_small_value()); + switch (t) { + case type::Float: return value::from_float(lean_unbox_float(field)); + case type::Float32: return value::from_float32(lean_unbox_float32(field)); + case type::UInt8: return lean_unbox(field); + case type::UInt16: return lean_unbox(field); + case type::UInt32: return lean_unbox_uint32(field); + case type::UInt64: return lean_unbox_uint64(field); + case type::USize: return lean_unbox_usize(field); + default: lean_unreachable(); + } + } return cnstr_get(var(expr_proj_obj(e)).m_obj, expr_proj_idx(e).get_small_value()); case expr_kind::UProj: // USize field access return cnstr_get_usize(var(expr_uproj_obj(e)).m_obj, expr_uproj_idx(e).get_small_value()); @@ -664,7 +701,7 @@ class interpreter { check_system(); break; } - value v = eval_expr(fn_body_vdecl_expr(b), fn_body_vdecl_type(b)); + value v = eval_expr(fn_body_vdecl_expr(b), fn_body_vdecl_type_ref(b)); // NOTE: `var` must be called *after* `eval_expr` because the stack may get resized and invalidate // the pointer var(fn_body_vdecl_var(b)) = v; @@ -841,6 +878,8 @@ class interpreter { symbol_cache_entry e_new { get_decl(fn), {nullptr, false} }; if (m_prefer_native || decl_tag(e_new.m_decl) == decl_kind::Extern || has_init_attribute(m_env, fn)) { string_ref mangled = get_symbol_stem(m_env, fn); + // TODO: this does not always cause the right mangled name to be generated + // compare a_ -> l_a__ and a_._boxed -> l_a___00__boxed string_ref boxed_mangled(string_append(mangled.to_obj_arg(), g_boxed_mangled_suffix->raw())); // check for boxed version first if (void *p_boxed = lookup_symbol_in_cur_exe(boxed_mangled.data())) { @@ -891,8 +930,6 @@ class interpreter { symbol_cache_entry e = lookup_symbol(fn); if (e.m_native.m_addr) { // we can assume that all native code has been initialized (see e.g. `evalConst`) - - // constants do not have boxed wrappers, but we'll survive switch (t) { case type::Float: return value::from_float(*static_cast(e.m_native.m_addr)); case type::Float32: return value::from_float32(*static_cast(e.m_native.m_addr)); @@ -909,7 +946,9 @@ class interpreter { return *static_cast(e.m_native.m_addr); case type::Struct: case type::Union: - throw exception("not implemented yet"); + // we generate boxed wrappers specifically for struct/union constants + lean_assert(e.m_native.m_boxed); + return *static_cast(e.m_native.m_addr); } } diff --git a/stage0/src/library/ir_interpreter.cpp b/stage0/src/library/ir_interpreter.cpp index 321e630f0480..a60fa5d22519 100644 --- a/stage0/src/library/ir_interpreter.cpp +++ b/stage0/src/library/ir_interpreter.cpp @@ -63,8 +63,7 @@ typedef object_ref lit_val; typedef object_ref ctor_info; type to_type(object * obj) { - if (!is_scalar(obj)) throw exception("unsupported IRType"); - else return static_cast(unbox(obj)); + return static_cast(lean_obj_tag(obj)); } type cnstr_get_type(object_ref const & o, unsigned i) { return to_type(cnstr_get(o.raw(), i)); } @@ -97,13 +96,13 @@ var_id const & expr_reuse_obj(expr const & e) { lean_assert(expr_tag(e) == expr_ ctor_info const & expr_reuse_ctor(expr const & e) { lean_assert(expr_tag(e) == expr_kind::Reuse); return cnstr_get_ref_t(e, 1); } bool expr_reuse_update_header(expr const & e) { lean_assert(expr_tag(e) == expr_kind::Reuse); return get_bool_field(e.raw(), 3); } array_ref const & expr_reuse_args(expr const & e) { lean_assert(expr_tag(e) == expr_kind::Reuse); return cnstr_get_ref_t>(e, 2); } -nat const & expr_proj_idx(expr const & e) { lean_assert(expr_tag(e) == expr_kind::Proj); return cnstr_get_ref_t(e, 0); } -var_id const & expr_proj_obj(expr const & e) { lean_assert(expr_tag(e) == expr_kind::Proj); return cnstr_get_ref_t(e, 1); } -nat const & expr_uproj_idx(expr const & e) { lean_assert(expr_tag(e) == expr_kind::UProj); return cnstr_get_ref_t(e, 0); } -var_id const & expr_uproj_obj(expr const & e) { lean_assert(expr_tag(e) == expr_kind::UProj); return cnstr_get_ref_t(e, 1); } -nat const & expr_sproj_idx(expr const & e) { lean_assert(expr_tag(e) == expr_kind::SProj); return cnstr_get_ref_t(e, 0); } -nat const & expr_sproj_offset(expr const & e) { lean_assert(expr_tag(e) == expr_kind::SProj); return cnstr_get_ref_t(e, 1); } -var_id const & expr_sproj_obj(expr const & e) { lean_assert(expr_tag(e) == expr_kind::SProj); return cnstr_get_ref_t(e, 2); } +nat const & expr_proj_idx(expr const & e) { lean_assert(expr_tag(e) == expr_kind::Proj); return cnstr_get_ref_t(e, 1); } +var_id const & expr_proj_obj(expr const & e) { lean_assert(expr_tag(e) == expr_kind::Proj); return cnstr_get_ref_t(e, 2); } +nat const & expr_uproj_idx(expr const & e) { lean_assert(expr_tag(e) == expr_kind::UProj); return cnstr_get_ref_t(e, 1); } +var_id const & expr_uproj_obj(expr const & e) { lean_assert(expr_tag(e) == expr_kind::UProj); return cnstr_get_ref_t(e, 2); } +nat const & expr_sproj_idx(expr const & e) { lean_assert(expr_tag(e) == expr_kind::SProj); return cnstr_get_ref_t(e, 1); } +nat const & expr_sproj_offset(expr const & e) { lean_assert(expr_tag(e) == expr_kind::SProj); return cnstr_get_ref_t(e, 2); } +var_id const & expr_sproj_obj(expr const & e) { lean_assert(expr_tag(e) == expr_kind::SProj); return cnstr_get_ref_t(e, 3); } fun_id const & expr_fap_fun(expr const & e) { lean_assert(expr_tag(e) == expr_kind::FAp); return cnstr_get_ref_t(e, 0); } array_ref const & expr_fap_args(expr const & e) { lean_assert(expr_tag(e) == expr_kind::FAp); return cnstr_get_ref_t>(e, 1); } fun_id const & expr_pap_fun(expr const & e) { lean_assert(expr_tag(e) == expr_kind::PAp); return cnstr_get_ref_t(e, 0); } @@ -133,6 +132,7 @@ enum class fn_body_kind { VDecl, JDecl, Set, SetTag, USet, SSet, Inc, Dec, Del, fn_body_kind fn_body_tag(fn_body const & a) { return is_scalar(a.raw()) ? static_cast(unbox(a.raw())) : static_cast(cnstr_tag(a.raw())); } var_id const & fn_body_vdecl_var(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::VDecl); return cnstr_get_ref_t(b, 0); } type fn_body_vdecl_type(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::VDecl); return cnstr_get_type(b, 1); } +object_ref const & fn_body_vdecl_type_ref(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::VDecl); return cnstr_get_ref(b, 1); } expr const & fn_body_vdecl_expr(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::VDecl); return cnstr_get_ref_t(b, 2); } fn_body const & fn_body_vdecl_cont(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::VDecl); return cnstr_get_ref_t(b, 3); } jp_id const & fn_body_jdecl_id(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::JDecl); return cnstr_get_ref_t(b, 0); } @@ -147,15 +147,15 @@ var_id const & fn_body_set_tag_var(fn_body const & b) { lean_assert(fn_body_tag( nat const & fn_body_set_tag_cidx(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::SetTag); return cnstr_get_ref_t(b, 1); } fn_body const & fn_body_set_tag_cont(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::SetTag); return cnstr_get_ref_t(b, 2); } var_id const & fn_body_uset_target(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::USet); return cnstr_get_ref_t(b, 0); } -nat const & fn_body_uset_idx(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::USet); return cnstr_get_ref_t(b, 1); } -var_id const & fn_body_uset_source(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::USet); return cnstr_get_ref_t(b, 2); } -fn_body const & fn_body_uset_cont(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::USet); return cnstr_get_ref_t(b, 3); } +nat const & fn_body_uset_idx(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::USet); return cnstr_get_ref_t(b, 2); } +var_id const & fn_body_uset_source(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::USet); return cnstr_get_ref_t(b, 3); } +fn_body const & fn_body_uset_cont(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::USet); return cnstr_get_ref_t(b, 4); } var_id const & fn_body_sset_target(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::SSet); return cnstr_get_ref_t(b, 0); } -nat const & fn_body_sset_idx(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::SSet); return cnstr_get_ref_t(b, 1); } -nat const & fn_body_sset_offset(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::SSet); return cnstr_get_ref_t(b, 2); } -var_id const & fn_body_sset_source(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::SSet); return cnstr_get_ref_t(b, 3); } -type fn_body_sset_type(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::SSet); return cnstr_get_type(b, 4); } -fn_body const & fn_body_sset_cont(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::SSet); return cnstr_get_ref_t(b, 5); } +nat const & fn_body_sset_idx(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::SSet); return cnstr_get_ref_t(b, 2); } +nat const & fn_body_sset_offset(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::SSet); return cnstr_get_ref_t(b, 3); } +var_id const & fn_body_sset_source(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::SSet); return cnstr_get_ref_t(b, 4); } +type fn_body_sset_type(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::SSet); return cnstr_get_type(b, 5); } +fn_body const & fn_body_sset_cont(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::SSet); return cnstr_get_ref_t(b, 6); } var_id const & fn_body_inc_var(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::Inc); return cnstr_get_ref_t(b, 0); } nat const & fn_body_inc_val(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::Inc); return cnstr_get_ref_t(b, 1); } bool fn_body_inc_maybe_scalar(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::Inc); return get_bool_field(b.raw(), 3); } @@ -221,8 +221,8 @@ std::string format_fn_body_head(fn_body const & b) { } static bool type_is_scalar(type t) { - return t != type::Object && t != type::Tagged && t != type::TObject && t != type::Irrelevant - && t != type::Void; + return t == type::Float || t == type::Float32 || t == type::UInt8 || t == type::UInt16 || + t == type::UInt32 || t == type::UInt64 || t == type::USize; } extern "C" object* lean_get_regular_init_fn_name_for(object* env, object* fn); @@ -277,10 +277,11 @@ object * box_t(value v, type t) { case type::TObject: case type::Irrelevant: case type::Void: - return v.m_obj; + // Note: structs and unions are not supported by the interpreter + // and thus are passed around in boxed form, i.e. unbox and box are no-ops case type::Struct: case type::Union: - throw exception("not implemented yet"); + return v.m_obj; } lean_unreachable(); } @@ -302,7 +303,9 @@ value unbox_t(object * o, type t) { break; case type::Struct: case type::Union: - throw exception("not implemented yet"); + // Note: structs and unions are not supported by the interpreter + // and thus are passed around in boxed form, i.e. unbox and box are no-ops + return value(o); } lean_unreachable(); } @@ -443,7 +446,7 @@ class interpreter { } /** \brief Allocate constructor object with given tag and arguments */ - object * alloc_ctor(ctor_info const & i, array_ref const & args) { + object * alloc_ctor(ctor_info const & i, array_ref const & args, object_ref const & tref) { size_t tag = ctor_info_tag(i).get_small_value(); // number of boxed object fields size_t size = ctor_info_size(i).get_small_value(); @@ -456,8 +459,27 @@ class interpreter { return box(tag); } else { object *o = alloc_cnstr(tag, size, usize * sizeof(void *) + ssize); - for (size_t i = 0; i < args.size(); i++) { - cnstr_set(o, i, eval_arg(args[i]).m_obj); + if (is_scalar(tref.raw())) { + for (size_t i = 0; i < args.size(); i++) { + cnstr_set(o, i, eval_arg(args[i]).m_obj); + } + } else { + // struct or union + object_ref const * t_ref = &tref; + if (to_type(tref.raw()) == type::Union) { + array_ref const & types = cnstr_get_ref_t>(tref, 1); + lean_assert(tag < types.size()); + t_ref = &types[tag]; + } + lean_assert(to_type(t_ref->raw()) == type::Struct); + array_ref const & types = cnstr_get_ref_t>(*t_ref, 1); + lean_assert(types.size() == args.size()); + for (size_t i = 0; i < args.size(); i++) { + type t = to_type(types[i].raw()); + value arg = eval_arg(args[i]); + object * val = box_t(arg, t); + cnstr_set(o, i, val); + } } return o; } @@ -476,10 +498,11 @@ class interpreter { return cls; } - value eval_expr(expr const & e, type t) { + value eval_expr(expr const & e, object_ref const & tref) { + type t = to_type(tref.raw()); switch (expr_tag(e)) { case expr_kind::Ctor: - return value { alloc_ctor(expr_ctor_info(e), expr_ctor_args(e)) }; + return value { alloc_ctor(expr_ctor_info(e), expr_ctor_args(e), tref) }; case expr_kind::Reset: { // release fields if unique reference in preparation for `Reuse` below object * o = var(expr_reset_obj(e)).m_obj; if (is_exclusive(o)) { @@ -497,7 +520,7 @@ class interpreter { // check if `Reset` above had a unique reference it consumed if (is_scalar(o)) { // fall back to regular allocation - return alloc_ctor(expr_reuse_ctor(e), expr_reuse_args(e)); + return alloc_ctor(expr_reuse_ctor(e), expr_reuse_args(e), tref); } else { // create new constructor object in-place if (expr_reuse_update_header(e)) { @@ -510,6 +533,20 @@ class interpreter { } } case expr_kind::Proj: // object field access + if (LEAN_UNLIKELY(type_is_scalar(t))) { + // possible for unboxed structs + object * field = cnstr_get(var(expr_proj_obj(e)).m_obj, expr_proj_idx(e).get_small_value()); + switch (t) { + case type::Float: return value::from_float(lean_unbox_float(field)); + case type::Float32: return value::from_float32(lean_unbox_float32(field)); + case type::UInt8: return lean_unbox(field); + case type::UInt16: return lean_unbox(field); + case type::UInt32: return lean_unbox_uint32(field); + case type::UInt64: return lean_unbox_uint64(field); + case type::USize: return lean_unbox_usize(field); + default: lean_unreachable(); + } + } return cnstr_get(var(expr_proj_obj(e)).m_obj, expr_proj_idx(e).get_small_value()); case expr_kind::UProj: // USize field access return cnstr_get_usize(var(expr_uproj_obj(e)).m_obj, expr_uproj_idx(e).get_small_value()); @@ -664,7 +701,7 @@ class interpreter { check_system(); break; } - value v = eval_expr(fn_body_vdecl_expr(b), fn_body_vdecl_type(b)); + value v = eval_expr(fn_body_vdecl_expr(b), fn_body_vdecl_type_ref(b)); // NOTE: `var` must be called *after* `eval_expr` because the stack may get resized and invalidate // the pointer var(fn_body_vdecl_var(b)) = v; @@ -841,6 +878,8 @@ class interpreter { symbol_cache_entry e_new { get_decl(fn), {nullptr, false} }; if (m_prefer_native || decl_tag(e_new.m_decl) == decl_kind::Extern || has_init_attribute(m_env, fn)) { string_ref mangled = get_symbol_stem(m_env, fn); + // TODO: this does not always cause the right mangled name to be generated + // compare a_ -> l_a__ and a_._boxed -> l_a___00__boxed string_ref boxed_mangled(string_append(mangled.to_obj_arg(), g_boxed_mangled_suffix->raw())); // check for boxed version first if (void *p_boxed = lookup_symbol_in_cur_exe(boxed_mangled.data())) { @@ -891,8 +930,6 @@ class interpreter { symbol_cache_entry e = lookup_symbol(fn); if (e.m_native.m_addr) { // we can assume that all native code has been initialized (see e.g. `evalConst`) - - // constants do not have boxed wrappers, but we'll survive switch (t) { case type::Float: return value::from_float(*static_cast(e.m_native.m_addr)); case type::Float32: return value::from_float32(*static_cast(e.m_native.m_addr)); @@ -909,7 +946,9 @@ class interpreter { return *static_cast(e.m_native.m_addr); case type::Struct: case type::Union: - throw exception("not implemented yet"); + // we generate boxed wrappers specifically for struct/union constants + lean_assert(e.m_native.m_boxed); + return *static_cast(e.m_native.m_addr); } } diff --git a/stage0/stdlib/Lake/Config/Meta.c b/stage0/stdlib/Lake/Config/Meta.c index 497948082f4f..e3d99e8e7397 100644 --- a/stage0/stdlib/Lake/Config/Meta.c +++ b/stage0/stdlib/Lake/Config/Meta.c @@ -12488,10 +12488,10 @@ lean_object* x_46; x_46 = l___private_Lake_Config_Meta_0__Lake_mkParentFieldView___lam__0(x_44); lean_dec(x_44); x_27 = x_39; -x_28 = x_41; -x_29 = x_40; -x_30 = x_42; -x_31 = x_43; +x_28 = x_42; +x_29 = x_43; +x_30 = x_40; +x_31 = x_41; x_32 = x_46; goto block_35; } @@ -12509,10 +12509,10 @@ lean_dec(x_49); lean_ctor_set(x_47, 0, x_50); x_51 = l_Lean_MacroScopesView_review(x_47); x_27 = x_39; -x_28 = x_41; -x_29 = x_40; -x_30 = x_42; -x_31 = x_43; +x_28 = x_42; +x_29 = x_43; +x_30 = x_40; +x_31 = x_41; x_32 = x_51; goto block_35; } @@ -12537,10 +12537,10 @@ lean_ctor_set(x_57, 2, x_54); lean_ctor_set(x_57, 3, x_55); x_58 = l_Lean_MacroScopesView_review(x_57); x_27 = x_39; -x_28 = x_41; -x_29 = x_40; -x_30 = x_42; -x_31 = x_43; +x_28 = x_42; +x_29 = x_43; +x_30 = x_40; +x_31 = x_41; x_32 = x_58; goto block_35; } @@ -12741,13 +12741,13 @@ return x_25; { uint8_t x_33; lean_object* x_34; x_33 = 0; -x_34 = l_Lean_mkIdentFrom(x_28, x_32, x_33); -lean_dec(x_28); +x_34 = l_Lean_mkIdentFrom(x_31, x_32, x_33); +lean_dec(x_31); x_8 = x_27; -x_9 = x_29; +x_9 = x_30; x_10 = x_34; -x_11 = x_30; -x_12 = x_31; +x_11 = x_28; +x_12 = x_29; goto block_26; } } @@ -12844,10 +12844,10 @@ lean_object* x_146; x_146 = l___private_Lake_Config_Meta_0__Lake_mkParentFieldView___lam__0(x_144); lean_dec(x_144); x_126 = x_139; -x_127 = x_141; -x_128 = x_140; -x_129 = x_142; -x_130 = x_143; +x_127 = x_142; +x_128 = x_143; +x_129 = x_140; +x_130 = x_141; x_131 = x_146; goto block_134; } @@ -12886,10 +12886,10 @@ lean_ctor_set(x_154, 2, x_150); lean_ctor_set(x_154, 3, x_151); x_155 = l_Lean_MacroScopesView_review(x_154); x_126 = x_139; -x_127 = x_141; -x_128 = x_140; -x_129 = x_142; -x_130 = x_143; +x_127 = x_142; +x_128 = x_143; +x_129 = x_140; +x_130 = x_141; x_131 = x_155; goto block_134; } @@ -13093,13 +13093,13 @@ return x_124; { uint8_t x_132; lean_object* x_133; x_132 = 0; -x_133 = l_Lean_mkIdentFrom(x_127, x_131, x_132); -lean_dec(x_127); +x_133 = l_Lean_mkIdentFrom(x_130, x_131, x_132); +lean_dec(x_130); x_107 = x_126; -x_108 = x_128; +x_108 = x_129; x_109 = x_133; -x_110 = x_129; -x_111 = x_130; +x_110 = x_127; +x_111 = x_128; goto block_125; } } @@ -13603,7 +13603,7 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; size_t x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; size_t x_107; lean_object* x_108; lean_object* x_109; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; size_t x_131; lean_object* x_132; lean_object* x_133; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; size_t x_163; lean_object* x_164; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; size_t x_186; lean_object* x_187; lean_object* x_188; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; size_t x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; size_t x_254; lean_object* x_255; lean_object* x_256; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; size_t x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; size_t x_304; lean_object* x_305; lean_object* x_306; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_455; lean_object* x_456; uint8_t x_457; +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; size_t x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; size_t x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; size_t x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; size_t x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; size_t x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; size_t x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; size_t x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; size_t x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; size_t x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_455; lean_object* x_456; uint8_t x_457; x_14 = lean_unsigned_to_nat(1u); x_46 = l_Lean_Syntax_getArg(x_1, x_14); x_47 = lean_unsigned_to_nat(2u); @@ -13658,9 +13658,9 @@ goto block_454; block_45: { lean_object* x_24; lean_object* x_25; -x_24 = lean_array_get_size(x_20); -lean_dec_ref(x_20); -x_25 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls(x_23, x_16, x_24, x_22, x_18, x_21, x_17); +x_24 = lean_array_get_size(x_22); +lean_dec_ref(x_22); +x_25 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls(x_23, x_19, x_24, x_15, x_20, x_17, x_21); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -13670,13 +13670,13 @@ if (x_26 == 0) lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; x_27 = lean_ctor_get(x_25, 0); x_28 = l_Lake_expandConfigDecl___closed__1; -x_29 = lean_array_push(x_28, x_15); +x_29 = lean_array_push(x_28, x_18); x_30 = l_Array_append___redArg(x_29, x_27); lean_dec(x_27); x_31 = lean_box(2); x_32 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_19); +lean_ctor_set(x_32, 1, x_16); lean_ctor_set(x_32, 2, x_30); lean_ctor_set(x_25, 0, x_32); return x_25; @@ -13690,13 +13690,13 @@ lean_inc(x_34); lean_inc(x_33); lean_dec(x_25); x_35 = l_Lake_expandConfigDecl___closed__1; -x_36 = lean_array_push(x_35, x_15); +x_36 = lean_array_push(x_35, x_18); x_37 = l_Array_append___redArg(x_36, x_33); lean_dec(x_33); x_38 = lean_box(2); x_39 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_19); +lean_ctor_set(x_39, 1, x_16); lean_ctor_set(x_39, 2, x_37); x_40 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_40, 0, x_39); @@ -13707,8 +13707,8 @@ return x_40; else { uint8_t x_41; -lean_dec(x_19); -lean_dec(x_15); +lean_dec(x_18); +lean_dec(x_16); x_41 = !lean_is_exclusive(x_25); if (x_41 == 0) { @@ -13732,40 +13732,40 @@ return x_44; block_90: { lean_object* x_69; lean_object* x_70; lean_object* x_71; size_t x_72; lean_object* x_73; size_t x_74; lean_object* x_75; size_t x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -lean_dec(x_51); +lean_dec(x_50); lean_inc_ref(x_62); x_69 = l_Array_append___redArg(x_62, x_68); lean_dec_ref(x_68); -lean_inc(x_56); -lean_inc(x_57); +lean_inc(x_51); +lean_inc(x_65); x_70 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_70, 0, x_57); -lean_ctor_set(x_70, 1, x_56); +lean_ctor_set(x_70, 0, x_65); +lean_ctor_set(x_70, 1, x_51); lean_ctor_set(x_70, 2, x_69); x_71 = l_Lake_expandConfigDecl___closed__3; -x_72 = lean_array_size(x_64); -x_73 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_expandConfigDecl_spec__5___redArg(x_72, x_65, x_64); +x_72 = lean_array_size(x_60); +x_73 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_expandConfigDecl_spec__5___redArg(x_72, x_55, x_60); x_74 = lean_array_size(x_73); -x_75 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_expandConfigDecl_spec__5___redArg(x_74, x_65, x_73); +x_75 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_expandConfigDecl_spec__5___redArg(x_74, x_55, x_73); x_76 = lean_array_size(x_75); -x_77 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_expandConfigDecl_spec__5___redArg(x_76, x_65, x_75); +x_77 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_expandConfigDecl_spec__5___redArg(x_76, x_55, x_75); x_78 = l_Array_append___redArg(x_62, x_77); lean_dec_ref(x_77); -lean_inc(x_56); -lean_inc(x_57); +lean_inc(x_51); +lean_inc(x_65); x_79 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_79, 0, x_57); -lean_ctor_set(x_79, 1, x_56); +lean_ctor_set(x_79, 0, x_65); +lean_ctor_set(x_79, 1, x_51); lean_ctor_set(x_79, 2, x_78); -lean_inc(x_57); -x_80 = l_Lean_Syntax_node1(x_57, x_71, x_79); -lean_inc(x_56); -lean_inc(x_57); -x_81 = l_Lean_Syntax_node3(x_57, x_56, x_66, x_70, x_80); -lean_inc(x_57); -x_82 = l_Lean_Syntax_node6(x_57, x_54, x_53, x_48, x_60, x_58, x_81, x_49); +lean_inc(x_65); +x_80 = l_Lean_Syntax_node1(x_65, x_71, x_79); +lean_inc(x_51); +lean_inc(x_65); +x_81 = l_Lean_Syntax_node3(x_65, x_51, x_59, x_70, x_80); +lean_inc(x_65); +x_82 = l_Lean_Syntax_node6(x_65, x_66, x_54, x_48, x_61, x_63, x_81, x_52); lean_inc(x_9); -x_83 = l_Lean_Syntax_node2(x_57, x_59, x_9, x_82); +x_83 = l_Lean_Syntax_node2(x_65, x_67, x_9, x_82); x_84 = l_Lean_Syntax_getArg(x_9, x_47); lean_dec(x_9); x_85 = l_Lean_Syntax_getOptional_x3f(x_84); @@ -13774,14 +13774,14 @@ if (lean_obj_tag(x_85) == 0) { lean_object* x_86; x_86 = lean_box(0); -x_15 = x_83; -x_16 = x_50; -x_17 = x_52; -x_18 = x_63; +x_15 = x_49; +x_16 = x_51; +x_17 = x_53; +x_18 = x_83; x_19 = x_56; -x_20 = x_55; -x_21 = x_67; -x_22 = x_61; +x_20 = x_57; +x_21 = x_64; +x_22 = x_58; x_23 = x_86; goto block_45; } @@ -13791,14 +13791,14 @@ uint8_t x_87; x_87 = !lean_is_exclusive(x_85); if (x_87 == 0) { -x_15 = x_83; -x_16 = x_50; -x_17 = x_52; -x_18 = x_63; +x_15 = x_49; +x_16 = x_51; +x_17 = x_53; +x_18 = x_83; x_19 = x_56; -x_20 = x_55; -x_21 = x_67; -x_22 = x_61; +x_20 = x_57; +x_21 = x_64; +x_22 = x_58; x_23 = x_85; goto block_45; } @@ -13810,14 +13810,14 @@ lean_inc(x_88); lean_dec(x_85); x_89 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_89, 0, x_88); -x_15 = x_83; -x_16 = x_50; -x_17 = x_52; -x_18 = x_63; +x_15 = x_49; +x_16 = x_51; +x_17 = x_53; +x_18 = x_83; x_19 = x_56; -x_20 = x_55; -x_21 = x_67; -x_22 = x_61; +x_20 = x_57; +x_21 = x_64; +x_22 = x_58; x_23 = x_89; goto block_45; } @@ -13840,87 +13840,87 @@ x_58 = x_100; x_59 = x_101; x_60 = x_102; x_61 = x_103; -x_62 = x_104; -x_63 = x_105; -x_64 = x_106; -x_65 = x_107; -x_66 = x_108; -x_67 = x_109; +x_62 = x_105; +x_63 = x_104; +x_64 = x_109; +x_65 = x_108; +x_66 = x_107; +x_67 = x_106; x_68 = x_110; goto block_90; } block_143: { lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; -lean_inc_ref(x_126); -x_134 = l_Array_append___redArg(x_126, x_133); +lean_inc_ref(x_128); +x_134 = l_Array_append___redArg(x_128, x_133); lean_dec_ref(x_133); -lean_inc(x_119); -lean_inc(x_121); +lean_inc(x_114); +lean_inc(x_132); x_135 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_135, 0, x_121); -lean_ctor_set(x_135, 1, x_119); +lean_ctor_set(x_135, 0, x_132); +lean_ctor_set(x_135, 1, x_114); lean_ctor_set(x_135, 2, x_134); -lean_inc(x_121); -x_136 = l_Lean_Syntax_node3(x_121, x_127, x_113, x_129, x_135); -lean_inc(x_119); -lean_inc(x_121); -x_137 = l_Lean_Syntax_node1(x_121, x_119, x_136); +lean_inc(x_132); +x_136 = l_Lean_Syntax_node3(x_132, x_127, x_119, x_123, x_135); +lean_inc(x_114); +lean_inc(x_132); +x_137 = l_Lean_Syntax_node1(x_132, x_114, x_136); x_138 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls_spec__3_spec__3___closed__32; -lean_inc(x_121); +lean_inc(x_132); x_139 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_139, 0, x_121); +lean_ctor_set(x_139, 0, x_132); lean_ctor_set(x_139, 1, x_138); -if (lean_obj_tag(x_122) == 0) +if (lean_obj_tag(x_125) == 0) { x_91 = x_112; -x_92 = x_114; -x_93 = x_115; -x_94 = x_116; -x_95 = x_117; -x_96 = x_118; -x_97 = x_120; -x_98 = x_119; +x_92 = x_113; +x_93 = x_114; +x_94 = x_115; +x_95 = x_116; +x_96 = x_117; +x_97 = x_118; +x_98 = x_120; x_99 = x_121; -x_100 = x_137; -x_101 = x_123; +x_100 = x_122; +x_101 = x_139; x_102 = x_124; -x_103 = x_125; -x_104 = x_126; +x_103 = x_126; +x_104 = x_137; x_105 = x_128; -x_106 = x_130; -x_107 = x_131; -x_108 = x_139; -x_109 = x_132; +x_106 = x_129; +x_107 = x_130; +x_108 = x_132; +x_109 = x_131; goto block_111; } else { lean_object* x_140; -x_140 = lean_ctor_get(x_122, 0); +x_140 = lean_ctor_get(x_125, 0); lean_inc(x_140); -lean_dec_ref(x_122); +lean_dec_ref(x_125); if (lean_obj_tag(x_140) == 0) { x_91 = x_112; -x_92 = x_114; -x_93 = x_115; -x_94 = x_116; -x_95 = x_117; -x_96 = x_118; -x_97 = x_120; -x_98 = x_119; +x_92 = x_113; +x_93 = x_114; +x_94 = x_115; +x_95 = x_116; +x_96 = x_117; +x_97 = x_118; +x_98 = x_120; x_99 = x_121; -x_100 = x_137; -x_101 = x_123; +x_100 = x_122; +x_101 = x_139; x_102 = x_124; -x_103 = x_125; -x_104 = x_126; +x_103 = x_126; +x_104 = x_137; x_105 = x_128; -x_106 = x_130; -x_107 = x_131; -x_108 = x_139; -x_109 = x_132; +x_106 = x_129; +x_107 = x_130; +x_108 = x_132; +x_109 = x_131; goto block_111; } else @@ -13931,24 +13931,24 @@ lean_inc(x_141); lean_dec_ref(x_140); x_142 = l_Lake_expandConfigDecl___lam__0(x_141); x_49 = x_112; -x_50 = x_114; -x_51 = x_115; -x_52 = x_116; -x_53 = x_117; -x_54 = x_118; -x_55 = x_120; -x_56 = x_119; +x_50 = x_113; +x_51 = x_114; +x_52 = x_115; +x_53 = x_116; +x_54 = x_117; +x_55 = x_118; +x_56 = x_120; x_57 = x_121; -x_58 = x_137; -x_59 = x_123; +x_58 = x_122; +x_59 = x_139; x_60 = x_124; -x_61 = x_125; -x_62 = x_126; -x_63 = x_128; -x_64 = x_130; -x_65 = x_131; -x_66 = x_139; -x_67 = x_132; +x_61 = x_126; +x_62 = x_128; +x_63 = x_137; +x_64 = x_131; +x_65 = x_132; +x_66 = x_130; +x_67 = x_129; x_68 = x_142; goto block_90; } @@ -13971,100 +13971,100 @@ x_121 = x_153; x_122 = x_154; x_123 = x_155; x_124 = x_156; -x_125 = x_157; -x_126 = x_158; +x_125 = x_158; +x_126 = x_157; x_127 = x_159; x_128 = x_160; -x_129 = x_161; -x_130 = x_162; -x_131 = x_163; -x_132 = x_164; +x_129 = x_164; +x_130 = x_163; +x_131 = x_162; +x_132 = x_161; x_133 = x_165; goto block_143; } block_200: { 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_inc_ref(x_180); -x_189 = l_Array_append___redArg(x_180, x_188); +lean_inc_ref(x_181); +x_189 = l_Array_append___redArg(x_181, x_188); lean_dec_ref(x_188); -lean_inc(x_174); -lean_inc(x_176); +lean_inc(x_169); +lean_inc(x_187); x_190 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_190, 0, x_176); -lean_ctor_set(x_190, 1, x_174); +lean_ctor_set(x_190, 0, x_187); +lean_ctor_set(x_190, 1, x_169); lean_ctor_set(x_190, 2, x_189); -lean_inc(x_176); -x_191 = l_Lean_Syntax_node2(x_176, x_185, x_182, x_190); +lean_inc(x_187); +x_191 = l_Lean_Syntax_node2(x_187, x_176, x_182, x_190); x_192 = l_Lake_configDecl___closed__32; x_193 = l_Lake_configDecl___closed__33; -lean_inc(x_176); +lean_inc(x_187); x_194 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_194, 0, x_176); +lean_ctor_set(x_194, 0, x_187); lean_ctor_set(x_194, 1, x_192); -lean_inc_ref(x_180); -x_195 = l_Array_append___redArg(x_180, x_172); -lean_dec_ref(x_172); -lean_inc(x_174); -lean_inc(x_176); +lean_inc_ref(x_181); +x_195 = l_Array_append___redArg(x_181, x_183); +lean_dec_ref(x_183); +lean_inc(x_169); +lean_inc(x_187); x_196 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_196, 0, x_176); -lean_ctor_set(x_196, 1, x_174); +lean_ctor_set(x_196, 0, x_187); +lean_ctor_set(x_196, 1, x_169); lean_ctor_set(x_196, 2, x_195); -if (lean_obj_tag(x_183) == 0) +if (lean_obj_tag(x_178) == 0) { x_144 = x_167; -x_145 = x_194; -x_146 = x_168; -x_147 = x_169; -x_148 = x_170; -x_149 = x_171; +x_145 = x_168; +x_146 = x_169; +x_147 = x_170; +x_148 = x_171; +x_149 = x_172; x_150 = x_173; -x_151 = x_174; -x_152 = x_175; -x_153 = x_176; -x_154 = x_178; -x_155 = x_177; -x_156 = x_191; -x_157 = x_179; +x_151 = x_194; +x_152 = x_174; +x_153 = x_175; +x_154 = x_177; +x_155 = x_196; +x_156 = x_179; +x_157 = x_191; x_158 = x_180; x_159 = x_193; x_160 = x_181; -x_161 = x_196; +x_161 = x_187; x_162 = x_184; -x_163 = x_186; -x_164 = x_187; +x_163 = x_185; +x_164 = x_186; goto block_166; } else { lean_object* x_197; -x_197 = lean_ctor_get(x_183, 0); +x_197 = lean_ctor_get(x_178, 0); lean_inc(x_197); -lean_dec_ref(x_183); +lean_dec_ref(x_178); if (lean_obj_tag(x_197) == 0) { x_144 = x_167; -x_145 = x_194; -x_146 = x_168; -x_147 = x_169; -x_148 = x_170; -x_149 = x_171; +x_145 = x_168; +x_146 = x_169; +x_147 = x_170; +x_148 = x_171; +x_149 = x_172; x_150 = x_173; -x_151 = x_174; -x_152 = x_175; -x_153 = x_176; -x_154 = x_178; -x_155 = x_177; -x_156 = x_191; -x_157 = x_179; +x_151 = x_194; +x_152 = x_174; +x_153 = x_175; +x_154 = x_177; +x_155 = x_196; +x_156 = x_179; +x_157 = x_191; x_158 = x_180; x_159 = x_193; x_160 = x_181; -x_161 = x_196; +x_161 = x_187; x_162 = x_184; -x_163 = x_186; -x_164 = x_187; +x_163 = x_185; +x_164 = x_186; goto block_166; } else @@ -14075,25 +14075,25 @@ lean_inc(x_198); lean_dec_ref(x_197); x_199 = l_Lake_expandConfigDecl___lam__0(x_198); x_112 = x_167; -x_113 = x_194; -x_114 = x_168; -x_115 = x_169; -x_116 = x_170; -x_117 = x_171; +x_113 = x_168; +x_114 = x_169; +x_115 = x_170; +x_116 = x_171; +x_117 = x_172; x_118 = x_173; -x_119 = x_174; -x_120 = x_175; -x_121 = x_176; -x_122 = x_178; -x_123 = x_177; -x_124 = x_191; -x_125 = x_179; -x_126 = x_180; +x_119 = x_194; +x_120 = x_174; +x_121 = x_175; +x_122 = x_177; +x_123 = x_196; +x_124 = x_179; +x_125 = x_180; +x_126 = x_191; x_127 = x_193; x_128 = x_181; -x_129 = x_196; -x_130 = x_184; -x_131 = x_186; +x_129 = x_186; +x_130 = x_185; +x_131 = x_184; x_132 = x_187; x_133 = x_199; goto block_143; @@ -14103,14 +14103,14 @@ goto block_143; block_240: { lean_object* x_218; lean_object* x_219; uint8_t x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; size_t x_231; lean_object* x_232; size_t x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; -lean_dec(x_214); -lean_dec_ref(x_208); -lean_dec(x_205); +lean_dec(x_212); +lean_dec_ref(x_210); +lean_dec(x_204); x_218 = lean_array_get_size(x_216); x_219 = l_Array_filterMapM___at___00Lake_expandConfigDecl_spec__2(x_216, x_8, x_218); x_220 = 0; -x_221 = l_Lean_SourceInfo_fromRef(x_210, x_220); -lean_dec(x_210); +x_221 = l_Lean_SourceInfo_fromRef(x_203, x_220); +lean_dec(x_203); x_222 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls_spec__2___closed__51; x_223 = l_Lake_expandConfigDecl___closed__4; x_224 = l_Lake_expandConfigDecl___closed__5; @@ -14124,11 +14124,11 @@ x_227 = l_Lean_Syntax_node1(x_221, x_225, x_226); x_228 = l___private_Lake_Config_Meta_0__Lake_mkFieldView___closed__7; x_229 = l___private_Lake_Config_Meta_0__Lake_instCoeIdentTSyntaxConsSyntaxNodeKindMkStr4Nil__lake___lam__0___closed__4; x_230 = l___private_Lake_Config_Meta_0__Lake_instCoeIdentTSyntaxConsSyntaxNodeKindMkStr4Nil__lake___lam__0___closed__5; -x_231 = lean_array_size(x_206); -lean_inc_ref(x_206); -x_232 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_expandConfigDecl_spec__3___redArg(x_231, x_213, x_206); +x_231 = lean_array_size(x_209); +lean_inc_ref(x_209); +x_232 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_expandConfigDecl_spec__3___redArg(x_231, x_207, x_209); x_233 = lean_array_size(x_232); -x_234 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_expandConfigDecl_spec__4___redArg(x_233, x_213, x_232); +x_234 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_expandConfigDecl_spec__4___redArg(x_233, x_207, x_232); x_235 = l_Array_append___redArg(x_230, x_234); lean_dec_ref(x_234); lean_inc(x_221); @@ -14136,63 +14136,63 @@ x_236 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_236, 0, x_221); lean_ctor_set(x_236, 1, x_229); lean_ctor_set(x_236, 2, x_235); -if (lean_obj_tag(x_212) == 1) +if (lean_obj_tag(x_214) == 1) { lean_object* x_237; lean_object* x_238; -x_237 = lean_ctor_get(x_212, 0); +x_237 = lean_ctor_get(x_214, 0); lean_inc(x_237); -lean_dec_ref(x_212); +lean_dec_ref(x_214); x_238 = l_Array_mkArray1___redArg(x_237); -x_167 = x_201; -x_168 = x_203; -x_169 = x_202; -x_170 = x_217; -x_171 = x_227; -x_172 = x_204; -x_173 = x_224; -x_174 = x_229; -x_175 = x_206; -x_176 = x_221; -x_177 = x_222; -x_178 = x_207; -x_179 = x_209; -x_180 = x_230; -x_181 = x_216; +x_167 = x_202; +x_168 = x_201; +x_169 = x_229; +x_170 = x_205; +x_171 = x_206; +x_172 = x_227; +x_173 = x_207; +x_174 = x_208; +x_175 = x_216; +x_176 = x_228; +x_177 = x_209; +x_178 = x_211; +x_179 = x_219; +x_180 = x_213; +x_181 = x_230; x_182 = x_236; -x_183 = x_211; -x_184 = x_219; -x_185 = x_228; -x_186 = x_213; -x_187 = x_215; +x_183 = x_215; +x_184 = x_217; +x_185 = x_224; +x_186 = x_222; +x_187 = x_221; x_188 = x_238; goto block_200; } else { lean_object* x_239; -lean_dec(x_212); +lean_dec(x_214); x_239 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls_spec__2___closed__53; -x_167 = x_201; -x_168 = x_203; -x_169 = x_202; -x_170 = x_217; -x_171 = x_227; -x_172 = x_204; -x_173 = x_224; -x_174 = x_229; -x_175 = x_206; -x_176 = x_221; -x_177 = x_222; -x_178 = x_207; -x_179 = x_209; -x_180 = x_230; -x_181 = x_216; +x_167 = x_202; +x_168 = x_201; +x_169 = x_229; +x_170 = x_205; +x_171 = x_206; +x_172 = x_227; +x_173 = x_207; +x_174 = x_208; +x_175 = x_216; +x_176 = x_228; +x_177 = x_209; +x_178 = x_211; +x_179 = x_219; +x_180 = x_213; +x_181 = x_230; x_182 = x_236; -x_183 = x_211; -x_184 = x_219; -x_185 = x_228; -x_186 = x_213; -x_187 = x_215; +x_183 = x_215; +x_184 = x_217; +x_185 = x_224; +x_186 = x_222; +x_187 = x_221; x_188 = x_239; goto block_200; } @@ -14219,9 +14219,9 @@ x_209 = x_249; x_210 = x_250; x_211 = x_251; x_212 = x_252; -x_213 = x_254; -x_214 = x_253; -x_215 = x_255; +x_213 = x_253; +x_214 = x_255; +x_215 = x_254; x_216 = x_257; x_217 = x_258; goto block_240; @@ -14229,17 +14229,17 @@ goto block_240; else { uint8_t x_259; -lean_dec_ref(x_255); +lean_dec(x_255); +lean_dec_ref(x_254); lean_dec(x_253); lean_dec(x_252); lean_dec(x_251); -lean_dec(x_250); -lean_dec(x_249); -lean_dec_ref(x_248); -lean_dec(x_247); +lean_dec_ref(x_250); +lean_dec_ref(x_249); +lean_dec(x_248); lean_dec_ref(x_246); lean_dec(x_245); -lean_dec_ref(x_244); +lean_dec(x_244); lean_dec(x_243); lean_dec(x_242); lean_dec(x_241); @@ -14277,20 +14277,20 @@ lean_dec_ref(x_281); x_201 = x_264; x_202 = x_265; x_203 = x_266; -x_204 = x_280; -x_205 = x_267; -x_206 = x_268; -x_207 = x_269; -x_208 = x_270; -x_209 = x_271; -x_210 = x_272; -x_211 = x_275; -x_212 = x_276; -x_213 = x_277; -x_214 = x_278; -x_215 = x_279; -x_216 = x_274; -x_217 = x_273; +x_204 = x_267; +x_205 = x_269; +x_206 = x_270; +x_207 = x_271; +x_208 = x_272; +x_209 = x_274; +x_210 = x_275; +x_211 = x_276; +x_212 = x_277; +x_213 = x_278; +x_214 = x_279; +x_215 = x_280; +x_216 = x_273; +x_217 = x_268; goto block_240; } else @@ -14305,43 +14305,43 @@ lean_dec_ref(x_281); x_201 = x_264; x_202 = x_265; x_203 = x_266; -x_204 = x_280; -x_205 = x_267; -x_206 = x_268; -x_207 = x_269; -x_208 = x_270; -x_209 = x_271; -x_210 = x_272; -x_211 = x_275; -x_212 = x_276; -x_213 = x_277; -x_214 = x_278; -x_215 = x_279; -x_216 = x_274; -x_217 = x_273; +x_204 = x_267; +x_205 = x_269; +x_206 = x_270; +x_207 = x_271; +x_208 = x_272; +x_209 = x_274; +x_210 = x_275; +x_211 = x_276; +x_212 = x_277; +x_213 = x_278; +x_214 = x_279; +x_215 = x_280; +x_216 = x_273; +x_217 = x_268; goto block_240; } else { size_t x_285; lean_object* x_286; x_285 = lean_usize_of_nat(x_282); -lean_inc_ref(x_279); -x_286 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lake_expandConfigDecl_spec__6(x_281, x_277, x_285, x_274, x_279, x_273); +lean_inc_ref(x_270); +x_286 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lake_expandConfigDecl_spec__6(x_281, x_271, x_285, x_273, x_270, x_268); lean_dec_ref(x_281); x_241 = x_264; x_242 = x_265; x_243 = x_266; -x_244 = x_280; -x_245 = x_267; -x_246 = x_268; -x_247 = x_269; -x_248 = x_270; -x_249 = x_271; -x_250 = x_272; -x_251 = x_275; -x_252 = x_276; +x_244 = x_267; +x_245 = x_269; +x_246 = x_270; +x_247 = x_271; +x_248 = x_272; +x_249 = x_274; +x_250 = x_275; +x_251 = x_276; +x_252 = x_277; x_253 = x_278; -x_254 = x_277; +x_254 = x_280; x_255 = x_279; x_256 = x_286; goto block_263; @@ -14351,23 +14351,23 @@ else { size_t x_287; lean_object* x_288; x_287 = lean_usize_of_nat(x_282); -lean_inc_ref(x_279); -x_288 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lake_expandConfigDecl_spec__6(x_281, x_277, x_287, x_274, x_279, x_273); +lean_inc_ref(x_270); +x_288 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lake_expandConfigDecl_spec__6(x_281, x_271, x_287, x_273, x_270, x_268); lean_dec_ref(x_281); x_241 = x_264; x_242 = x_265; x_243 = x_266; -x_244 = x_280; -x_245 = x_267; -x_246 = x_268; -x_247 = x_269; -x_248 = x_270; -x_249 = x_271; -x_250 = x_272; -x_251 = x_275; -x_252 = x_276; +x_244 = x_267; +x_245 = x_269; +x_246 = x_270; +x_247 = x_271; +x_248 = x_272; +x_249 = x_274; +x_250 = x_275; +x_251 = x_276; +x_252 = x_277; x_253 = x_278; -x_254 = x_277; +x_254 = x_280; x_255 = x_279; x_256 = x_288; goto block_263; @@ -14378,11 +14378,11 @@ goto block_263; { size_t x_307; lean_object* x_308; x_307 = lean_array_size(x_306); -lean_inc_ref(x_305); -x_308 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_expandConfigDecl_spec__1(x_307, x_304, x_306, x_305, x_296); +lean_inc_ref(x_295); +x_308 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_expandConfigDecl_spec__1(x_307, x_296, x_306, x_295, x_297); if (lean_obj_tag(x_308) == 0) { -if (lean_obj_tag(x_294) == 0) +if (lean_obj_tag(x_305) == 0) { lean_object* x_309; lean_object* x_310; lean_object* x_311; x_309 = lean_ctor_get(x_308, 0); @@ -14395,18 +14395,18 @@ x_264 = x_290; x_265 = x_291; x_266 = x_292; x_267 = x_293; -x_268 = x_295; -x_269 = x_297; -x_270 = x_298; -x_271 = x_299; -x_272 = x_300; -x_273 = x_310; -x_274 = x_309; -x_275 = x_301; -x_276 = x_302; -x_277 = x_304; +x_268 = x_310; +x_269 = x_294; +x_270 = x_295; +x_271 = x_296; +x_272 = x_298; +x_273 = x_309; +x_274 = x_299; +x_275 = x_300; +x_276 = x_301; +x_277 = x_302; x_278 = x_303; -x_279 = x_305; +x_279 = x_304; x_280 = x_311; goto block_289; } @@ -14418,25 +14418,25 @@ lean_inc(x_312); x_313 = lean_ctor_get(x_308, 1); lean_inc(x_313); lean_dec_ref(x_308); -x_314 = lean_ctor_get(x_294, 0); +x_314 = lean_ctor_get(x_305, 0); lean_inc(x_314); -lean_dec_ref(x_294); +lean_dec_ref(x_305); x_264 = x_290; x_265 = x_291; x_266 = x_292; x_267 = x_293; -x_268 = x_295; -x_269 = x_297; -x_270 = x_298; -x_271 = x_299; -x_272 = x_300; -x_273 = x_313; -x_274 = x_312; -x_275 = x_301; -x_276 = x_302; -x_277 = x_304; +x_268 = x_313; +x_269 = x_294; +x_270 = x_295; +x_271 = x_296; +x_272 = x_298; +x_273 = x_312; +x_274 = x_299; +x_275 = x_300; +x_276 = x_301; +x_277 = x_302; x_278 = x_303; -x_279 = x_305; +x_279 = x_304; x_280 = x_314; goto block_289; } @@ -14444,14 +14444,14 @@ goto block_289; else { uint8_t x_315; -lean_dec_ref(x_305); +lean_dec(x_305); +lean_dec(x_304); lean_dec(x_303); lean_dec(x_302); lean_dec(x_301); -lean_dec(x_300); -lean_dec(x_299); -lean_dec_ref(x_298); -lean_dec(x_297); +lean_dec_ref(x_300); +lean_dec_ref(x_299); +lean_dec(x_298); lean_dec_ref(x_295); lean_dec(x_294); lean_dec(x_293); @@ -14522,22 +14522,22 @@ if (lean_obj_tag(x_326) == 0) { lean_object* x_347; x_347 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___closed__0; -x_290 = x_340; -x_291 = x_336; -x_292 = x_342; +x_290 = x_336; +x_291 = x_346; +x_292 = x_332; x_293 = x_338; -x_294 = x_323; -x_295 = x_331; -x_296 = x_335; -x_297 = x_325; -x_298 = x_337; -x_299 = x_346; -x_300 = x_332; +x_294 = x_340; +x_295 = x_327; +x_296 = x_344; +x_297 = x_335; +x_298 = x_342; +x_299 = x_331; +x_300 = x_337; x_301 = x_322; -x_302 = x_324; -x_303 = x_341; -x_304 = x_344; -x_305 = x_327; +x_302 = x_341; +x_303 = x_325; +x_304 = x_323; +x_305 = x_324; x_306 = x_347; goto block_319; } @@ -14547,22 +14547,22 @@ lean_object* x_348; x_348 = lean_ctor_get(x_326, 0); lean_inc(x_348); lean_dec_ref(x_326); -x_290 = x_340; -x_291 = x_336; -x_292 = x_342; +x_290 = x_336; +x_291 = x_346; +x_292 = x_332; x_293 = x_338; -x_294 = x_323; -x_295 = x_331; -x_296 = x_335; -x_297 = x_325; -x_298 = x_337; -x_299 = x_346; -x_300 = x_332; +x_294 = x_340; +x_295 = x_327; +x_296 = x_344; +x_297 = x_335; +x_298 = x_342; +x_299 = x_331; +x_300 = x_337; x_301 = x_322; -x_302 = x_324; -x_303 = x_341; -x_304 = x_344; -x_305 = x_327; +x_302 = x_341; +x_303 = x_325; +x_304 = x_323; +x_305 = x_324; x_306 = x_348; goto block_319; } @@ -14657,22 +14657,22 @@ if (lean_obj_tag(x_326) == 0) { lean_object* x_376; x_376 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___closed__0; -x_290 = x_369; -x_291 = x_365; -x_292 = x_371; +x_290 = x_365; +x_291 = x_375; +x_292 = x_360; x_293 = x_367; -x_294 = x_323; -x_295 = x_359; -x_296 = x_364; -x_297 = x_325; -x_298 = x_366; -x_299 = x_375; -x_300 = x_360; +x_294 = x_369; +x_295 = x_361; +x_296 = x_373; +x_297 = x_364; +x_298 = x_371; +x_299 = x_359; +x_300 = x_366; x_301 = x_322; -x_302 = x_324; -x_303 = x_370; -x_304 = x_373; -x_305 = x_361; +x_302 = x_370; +x_303 = x_325; +x_304 = x_323; +x_305 = x_324; x_306 = x_376; goto block_319; } @@ -14682,22 +14682,22 @@ lean_object* x_377; x_377 = lean_ctor_get(x_326, 0); lean_inc(x_377); lean_dec_ref(x_326); -x_290 = x_369; -x_291 = x_365; -x_292 = x_371; +x_290 = x_365; +x_291 = x_375; +x_292 = x_360; x_293 = x_367; -x_294 = x_323; -x_295 = x_359; -x_296 = x_364; -x_297 = x_325; -x_298 = x_366; -x_299 = x_375; -x_300 = x_360; +x_294 = x_369; +x_295 = x_361; +x_296 = x_373; +x_297 = x_364; +x_298 = x_371; +x_299 = x_359; +x_300 = x_366; x_301 = x_322; -x_302 = x_324; -x_303 = x_370; -x_304 = x_373; -x_305 = x_361; +x_302 = x_370; +x_303 = x_325; +x_304 = x_323; +x_305 = x_324; x_306 = x_377; goto block_319; } @@ -14742,8 +14742,8 @@ return x_381; block_394: { lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; -x_390 = l_Lean_Syntax_getArg(x_386, x_47); -lean_dec(x_386); +x_390 = l_Lean_Syntax_getArg(x_384, x_47); +lean_dec(x_384); x_391 = l_Lean_Syntax_getArgs(x_390); lean_dec(x_390); x_392 = lean_alloc_ctor(1, 1, 0); @@ -14751,8 +14751,8 @@ lean_ctor_set(x_392, 0, x_387); x_393 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_393, 0, x_391); x_322 = x_383; -x_323 = x_384; -x_324 = x_385; +x_323 = x_385; +x_324 = x_386; x_325 = x_392; x_326 = x_393; x_327 = x_388; @@ -14846,9 +14846,9 @@ lean_dec(x_411); x_417 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_417, 0, x_416); x_383 = x_397; -x_384 = x_396; +x_384 = x_401; x_385 = x_395; -x_386 = x_401; +x_386 = x_396; x_387 = x_417; x_388 = x_398; x_389 = x_399; @@ -14861,9 +14861,9 @@ lean_object* x_418; lean_dec(x_411); x_418 = lean_box(0); x_383 = x_397; -x_384 = x_396; +x_384 = x_401; x_385 = x_395; -x_386 = x_401; +x_386 = x_396; x_387 = x_418; x_388 = x_398; x_389 = x_399; @@ -14878,8 +14878,8 @@ lean_object* x_419; lean_dec(x_401); x_419 = lean_box(0); x_322 = x_397; -x_323 = x_396; -x_324 = x_395; +x_323 = x_395; +x_324 = x_396; x_325 = x_419; x_326 = x_419; x_327 = x_398; @@ -14890,13 +14890,13 @@ goto block_382; block_429: { lean_object* x_426; lean_object* x_427; lean_object* x_428; -x_426 = l_Lean_Syntax_getArgs(x_422); -lean_dec(x_422); +x_426 = l_Lean_Syntax_getArgs(x_421); +lean_dec(x_421); x_427 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_427, 0, x_426); x_428 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_428, 0, x_423); -x_395 = x_421; +x_395 = x_422; x_396 = x_427; x_397 = x_428; x_398 = x_424; @@ -14987,8 +14987,8 @@ x_450 = l_Lean_Syntax_getArg(x_445, x_8); lean_dec(x_445); x_451 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_451, 0, x_450); -x_421 = x_430; -x_422 = x_444; +x_421 = x_444; +x_422 = x_430; x_423 = x_451; x_424 = x_431; x_425 = x_432; @@ -15000,8 +15000,8 @@ else lean_object* x_452; lean_dec(x_445); x_452 = lean_box(0); -x_421 = x_430; -x_422 = x_444; +x_421 = x_444; +x_422 = x_430; x_423 = x_452; x_424 = x_431; x_425 = x_432; diff --git a/stage0/stdlib/Lake/Load/Toml.c b/stage0/stdlib/Lake/Load/Toml.c index 257a6314797a..07cd9bbcf9ad 100644 --- a/stage0/stdlib/Lake/Load/Toml.c +++ b/stage0/stdlib/Lake/Load/Toml.c @@ -11540,9 +11540,9 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_dec(x_8); -lean_dec_ref(x_7); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); +lean_dec_ref(x_5); lean_dec(x_4); x_11 = lean_box(0); x_12 = lean_alloc_ctor(1, 2, 0); @@ -11554,10 +11554,10 @@ else { lean_object* x_13; lean_object* x_14; x_13 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_13, 0, x_6); -lean_ctor_set(x_13, 1, x_7); -lean_ctor_set(x_13, 2, x_5); -lean_ctor_set(x_13, 3, x_4); +lean_ctor_set(x_13, 0, x_7); +lean_ctor_set(x_13, 1, x_5); +lean_ctor_set(x_13, 2, x_4); +lean_ctor_set(x_13, 3, x_6); lean_ctor_set(x_13, 4, x_8); x_14 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_14, 0, x_13); @@ -11574,10 +11574,10 @@ x_23 = lean_box(1); x_24 = l_Lake_Toml_RBDict_findEntry_x3f___redArg(x_22, x_21, x_1); if (lean_obj_tag(x_24) == 0) { -x_4 = x_16; -x_5 = x_19; -x_6 = x_17; -x_7 = x_18; +x_4 = x_19; +x_5 = x_16; +x_6 = x_18; +x_7 = x_17; x_8 = x_23; x_9 = x_20; goto block_15; @@ -11609,10 +11609,10 @@ lean_inc(x_31); x_32 = lean_ctor_get(x_30, 1); lean_inc(x_32); lean_dec_ref(x_30); -x_4 = x_16; -x_5 = x_19; -x_6 = x_17; -x_7 = x_18; +x_4 = x_19; +x_5 = x_16; +x_6 = x_18; +x_7 = x_17; x_8 = x_31; x_9 = x_32; goto block_15; @@ -11623,10 +11623,10 @@ lean_object* x_33; x_33 = lean_ctor_get(x_30, 1); lean_inc(x_33); lean_dec_ref(x_30); -x_4 = x_16; -x_5 = x_19; -x_6 = x_17; -x_7 = x_18; +x_4 = x_19; +x_5 = x_16; +x_6 = x_18; +x_7 = x_17; x_8 = x_23; x_9 = x_33; goto block_15; @@ -11638,10 +11638,10 @@ lean_object* x_34; x_34 = lean_ctor_get(x_27, 1); lean_inc(x_34); lean_dec_ref(x_27); -x_4 = x_16; -x_5 = x_19; -x_6 = x_17; -x_7 = x_18; +x_4 = x_19; +x_5 = x_16; +x_6 = x_18; +x_7 = x_17; x_8 = x_23; x_9 = x_34; goto block_15; @@ -11652,7 +11652,7 @@ goto block_15; { if (lean_obj_tag(x_36) == 1) { -if (lean_obj_tag(x_37) == 0) +if (lean_obj_tag(x_39) == 0) { uint8_t x_41; x_41 = !lean_is_exclusive(x_36); @@ -11726,9 +11726,9 @@ x_59 = l_Lake_Toml_RBDict_findEntry_x3f___redArg(x_58, x_57, x_1); if (lean_obj_tag(x_59) == 0) { x_36 = x_52; -x_37 = x_53; +x_37 = x_55; x_38 = x_54; -x_39 = x_55; +x_39 = x_53; x_40 = x_56; goto block_51; } @@ -11754,9 +11754,9 @@ x_65 = lean_ctor_get(x_63, 1); lean_inc(x_65); lean_dec_ref(x_63); lean_ctor_set(x_59, 0, x_64); -x_16 = x_53; +x_16 = x_55; x_17 = x_54; -x_18 = x_55; +x_18 = x_53; x_19 = x_59; x_20 = x_65; goto block_35; @@ -11769,9 +11769,9 @@ x_66 = lean_ctor_get(x_63, 1); lean_inc(x_66); lean_dec_ref(x_63); x_36 = x_52; -x_37 = x_53; +x_37 = x_55; x_38 = x_54; -x_39 = x_55; +x_39 = x_53; x_40 = x_66; goto block_51; } @@ -11797,9 +11797,9 @@ lean_inc(x_71); lean_dec_ref(x_69); x_72 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_72, 0, x_70); -x_16 = x_53; +x_16 = x_55; x_17 = x_54; -x_18 = x_55; +x_18 = x_53; x_19 = x_72; x_20 = x_71; goto block_35; @@ -11811,9 +11811,9 @@ x_73 = lean_ctor_get(x_69, 1); lean_inc(x_73); lean_dec_ref(x_69); x_36 = x_52; -x_37 = x_53; +x_37 = x_55; x_38 = x_54; -x_39 = x_55; +x_39 = x_53; x_40 = x_73; goto block_51; } @@ -11884,10 +11884,10 @@ x_94 = l_Lake_Dependency_decodeToml___closed__7; x_95 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_95, 0, x_93); lean_ctor_set(x_95, 1, x_94); -x_96 = lean_array_push(x_91, x_95); +x_96 = lean_array_push(x_92, x_95); x_97 = lean_box(0); x_75 = x_90; -x_76 = x_92; +x_76 = x_91; x_77 = x_97; x_78 = x_96; goto block_89; @@ -11897,7 +11897,7 @@ goto block_89; if (lean_obj_tag(x_102) == 0) { lean_object* x_103; lean_object* x_104; lean_object* x_105; -lean_dec(x_101); +lean_dec(x_100); x_103 = lean_ctor_get(x_102, 0); lean_inc(x_103); x_104 = lean_ctor_get(x_102, 1); @@ -11906,7 +11906,7 @@ lean_dec_ref(x_102); x_105 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_105, 0, x_103); x_75 = x_99; -x_76 = x_100; +x_76 = x_101; x_77 = x_105; x_78 = x_104; goto block_89; @@ -11918,8 +11918,8 @@ x_106 = lean_ctor_get(x_102, 1); lean_inc(x_106); lean_dec_ref(x_102); x_75 = x_99; -x_76 = x_100; -x_77 = x_101; +x_76 = x_101; +x_77 = x_100; x_78 = x_106; goto block_89; } @@ -11945,13 +11945,13 @@ goto block_89; lean_object* x_121; lean_object* x_122; lean_inc(x_116); x_121 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_121, 0, x_117); +lean_ctor_set(x_121, 0, x_118); lean_ctor_set(x_121, 1, x_116); lean_ctor_set(x_121, 2, x_119); x_122 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_122, 0, x_121); x_75 = x_116; -x_76 = x_118; +x_76 = x_117; x_77 = x_122; x_78 = x_120; goto block_89; @@ -11960,14 +11960,14 @@ goto block_89; { lean_object* x_130; lean_object* x_131; x_130 = l_Lake_DependencySrc_decodeToml___closed__1; -x_131 = l_Lake_Toml_RBDict_findEntry_x3f___redArg(x_127, x_130, x_125); +x_131 = l_Lake_Toml_RBDict_findEntry_x3f___redArg(x_127, x_130, x_126); if (lean_obj_tag(x_131) == 0) { lean_object* x_132; x_132 = lean_box(0); x_116 = x_124; -x_117 = x_128; -x_118 = x_126; +x_117 = x_125; +x_118 = x_128; x_119 = x_132; x_120 = x_129; goto block_123; @@ -11994,8 +11994,8 @@ lean_inc(x_138); lean_dec_ref(x_136); lean_ctor_set(x_131, 0, x_137); x_116 = x_124; -x_117 = x_128; -x_118 = x_126; +x_117 = x_125; +x_118 = x_128; x_119 = x_131; x_120 = x_138; goto block_123; @@ -12009,8 +12009,8 @@ lean_inc(x_139); lean_dec_ref(x_136); x_140 = lean_box(0); x_116 = x_124; -x_117 = x_128; -x_118 = x_126; +x_117 = x_125; +x_118 = x_128; x_119 = x_140; x_120 = x_139; goto block_123; @@ -12037,8 +12037,8 @@ lean_dec_ref(x_143); x_146 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_146, 0, x_144); x_116 = x_124; -x_117 = x_128; -x_118 = x_126; +x_117 = x_125; +x_118 = x_128; x_119 = x_146; x_120 = x_145; goto block_123; @@ -12051,8 +12051,8 @@ lean_inc(x_147); lean_dec_ref(x_143); x_148 = lean_box(0); x_116 = x_124; -x_117 = x_128; -x_118 = x_126; +x_117 = x_125; +x_118 = x_128; x_119 = x_148; x_120 = x_147; goto block_123; @@ -12108,8 +12108,8 @@ lean_inc(x_165); lean_dec(x_160); x_166 = l_Lake_DependencySrc_decodeToml(x_163, x_165, x_164); x_99 = x_150; -x_100 = x_151; -x_101 = x_161; +x_100 = x_161; +x_101 = x_151; x_102 = x_166; goto block_107; } @@ -12276,8 +12276,8 @@ x_199 = lean_ctor_get(x_197, 1); lean_inc(x_199); lean_dec_ref(x_197); x_124 = x_150; -x_125 = x_191; -x_126 = x_151; +x_125 = x_151; +x_126 = x_191; x_127 = x_153; x_128 = x_198; x_129 = x_199; @@ -12290,8 +12290,8 @@ x_200 = lean_ctor_get(x_197, 1); lean_inc(x_200); lean_dec_ref(x_197); x_124 = x_150; -x_125 = x_191; -x_126 = x_151; +x_125 = x_151; +x_126 = x_191; x_127 = x_153; x_128 = x_193; x_129 = x_200; @@ -12305,8 +12305,8 @@ x_201 = lean_ctor_get(x_194, 1); lean_inc(x_201); lean_dec_ref(x_194); x_124 = x_150; -x_125 = x_191; -x_126 = x_151; +x_125 = x_151; +x_126 = x_191; x_127 = x_153; x_128 = x_193; x_129 = x_201; @@ -12320,8 +12320,8 @@ x_202 = lean_ctor_get(x_169, 0); lean_inc(x_202); lean_dec(x_169); x_90 = x_150; -x_91 = x_152; -x_92 = x_151; +x_91 = x_151; +x_92 = x_152; x_93 = x_202; goto block_98; } @@ -62343,9 +62343,9 @@ goto block_52; block_44: { lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_nat_add(x_38, x_40); +x_41 = lean_nat_add(x_39, x_40); lean_dec(x_40); -lean_dec(x_38); +lean_dec(x_39); if (lean_is_scalar(x_35)) { x_42 = lean_alloc_ctor(0, 5, 0); } else { @@ -62364,7 +62364,7 @@ if (lean_is_scalar(x_25)) { lean_ctor_set(x_43, 0, x_37); lean_ctor_set(x_43, 1, x_28); lean_ctor_set(x_43, 2, x_29); -lean_ctor_set(x_43, 3, x_39); +lean_ctor_set(x_43, 3, x_38); lean_ctor_set(x_43, 4, x_42); return x_43; } @@ -62390,8 +62390,8 @@ if (lean_obj_tag(x_31) == 0) lean_object* x_50; x_50 = lean_ctor_get(x_31, 0); lean_inc(x_50); -x_38 = x_49; -x_39 = x_48; +x_38 = x_48; +x_39 = x_49; x_40 = x_50; goto block_44; } @@ -62399,8 +62399,8 @@ else { lean_object* x_51; x_51 = lean_unsigned_to_nat(0u); -x_38 = x_49; -x_39 = x_48; +x_38 = x_48; +x_39 = x_49; x_40 = x_51; goto block_44; } @@ -62817,9 +62817,9 @@ goto block_154; block_147: { lean_object* x_144; lean_object* x_145; lean_object* x_146; -x_144 = lean_nat_add(x_141, x_143); +x_144 = lean_nat_add(x_142, x_143); lean_dec(x_143); -lean_dec(x_141); +lean_dec(x_142); if (lean_is_scalar(x_138)) { x_145 = lean_alloc_ctor(0, 5, 0); } else { @@ -62838,7 +62838,7 @@ if (lean_is_scalar(x_128)) { lean_ctor_set(x_146, 0, x_140); lean_ctor_set(x_146, 1, x_130); lean_ctor_set(x_146, 2, x_131); -lean_ctor_set(x_146, 3, x_142); +lean_ctor_set(x_146, 3, x_141); lean_ctor_set(x_146, 4, x_145); return x_146; } @@ -62864,8 +62864,8 @@ if (lean_obj_tag(x_133) == 0) lean_object* x_152; x_152 = lean_ctor_get(x_133, 0); lean_inc(x_152); -x_141 = x_151; -x_142 = x_150; +x_141 = x_150; +x_142 = x_151; x_143 = x_152; goto block_147; } @@ -62873,8 +62873,8 @@ else { lean_object* x_153; x_153 = lean_unsigned_to_nat(0u); -x_141 = x_151; -x_142 = x_150; +x_141 = x_150; +x_142 = x_151; x_143 = x_153; goto block_147; } @@ -65161,7 +65161,7 @@ lean_inc_ref(x_25); x_26 = l_Lake_Toml_loadToml(x_25); if (lean_obj_tag(x_26) == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; size_t x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; size_t x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_86; lean_object* x_87; lean_object* x_88; size_t x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_155; lean_object* x_156; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; size_t x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; size_t x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; size_t x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_155; lean_object* x_156; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; x_27 = lean_ctor_get(x_26, 0); lean_inc(x_27); lean_dec_ref(x_26); @@ -65216,23 +65216,23 @@ goto block_159; block_59: { uint8_t x_44; -x_44 = l_Array_isEmpty___redArg(x_28); +x_44 = l_Array_isEmpty___redArg(x_29); if (x_44 == 0) { lean_object* x_45; lean_object* x_46; uint8_t x_47; lean_dec_ref(x_43); lean_dec_ref(x_42); -lean_dec_ref(x_40); -lean_dec(x_39); +lean_dec_ref(x_41); +lean_dec(x_40); +lean_dec_ref(x_38); lean_dec_ref(x_37); -lean_dec_ref(x_36); lean_dec(x_35); -lean_dec(x_34); +lean_dec_ref(x_34); lean_dec_ref(x_33); -lean_dec_ref(x_32); +lean_dec(x_32); lean_dec(x_31); -lean_dec_ref(x_30); -lean_dec(x_29); +lean_dec(x_30); +lean_dec_ref(x_28); lean_dec_ref(x_20); lean_dec_ref(x_19); lean_dec_ref(x_18); @@ -65240,12 +65240,12 @@ lean_dec_ref(x_17); lean_dec_ref(x_16); lean_dec_ref(x_15); lean_dec(x_13); -x_45 = lean_array_get_size(x_28); +x_45 = lean_array_get_size(x_29); x_46 = lean_array_get_size(x_2); -x_47 = lean_nat_dec_lt(x_41, x_45); +x_47 = lean_nat_dec_lt(x_36, x_45); if (x_47 == 0) { -lean_dec_ref(x_28); +lean_dec_ref(x_29); lean_dec_ref(x_25); x_4 = x_46; x_5 = x_2; @@ -65261,7 +65261,7 @@ if (x_49 == 0) { if (x_47 == 0) { -lean_dec_ref(x_28); +lean_dec_ref(x_29); lean_dec_ref(x_25); x_4 = x_46; x_5 = x_2; @@ -65272,8 +65272,8 @@ else { size_t x_50; lean_object* x_51; x_50 = lean_usize_of_nat(x_45); -x_51 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lake_loadTomlConfig_spec__1(x_25, x_44, x_28, x_38, x_50, x_48, x_2); -lean_dec_ref(x_28); +x_51 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lake_loadTomlConfig_spec__1(x_25, x_44, x_29, x_39, x_50, x_48, x_2); +lean_dec_ref(x_29); x_9 = x_46; x_10 = x_51; goto block_12; @@ -65283,8 +65283,8 @@ else { size_t x_52; lean_object* x_53; x_52 = lean_usize_of_nat(x_45); -x_53 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lake_loadTomlConfig_spec__1(x_25, x_44, x_28, x_38, x_52, x_48, x_2); -lean_dec_ref(x_28); +x_53 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lake_loadTomlConfig_spec__1(x_25, x_44, x_29, x_39, x_52, x_48, x_2); +lean_dec_ref(x_29); x_9 = x_46; x_10 = x_53; goto block_12; @@ -65294,33 +65294,33 @@ goto block_12; else { lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -lean_dec_ref(x_28); +lean_dec_ref(x_29); lean_dec_ref(x_25); -x_54 = lean_ctor_get(x_42, 13); +x_54 = lean_ctor_get(x_38, 13); lean_inc_ref(x_54); -x_55 = lean_ctor_get(x_42, 15); +x_55 = lean_ctor_get(x_38, 15); lean_inc_ref(x_55); x_56 = lean_box(0); x_57 = lean_alloc_ctor(0, 24, 0); lean_ctor_set(x_57, 0, x_13); -lean_ctor_set(x_57, 1, x_29); -lean_ctor_set(x_57, 2, x_39); -lean_ctor_set(x_57, 3, x_35); +lean_ctor_set(x_57, 1, x_30); +lean_ctor_set(x_57, 2, x_31); +lean_ctor_set(x_57, 3, x_32); lean_ctor_set(x_57, 4, x_16); lean_ctor_set(x_57, 5, x_15); -lean_ctor_set(x_57, 6, x_42); +lean_ctor_set(x_57, 6, x_38); lean_ctor_set(x_57, 7, x_18); lean_ctor_set(x_57, 8, x_17); -lean_ctor_set(x_57, 9, x_30); +lean_ctor_set(x_57, 9, x_41); lean_ctor_set(x_57, 10, x_19); lean_ctor_set(x_57, 11, x_20); -lean_ctor_set(x_57, 12, x_36); -lean_ctor_set(x_57, 13, x_32); -lean_ctor_set(x_57, 14, x_34); -lean_ctor_set(x_57, 15, x_37); -lean_ctor_set(x_57, 16, x_31); -lean_ctor_set(x_57, 17, x_33); -lean_ctor_set(x_57, 18, x_40); +lean_ctor_set(x_57, 12, x_28); +lean_ctor_set(x_57, 13, x_33); +lean_ctor_set(x_57, 14, x_40); +lean_ctor_set(x_57, 15, x_42); +lean_ctor_set(x_57, 16, x_35); +lean_ctor_set(x_57, 17, x_34); +lean_ctor_set(x_57, 18, x_37); lean_ctor_set(x_57, 19, x_43); lean_ctor_set(x_57, 20, x_54); lean_ctor_set(x_57, 21, x_55); @@ -65339,58 +65339,58 @@ x_72 = l_System_FilePath_normalize(x_71); x_73 = lean_box(1); x_74 = lean_unsigned_to_nat(0u); x_75 = l_Lake_loadTomlConfig___closed__0; -if (lean_obj_tag(x_69) == 1) +if (lean_obj_tag(x_61) == 1) { lean_object* x_76; -x_76 = lean_ctor_get(x_69, 0); +x_76 = lean_ctor_get(x_61, 0); lean_inc(x_76); -lean_dec_ref(x_69); -x_28 = x_67; -x_29 = x_66; -x_30 = x_72; -x_31 = x_73; -x_32 = x_70; -x_33 = x_75; -x_34 = x_60; -x_35 = x_62; -x_36 = x_61; -x_37 = x_63; -x_38 = x_65; -x_39 = x_64; -x_40 = x_75; -x_41 = x_74; -x_42 = x_68; +lean_dec_ref(x_61); +x_28 = x_62; +x_29 = x_63; +x_30 = x_64; +x_31 = x_65; +x_32 = x_68; +x_33 = x_70; +x_34 = x_75; +x_35 = x_73; +x_36 = x_74; +x_37 = x_75; +x_38 = x_60; +x_39 = x_67; +x_40 = x_66; +x_41 = x_72; +x_42 = x_69; x_43 = x_76; goto block_59; } else { uint8_t x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; -lean_dec(x_69); +lean_dec(x_61); x_77 = 0; -lean_inc(x_66); -x_78 = l_Lean_Name_toString(x_66, x_77); +lean_inc(x_64); +x_78 = l_Lean_Name_toString(x_64, x_77); x_79 = l___private_Lake_Load_Toml_0__Lake_decodeTomlConfig___at___00Lake_LeanExeConfig_decodeToml_spec__0___closed__1; x_80 = lean_string_append(x_78, x_79); x_81 = l_Lake_loadTomlConfig___closed__1; x_82 = lean_string_append(x_80, x_81); x_83 = l_Lake_loadTomlConfig___closed__2; x_84 = lean_string_append(x_82, x_83); -x_28 = x_67; -x_29 = x_66; -x_30 = x_72; -x_31 = x_73; -x_32 = x_70; -x_33 = x_75; -x_34 = x_60; -x_35 = x_62; -x_36 = x_61; -x_37 = x_63; -x_38 = x_65; -x_39 = x_64; -x_40 = x_75; -x_41 = x_74; -x_42 = x_68; +x_28 = x_62; +x_29 = x_63; +x_30 = x_64; +x_31 = x_65; +x_32 = x_68; +x_33 = x_70; +x_34 = x_75; +x_35 = x_73; +x_36 = x_74; +x_37 = x_75; +x_38 = x_60; +x_39 = x_67; +x_40 = x_66; +x_41 = x_72; +x_42 = x_69; x_43 = x_84; goto block_59; } @@ -65398,45 +65398,45 @@ goto block_59; block_101: { lean_object* x_96; -x_96 = lean_ctor_get(x_92, 2); +x_96 = lean_ctor_get(x_86, 2); if (lean_obj_tag(x_96) == 0) { lean_object* x_97; lean_object* x_98; -x_97 = lean_ctor_get(x_92, 12); +x_97 = lean_ctor_get(x_86, 12); lean_inc(x_97); x_98 = l_Lake_loadTomlConfig___closed__3; x_60 = x_86; -x_61 = x_94; -x_62 = x_87; -x_63 = x_88; -x_64 = x_90; -x_65 = x_89; +x_61 = x_97; +x_62 = x_94; +x_63 = x_95; +x_64 = x_87; +x_65 = x_88; x_66 = x_91; -x_67 = x_95; -x_68 = x_92; -x_69 = x_97; -x_70 = x_93; +x_67 = x_90; +x_68 = x_89; +x_69 = x_93; +x_70 = x_92; x_71 = x_98; goto block_85; } else { lean_object* x_99; lean_object* x_100; -x_99 = lean_ctor_get(x_92, 12); +x_99 = lean_ctor_get(x_86, 12); lean_inc(x_99); x_100 = lean_ctor_get(x_96, 0); lean_inc(x_100); x_60 = x_86; -x_61 = x_94; -x_62 = x_87; -x_63 = x_88; -x_64 = x_90; -x_65 = x_89; +x_61 = x_99; +x_62 = x_94; +x_63 = x_95; +x_64 = x_87; +x_65 = x_88; x_66 = x_91; -x_67 = x_95; -x_68 = x_92; -x_69 = x_99; -x_70 = x_93; +x_67 = x_90; +x_68 = x_89; +x_69 = x_93; +x_70 = x_92; x_71 = x_100; goto block_85; } @@ -65455,12 +65455,12 @@ if (lean_obj_tag(x_116) == 0) { x_86 = x_102; x_87 = x_103; -x_88 = x_112; -x_89 = x_111; -x_90 = x_104; +x_88 = x_104; +x_89 = x_106; +x_90 = x_111; x_91 = x_105; -x_92 = x_106; -x_93 = x_107; +x_92 = x_107; +x_93 = x_112; x_94 = x_114; x_95 = x_109; goto block_101; @@ -65495,12 +65495,12 @@ lean_inc(x_124); lean_dec_ref(x_122); x_86 = x_102; x_87 = x_103; -x_88 = x_112; -x_89 = x_111; -x_90 = x_104; +x_88 = x_104; +x_89 = x_106; +x_90 = x_111; x_91 = x_105; -x_92 = x_106; -x_93 = x_107; +x_92 = x_107; +x_93 = x_112; x_94 = x_123; x_95 = x_124; goto block_101; @@ -65513,12 +65513,12 @@ lean_inc(x_125); lean_dec_ref(x_122); x_86 = x_102; x_87 = x_103; -x_88 = x_112; -x_89 = x_111; -x_90 = x_104; +x_88 = x_104; +x_89 = x_106; +x_90 = x_111; x_91 = x_105; -x_92 = x_106; -x_93 = x_107; +x_92 = x_107; +x_93 = x_112; x_94 = x_114; x_95 = x_125; goto block_101; @@ -65532,12 +65532,12 @@ lean_inc(x_126); lean_dec_ref(x_119); x_86 = x_102; x_87 = x_103; -x_88 = x_112; -x_89 = x_111; -x_90 = x_104; +x_88 = x_104; +x_89 = x_106; +x_90 = x_111; x_91 = x_105; -x_92 = x_106; -x_93 = x_107; +x_92 = x_107; +x_93 = x_112; x_94 = x_114; x_95 = x_126; goto block_101; @@ -65579,11 +65579,11 @@ lean_inc(x_27); x_143 = l_Lake_Toml_RBDict_findEntry_x3f___redArg(x_142, x_140, x_27); if (lean_obj_tag(x_143) == 0) { -x_102 = x_139; -x_103 = x_128; +x_102 = x_133; +x_103 = x_130; x_104 = x_131; -x_105 = x_130; -x_106 = x_133; +x_105 = x_139; +x_106 = x_128; x_107 = x_138; x_108 = x_141; x_109 = x_137; @@ -65617,11 +65617,11 @@ lean_inc(x_150); x_151 = lean_ctor_get(x_149, 1); lean_inc(x_151); lean_dec_ref(x_149); -x_102 = x_139; -x_103 = x_128; +x_102 = x_133; +x_103 = x_130; x_104 = x_131; -x_105 = x_130; -x_106 = x_133; +x_105 = x_139; +x_106 = x_128; x_107 = x_138; x_108 = x_150; x_109 = x_151; @@ -65633,11 +65633,11 @@ lean_object* x_152; x_152 = lean_ctor_get(x_149, 1); lean_inc(x_152); lean_dec_ref(x_149); -x_102 = x_139; -x_103 = x_128; +x_102 = x_133; +x_103 = x_130; x_104 = x_131; -x_105 = x_130; -x_106 = x_133; +x_105 = x_139; +x_106 = x_128; x_107 = x_138; x_108 = x_141; x_109 = x_152; @@ -65650,11 +65650,11 @@ lean_object* x_153; x_153 = lean_ctor_get(x_146, 1); lean_inc(x_153); lean_dec_ref(x_146); -x_102 = x_139; -x_103 = x_128; +x_102 = x_133; +x_103 = x_130; x_104 = x_131; -x_105 = x_130; -x_106 = x_133; +x_105 = x_139; +x_106 = x_128; x_107 = x_138; x_108 = x_141; x_109 = x_153; diff --git a/stage0/stdlib/Lean/Compiler/IR.c b/stage0/stdlib/Lean/Compiler/IR.c index 586c01ba3f7b..ed1b7e1f6a5b 100644 --- a/stage0/stdlib/Lean/Compiler/IR.c +++ b/stage0/stdlib/Lean/Compiler/IR.c @@ -16,23 +16,26 @@ extern "C" { static lean_object* l_Lean_IR_initFn___closed__5_00___x40_Lean_Compiler_IR_583794373____hygCtx___hyg_4_; static lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_640659120____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__1___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__5___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_583794373____hygCtx___hyg_4_; lean_object* l_Lean_IR_Decl_elimDead(lean_object*); lean_object* l_Lean_IR_Decl_simpCase(lean_object*); static lean_object* l_Lean_IR_initFn___closed__6_00___x40_Lean_Compiler_IR_583794373____hygCtx___hyg_4_; lean_object* l_Lean_IR_inferBorrow___redArg(lean_object*, lean_object*); static lean_object* l_Lean_IR_compile___closed__22; +static lean_object* l_Lean_IR_compile___closed__37; static lean_object* l_Lean_IR_compile___closed__13; LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_initFn_00___x40_Lean_Compiler_IR_640659120____hygCtx___hyg_2____boxed(lean_object*); static lean_object* l_Lean_IR_compile___closed__28; static lean_object* l_Lean_IR_initFn___closed__4_00___x40_Lean_Compiler_IR_583794373____hygCtx___hyg_4_; static lean_object* l_Lean_IR_compile___closed__24; static lean_object* l_Lean_IR_compile___closed__17; +lean_object* l_Lean_IR_log(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_KVMap_find(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_initFn___closed__5_00___x40_Lean_Compiler_IR_640659120____hygCtx___hyg_2_; lean_object* l_Lean_IR_addDecls(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_compile___closed__2; +static lean_object* l_Lean_IR_compile___closed__39; +static lean_object* l_Lean_IR_compile___closed__35; static lean_object* l_Lean_IR_compile___closed__21; static lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_initFn___closed__18_00___x40_Lean_Compiler_IR_640659120____hygCtx___hyg_2_; lean_object* l_Lean_IR_Decl_expandResetReuse(lean_object*); @@ -42,14 +45,15 @@ static lean_object* l_Lean_IR_compile___closed__31; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Option_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_583794373____hygCtx___hyg_4__spec__0(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_initFn___closed__20_00___x40_Lean_Compiler_IR_640659120____hygCtx___hyg_2_; -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__5(size_t, size_t, lean_object*); static lean_object* l_Lean_IR_compile___closed__3; lean_object* l___private_Lean_Compiler_IR_CompilerM_0__Lean_IR_logDeclsAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_compile___closed__11; static lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_initFn___closed__25_00___x40_Lean_Compiler_IR_640659120____hygCtx___hyg_2_; static lean_object* l_Lean_IR_compile___closed__26; lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*); +static lean_object* l_Lean_IR_compile___closed__34; static lean_object* l_Lean_IR_compile___closed__10; +static lean_object* l_Lean_IR_compile___closed__38; static lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_initFn___closed__12_00___x40_Lean_Compiler_IR_640659120____hygCtx___hyg_2_; static lean_object* l_Lean_IR_compile___closed__14; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__2___boxed(lean_object*, lean_object*, lean_object*); @@ -58,24 +62,29 @@ lean_object* lean_st_ref_get(lean_object*); lean_object* l_Lean_IR_explicitBoxing___redArg(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_initFn___closed__10_00___x40_Lean_Compiler_IR_640659120____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__3(size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__6(lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__6(size_t, size_t, lean_object*); +static lean_object* l_Lean_IR_compile___closed__36; static lean_object* l_Lean_IR_compile___closed__0; static lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_initFn___closed__6_00___x40_Lean_Compiler_IR_640659120____hygCtx___hyg_2_; +lean_object* l_Lean_IR_StructRC_visitDecl(lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_initFn___closed__16_00___x40_Lean_Compiler_IR_640659120____hygCtx___hyg_2_; static lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_640659120____hygCtx___hyg_2_; static lean_object* l_Lean_IR_compile___closed__18; static lean_object* l_Lean_IR_compile___closed__30; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__4___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_compile___closed__20; static lean_object* l_Lean_IR_compile___closed__5; static lean_object* l_Lean_IR_compile___closed__19; static lean_object* l_Lean_IR_initFn___closed__3_00___x40_Lean_Compiler_IR_583794373____hygCtx___hyg_4_; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_initFn___closed__3_00___x40_Lean_Compiler_IR_640659120____hygCtx___hyg_2_; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__7(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_IR_compile___closed__16; static lean_object* l_Lean_IR_compile___closed__23; -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_compile___closed__40; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__6___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_initFn___closed__13_00___x40_Lean_Compiler_IR_640659120____hygCtx___hyg_2_; lean_object* l_Lean_IR_checkDecls(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__1(size_t, size_t, lean_object*); @@ -87,7 +96,6 @@ static lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_initFn___closed__24_ static lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_initFn___closed__17_00___x40_Lean_Compiler_IR_640659120____hygCtx___hyg_2_; lean_object* l_Lean_IR_Decl_insertResetReuse(lean_object*, lean_object*); static lean_object* l_Lean_IR_compile___closed__25; -LEAN_EXPORT lean_object* l_Lean_Option_get___at___00Lean_IR_compile_spec__4___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__2(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Option_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_583794373____hygCtx___hyg_4__spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_initFn___closed__21_00___x40_Lean_Compiler_IR_640659120____hygCtx___hyg_2_; @@ -95,13 +103,19 @@ static lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_initFn___closed__11_ lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_initFn___closed__14_00___x40_Lean_Compiler_IR_640659120____hygCtx___hyg_2_; static lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_initFn___closed__4_00___x40_Lean_Compiler_IR_640659120____hygCtx___hyg_2_; +uint8_t l_Lean_KVMap_getBool(lean_object*, lean_object*, uint8_t); +LEAN_EXPORT uint8_t l_Lean_Option_get___at___00Lean_IR_compile_spec__5(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_initFn___closed__9_00___x40_Lean_Compiler_IR_640659120____hygCtx___hyg_2_; static lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_initFn___closed__15_00___x40_Lean_Compiler_IR_640659120____hygCtx___hyg_2_; extern lean_object* l_Lean_IR_tracePrefixOptionName; lean_object* l_Lean_IR_explicitRC___redArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Option_get___at___00Lean_IR_compile_spec__5___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_initFn_00___x40_Lean_Compiler_IR_583794373____hygCtx___hyg_4____boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_compile(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__0___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_IR_ExplicitBoxing_addBoxedVersions(lean_object*, lean_object*); +lean_object* l_Lean_IR_prepareBoxParams(lean_object*, lean_object*); static lean_object* l_Lean_IR_compile___closed__7; static lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_initFn___closed__22_00___x40_Lean_Compiler_IR_640659120____hygCtx___hyg_2_; size_t lean_usize_add(size_t, size_t); @@ -120,12 +134,13 @@ lean_object* l_Lean_Name_mkStr1(lean_object*); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_583794373____hygCtx___hyg_4_; lean_object* l_Lean_IR_updateSorryDep(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Lean_Option_get___at___00Lean_IR_compile_spec__4(lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_Lean_IR_compile___closed__4; lean_object* l_Lean_IR_Decl_normalizeIds(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_initFn_00___x40_Lean_Compiler_IR_583794373____hygCtx___hyg_4_(); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__4(size_t, size_t, lean_object*); static lean_object* l_Lean_IR_compile___closed__1; +static lean_object* l_Lean_IR_compile___closed__33; static lean_object* l_Lean_IR_compile___closed__6; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); static lean_object* l_Lean_IR_compile___closed__12; @@ -133,6 +148,7 @@ static lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_initFn___closed__8_0 static lean_object* l_Lean_IR_initFn___closed__7_00___x40_Lean_Compiler_IR_583794373____hygCtx___hyg_4_; static lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_initFn___closed__19_00___x40_Lean_Compiler_IR_640659120____hygCtx___hyg_2_; static lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_initFn___closed__2_00___x40_Lean_Compiler_IR_640659120____hygCtx___hyg_2_; +static lean_object* l_Lean_IR_compile___closed__32; static lean_object* l_Lean_IR_compile___closed__29; static lean_object* _init_l_Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_583794373____hygCtx___hyg_4_() { _start: @@ -370,7 +386,7 @@ x_5 = l_Lean_Option_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_5 return x_5; } } -LEAN_EXPORT uint8_t l_Lean_Option_get___at___00Lean_IR_compile_spec__4(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l_Lean_Option_get___at___00Lean_IR_compile_spec__5(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; @@ -446,7 +462,7 @@ lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x x_5 = lean_array_uget(x_3, x_2); x_6 = lean_unsigned_to_nat(0u); x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = l_Lean_IR_Decl_elimDead(x_5); +x_8 = l_Lean_IR_StructRC_visitDecl(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); @@ -506,7 +522,32 @@ goto _start; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__5(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__4(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = l_Lean_IR_Decl_elimDead(x_5); +x_9 = 1; +x_10 = lean_usize_add(x_2, x_9); +x_11 = lean_array_uset(x_7, x_2, x_8); +x_2 = x_10; +x_3 = x_11; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__6(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -531,7 +572,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__6(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__7(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -651,7 +692,7 @@ static lean_object* _init_l_Lean_IR_compile___closed__10() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("elim_dead", 9, 9); +x_1 = lean_mk_string_unchecked("struct_rc", 9, 9); return x_1; } } @@ -705,7 +746,7 @@ static lean_object* _init_l_Lean_IR_compile___closed__16() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("borrow", 6, 6); +x_1 = lean_mk_string_unchecked("elim_dead", 9, 9); return x_1; } } @@ -732,7 +773,7 @@ static lean_object* _init_l_Lean_IR_compile___closed__19() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("boxing", 6, 6); +x_1 = lean_mk_string_unchecked("borrow", 6, 6); return x_1; } } @@ -759,7 +800,7 @@ static lean_object* _init_l_Lean_IR_compile___closed__22() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("rc", 2, 2); +x_1 = lean_mk_string_unchecked("boxed_versions", 14, 14); return x_1; } } @@ -786,38 +827,119 @@ static lean_object* _init_l_Lean_IR_compile___closed__25() { _start: { lean_object* x_1; -x_1 = l_Lean_IR_compiler_reuse; +x_1 = lean_mk_string_unchecked("rc", 2, 2); return x_1; } } static lean_object* _init_l_Lean_IR_compile___closed__26() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_IR_compile___closed__25; +x_2 = l_Lean_Name_mkStr1(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_IR_compile___closed__27() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_compile___closed__26; +x_2 = l_Lean_IR_compile___closed__2; +x_3 = l_Lean_Name_append(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_IR_compile___closed__28() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_IR_compiler_reuse; +return x_1; +} +} +static lean_object* _init_l_Lean_IR_compile___closed__29() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_unchecked("expand_reset_reuse", 18, 18); return x_1; } } -static lean_object* _init_l_Lean_IR_compile___closed__27() { +static lean_object* _init_l_Lean_IR_compile___closed__30() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_IR_compile___closed__26; +x_1 = l_Lean_IR_compile___closed__29; x_2 = l_Lean_Name_mkStr1(x_1); return x_2; } } -static lean_object* _init_l_Lean_IR_compile___closed__28() { +static lean_object* _init_l_Lean_IR_compile___closed__31() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_IR_compile___closed__27; +x_1 = l_Lean_IR_compile___closed__30; x_2 = l_Lean_IR_compile___closed__2; x_3 = l_Lean_Name_append(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_IR_compile___closed__29() { +static lean_object* _init_l_Lean_IR_compile___closed__32() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("box_params", 10, 10); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_compile___closed__33() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_IR_compile___closed__32; +x_2 = l_Lean_Name_mkStr1(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_IR_compile___closed__34() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_compile___closed__33; +x_2 = l_Lean_IR_compile___closed__2; +x_3 = l_Lean_Name_append(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_IR_compile___closed__35() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("boxing", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_compile___closed__36() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_IR_compile___closed__35; +x_2 = l_Lean_Name_mkStr1(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_IR_compile___closed__37() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_compile___closed__36; +x_2 = l_Lean_IR_compile___closed__2; +x_3 = l_Lean_Name_append(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_IR_compile___closed__38() { _start: { lean_object* x_1; @@ -825,20 +947,20 @@ x_1 = lean_mk_string_unchecked("reset_reuse", 11, 11); return x_1; } } -static lean_object* _init_l_Lean_IR_compile___closed__30() { +static lean_object* _init_l_Lean_IR_compile___closed__39() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_IR_compile___closed__29; +x_1 = l_Lean_IR_compile___closed__38; x_2 = l_Lean_Name_mkStr1(x_1); return x_2; } } -static lean_object* _init_l_Lean_IR_compile___closed__31() { +static lean_object* _init_l_Lean_IR_compile___closed__40() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_IR_compile___closed__30; +x_1 = l_Lean_IR_compile___closed__39; x_2 = l_Lean_IR_compile___closed__2; x_3 = l_Lean_Name_append(x_2, x_1); return x_3; @@ -847,622 +969,780 @@ return x_3; LEAN_EXPORT lean_object* l_Lean_IR_compile(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = l_Lean_IR_compile___closed__1; -x_6 = l_Lean_IR_compile___closed__3; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_23 = l_Lean_IR_compile___closed__1; +x_24 = l_Lean_IR_compile___closed__2; +x_25 = l_Lean_IR_compile___closed__3; lean_inc_ref(x_1); -x_7 = l___private_Lean_Compiler_IR_CompilerM_0__Lean_IR_logDeclsAux(x_6, x_5, x_1, x_2, x_3); -if (lean_obj_tag(x_7) == 0) +x_26 = l___private_Lean_Compiler_IR_CompilerM_0__Lean_IR_logDeclsAux(x_25, x_23, x_1, x_2, x_3); +if (lean_obj_tag(x_26) == 0) { -lean_object* x_8; -lean_dec_ref(x_7); +lean_object* x_27; +lean_dec_ref(x_26); lean_inc_ref(x_1); -x_8 = l_Lean_IR_checkDecls(x_1, x_2, x_3); -if (lean_obj_tag(x_8) == 0) +x_27 = l_Lean_IR_checkDecls(x_1, x_2, x_3); +if (lean_obj_tag(x_27) == 0) { -size_t x_9; size_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_108; lean_object* x_109; -lean_dec_ref(x_8); -x_9 = lean_array_size(x_1); -x_10 = 0; -x_11 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__0(x_9, x_10, x_1); -x_12 = l_Lean_IR_compile___closed__5; -x_108 = l_Lean_IR_compile___closed__6; -lean_inc_ref(x_11); -x_109 = l___private_Lean_Compiler_IR_CompilerM_0__Lean_IR_logDeclsAux(x_108, x_12, x_11, x_2, x_3); -if (lean_obj_tag(x_109) == 0) -{ -lean_object* x_110; lean_object* x_111; uint8_t x_112; -lean_dec_ref(x_109); -x_110 = lean_ctor_get(x_2, 2); -x_111 = l_Lean_IR_compile___closed__25; -x_112 = l_Lean_Option_get___at___00Lean_IR_compile_spec__4(x_110, x_111); -if (x_112 == 0) -{ -x_50 = x_11; -x_51 = x_2; -x_52 = x_3; -x_53 = lean_box(0); -goto block_107; +size_t x_28; size_t x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_135; lean_object* x_136; +lean_dec_ref(x_27); +x_28 = lean_array_size(x_1); +x_29 = 0; +x_44 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__0(x_28, x_29, x_1); +x_45 = l_Lean_IR_compile___closed__5; +x_135 = l_Lean_IR_compile___closed__6; +lean_inc_ref(x_44); +x_136 = l___private_Lean_Compiler_IR_CompilerM_0__Lean_IR_logDeclsAux(x_135, x_45, x_44, x_2, x_3); +if (lean_obj_tag(x_136) == 0) +{ +lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; +lean_dec_ref(x_136); +x_137 = lean_st_ref_get(x_3); +x_138 = lean_ctor_get(x_137, 0); +lean_inc_ref(x_138); +lean_dec(x_137); +x_139 = l_Lean_IR_prepareBoxParams(x_138, x_44); +x_140 = l_Lean_IR_compile___closed__33; +x_141 = l_Lean_IR_compile___closed__34; +lean_inc_ref(x_139); +x_142 = l___private_Lean_Compiler_IR_CompilerM_0__Lean_IR_logDeclsAux(x_141, x_140, x_139, x_2, x_3); +if (lean_obj_tag(x_142) == 0) +{ +lean_object* x_143; +lean_dec_ref(x_142); +x_143 = l_Lean_IR_explicitBoxing___redArg(x_139, x_3); +if (lean_obj_tag(x_143) == 0) +{ +lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; +x_144 = lean_ctor_get(x_143, 0); +lean_inc(x_144); +lean_dec_ref(x_143); +x_145 = l_Lean_IR_compile___closed__36; +x_146 = l_Lean_IR_compile___closed__37; +lean_inc(x_144); +x_147 = l___private_Lean_Compiler_IR_CompilerM_0__Lean_IR_logDeclsAux(x_146, x_145, x_144, x_2, x_3); +if (lean_obj_tag(x_147) == 0) +{ +lean_object* x_148; lean_object* x_149; uint8_t x_150; +lean_dec_ref(x_147); +x_148 = lean_ctor_get(x_2, 2); +x_149 = l_Lean_IR_compile___closed__28; +x_150 = l_Lean_Option_get___at___00Lean_IR_compile_spec__5(x_148, x_149); +if (x_150 == 0) +{ +x_76 = x_144; +x_77 = x_2; +x_78 = x_3; +x_79 = lean_box(0); +goto block_134; } else { -lean_object* x_113; lean_object* x_114; size_t x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; -x_113 = lean_st_ref_get(x_3); -x_114 = lean_ctor_get(x_113, 0); -lean_inc_ref(x_114); -lean_dec(x_113); -x_115 = lean_array_size(x_11); -x_116 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__6(x_114, x_115, x_10, x_11); -x_117 = l_Lean_IR_compile___closed__30; -x_118 = l_Lean_IR_compile___closed__31; -lean_inc_ref(x_116); -x_119 = l___private_Lean_Compiler_IR_CompilerM_0__Lean_IR_logDeclsAux(x_118, x_117, x_116, x_2, x_3); -if (lean_obj_tag(x_119) == 0) -{ -lean_dec_ref(x_119); -x_50 = x_116; -x_51 = x_2; -x_52 = x_3; -x_53 = lean_box(0); -goto block_107; +lean_object* x_151; lean_object* x_152; size_t x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; +x_151 = lean_st_ref_get(x_3); +x_152 = lean_ctor_get(x_151, 0); +lean_inc_ref(x_152); +lean_dec(x_151); +x_153 = lean_array_size(x_144); +x_154 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__7(x_152, x_153, x_29, x_144); +x_155 = l_Lean_IR_compile___closed__39; +x_156 = l_Lean_IR_compile___closed__40; +lean_inc_ref(x_154); +x_157 = l___private_Lean_Compiler_IR_CompilerM_0__Lean_IR_logDeclsAux(x_156, x_155, x_154, x_2, x_3); +if (lean_obj_tag(x_157) == 0) +{ +lean_dec_ref(x_157); +x_76 = x_154; +x_77 = x_2; +x_78 = x_3; +x_79 = lean_box(0); +goto block_134; } else { -uint8_t x_120; -lean_dec_ref(x_116); +uint8_t x_158; +lean_dec_ref(x_154); lean_dec(x_3); lean_dec_ref(x_2); -x_120 = !lean_is_exclusive(x_119); -if (x_120 == 0) +x_158 = !lean_is_exclusive(x_157); +if (x_158 == 0) { -return x_119; +return x_157; } else { -lean_object* x_121; lean_object* x_122; -x_121 = lean_ctor_get(x_119, 0); -lean_inc(x_121); -lean_dec(x_119); -x_122 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_122, 0, x_121); -return x_122; +lean_object* x_159; lean_object* x_160; +x_159 = lean_ctor_get(x_157, 0); +lean_inc(x_159); +lean_dec(x_157); +x_160 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_160, 0, x_159); +return x_160; } } } } else { -uint8_t x_123; -lean_dec_ref(x_11); +uint8_t x_161; +lean_dec(x_144); lean_dec(x_3); lean_dec_ref(x_2); -x_123 = !lean_is_exclusive(x_109); -if (x_123 == 0) +x_161 = !lean_is_exclusive(x_147); +if (x_161 == 0) { -return x_109; +return x_147; } else { -lean_object* x_124; lean_object* x_125; -x_124 = lean_ctor_get(x_109, 0); -lean_inc(x_124); -lean_dec(x_109); -x_125 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_125, 0, x_124); -return x_125; -} -} -block_49: -{ -size_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_17 = lean_array_size(x_13); -x_18 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__0(x_17, x_10, x_13); -x_19 = l_Lean_IR_compile___closed__6; -lean_inc_ref(x_18); -x_20 = l___private_Lean_Compiler_IR_CompilerM_0__Lean_IR_logDeclsAux(x_19, x_12, x_18, x_14, x_15); -if (lean_obj_tag(x_20) == 0) -{ -lean_object* x_21; -lean_dec_ref(x_20); -x_21 = l_Lean_IR_updateSorryDep(x_18, x_14, x_15); -if (lean_obj_tag(x_21) == 0) -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -lean_dec_ref(x_21); -x_23 = l_Lean_IR_compile___closed__8; -x_24 = l_Lean_IR_compile___closed__9; -lean_inc(x_22); -x_25 = l___private_Lean_Compiler_IR_CompilerM_0__Lean_IR_logDeclsAux(x_24, x_23, x_22, x_14, x_15); -if (lean_obj_tag(x_25) == 0) -{ -lean_object* x_26; -lean_dec_ref(x_25); -lean_inc(x_22); -x_26 = l_Lean_IR_checkDecls(x_22, x_14, x_15); -if (lean_obj_tag(x_26) == 0) -{ -lean_object* x_27; -lean_dec_ref(x_26); -lean_inc(x_15); -lean_inc_ref(x_14); -x_27 = l_Lean_IR_toposortDecls(x_22, x_14, x_15); -lean_dec(x_22); -if (lean_obj_tag(x_27) == 0) -{ -lean_object* x_28; lean_object* x_29; -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -lean_dec_ref(x_27); -x_29 = l_Lean_IR_addDecls(x_28, x_14, x_15); -if (lean_obj_tag(x_29) == 0) -{ -lean_object* x_30; -lean_dec_ref(x_29); -x_30 = l_Lean_IR_inferMeta(x_28, x_14, x_15); -lean_dec(x_15); -lean_dec_ref(x_14); -if (lean_obj_tag(x_30) == 0) -{ -uint8_t x_31; -x_31 = !lean_is_exclusive(x_30); -if (x_31 == 0) -{ -lean_object* x_32; -x_32 = lean_ctor_get(x_30, 0); -lean_dec(x_32); -lean_ctor_set(x_30, 0, x_28); -return x_30; +lean_object* x_162; lean_object* x_163; +x_162 = lean_ctor_get(x_147, 0); +lean_inc(x_162); +lean_dec(x_147); +x_163 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_163, 0, x_162); +return x_163; +} +} } else { -lean_object* x_33; -lean_dec(x_30); -x_33 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_33, 0, x_28); -return x_33; +lean_dec(x_3); +lean_dec_ref(x_2); +return x_143; } } else { -uint8_t x_34; -lean_dec(x_28); -x_34 = !lean_is_exclusive(x_30); -if (x_34 == 0) +uint8_t x_164; +lean_dec_ref(x_139); +lean_dec(x_3); +lean_dec_ref(x_2); +x_164 = !lean_is_exclusive(x_142); +if (x_164 == 0) { -return x_30; +return x_142; } else { -lean_object* x_35; lean_object* x_36; -x_35 = lean_ctor_get(x_30, 0); -lean_inc(x_35); -lean_dec(x_30); -x_36 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_36, 0, x_35); -return x_36; +lean_object* x_165; lean_object* x_166; +x_165 = lean_ctor_get(x_142, 0); +lean_inc(x_165); +lean_dec(x_142); +x_166 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_166, 0, x_165); +return x_166; } } } else { -uint8_t x_37; -lean_dec(x_28); -lean_dec(x_15); -lean_dec_ref(x_14); -x_37 = !lean_is_exclusive(x_29); -if (x_37 == 0) +uint8_t x_167; +lean_dec_ref(x_44); +lean_dec(x_3); +lean_dec_ref(x_2); +x_167 = !lean_is_exclusive(x_136); +if (x_167 == 0) { -return x_29; +return x_136; } else { -lean_object* x_38; lean_object* x_39; -x_38 = lean_ctor_get(x_29, 0); -lean_inc(x_38); -lean_dec(x_29); -x_39 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_39, 0, x_38); -return x_39; +lean_object* x_168; lean_object* x_169; +x_168 = lean_ctor_get(x_136, 0); +lean_inc(x_168); +lean_dec(x_136); +x_169 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_169, 0, x_168); +return x_169; } } +block_43: +{ +if (x_35 == 0) +{ +lean_dec(x_33); +x_5 = x_34; +x_6 = x_30; +x_7 = x_31; +x_8 = lean_box(0); +goto block_22; } else { -lean_dec(x_15); -lean_dec_ref(x_14); -return x_27; -} +size_t x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_36 = lean_array_size(x_34); +lean_inc_ref(x_34); +x_37 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__1(x_36, x_29, x_34); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_33); +lean_ctor_set(x_38, 1, x_37); +x_39 = l_Lean_IR_log(x_38, x_30, x_31); +if (lean_obj_tag(x_39) == 0) +{ +lean_dec_ref(x_39); +x_5 = x_34; +x_6 = x_30; +x_7 = x_31; +x_8 = lean_box(0); +goto block_22; } else { uint8_t x_40; -lean_dec(x_22); -lean_dec(x_15); -lean_dec_ref(x_14); -x_40 = !lean_is_exclusive(x_26); +lean_dec_ref(x_34); +lean_dec(x_31); +lean_dec_ref(x_30); +x_40 = !lean_is_exclusive(x_39); if (x_40 == 0) { -return x_26; +return x_39; } else { lean_object* x_41; lean_object* x_42; -x_41 = lean_ctor_get(x_26, 0); +x_41 = lean_ctor_get(x_39, 0); lean_inc(x_41); -lean_dec(x_26); +lean_dec(x_39); x_42 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_42, 0, x_41); return x_42; } } } -else +} +block_75: { -uint8_t x_43; -lean_dec(x_22); -lean_dec(x_15); -lean_dec_ref(x_14); -x_43 = !lean_is_exclusive(x_25); -if (x_43 == 0) +size_t x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_50 = lean_array_size(x_46); +x_51 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__0(x_50, x_29, x_46); +x_52 = l_Lean_IR_compile___closed__6; +lean_inc_ref(x_51); +x_53 = l___private_Lean_Compiler_IR_CompilerM_0__Lean_IR_logDeclsAux(x_52, x_45, x_51, x_47, x_48); +if (lean_obj_tag(x_53) == 0) { -return x_25; +lean_object* x_54; +lean_dec_ref(x_53); +x_54 = l_Lean_IR_updateSorryDep(x_51, x_47, x_48); +if (lean_obj_tag(x_54) == 0) +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +lean_dec_ref(x_54); +x_56 = l_Lean_IR_compile___closed__8; +x_57 = l_Lean_IR_compile___closed__9; +lean_inc(x_55); +x_58 = l___private_Lean_Compiler_IR_CompilerM_0__Lean_IR_logDeclsAux(x_57, x_56, x_55, x_47, x_48); +if (lean_obj_tag(x_58) == 0) +{ +lean_object* x_59; +lean_dec_ref(x_58); +lean_inc(x_55); +x_59 = l_Lean_IR_checkDecls(x_55, x_47, x_48); +if (lean_obj_tag(x_59) == 0) +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; uint8_t x_64; +lean_dec_ref(x_59); +x_60 = lean_ctor_get(x_47, 2); +x_61 = l_Lean_IR_compile___closed__11; +x_62 = l_Lean_IR_compile___closed__12; +x_63 = 0; +x_64 = l_Lean_KVMap_getBool(x_60, x_62, x_63); +if (x_64 == 0) +{ +uint8_t x_65; +x_65 = l_Lean_KVMap_getBool(x_60, x_24, x_64); +x_30 = x_47; +x_31 = x_48; +x_32 = lean_box(0); +x_33 = x_61; +x_34 = x_55; +x_35 = x_65; +goto block_43; } else { -lean_object* x_44; lean_object* x_45; -x_44 = lean_ctor_get(x_25, 0); -lean_inc(x_44); -lean_dec(x_25); -x_45 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_45, 0, x_44); -return x_45; +x_30 = x_47; +x_31 = x_48; +x_32 = lean_box(0); +x_33 = x_61; +x_34 = x_55; +x_35 = x_64; +goto block_43; } } +else +{ +uint8_t x_66; +lean_dec(x_55); +lean_dec(x_48); +lean_dec_ref(x_47); +x_66 = !lean_is_exclusive(x_59); +if (x_66 == 0) +{ +return x_59; } else { -lean_dec(x_15); -lean_dec_ref(x_14); -return x_21; +lean_object* x_67; lean_object* x_68; +x_67 = lean_ctor_get(x_59, 0); +lean_inc(x_67); +lean_dec(x_59); +x_68 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_68, 0, x_67); +return x_68; +} } } else { -uint8_t x_46; -lean_dec_ref(x_18); -lean_dec(x_15); -lean_dec_ref(x_14); -x_46 = !lean_is_exclusive(x_20); -if (x_46 == 0) +uint8_t x_69; +lean_dec(x_55); +lean_dec(x_48); +lean_dec_ref(x_47); +x_69 = !lean_is_exclusive(x_58); +if (x_69 == 0) { -return x_20; +return x_58; } else { -lean_object* x_47; lean_object* x_48; -x_47 = lean_ctor_get(x_20, 0); -lean_inc(x_47); -lean_dec(x_20); -x_48 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_48, 0, x_47); -return x_48; +lean_object* x_70; lean_object* x_71; +x_70 = lean_ctor_get(x_58, 0); +lean_inc(x_70); +lean_dec(x_58); +x_71 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_71, 0, x_70); +return x_71; } } } -block_107: +else { -size_t x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_54 = lean_array_size(x_50); -x_55 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__1(x_54, x_10, x_50); -x_56 = l_Lean_IR_compile___closed__11; -x_57 = l_Lean_IR_compile___closed__12; -lean_inc_ref(x_55); -x_58 = l___private_Lean_Compiler_IR_CompilerM_0__Lean_IR_logDeclsAux(x_57, x_56, x_55, x_51, x_52); -if (lean_obj_tag(x_58) == 0) +lean_dec(x_48); +lean_dec_ref(x_47); +return x_54; +} +} +else { -size_t x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; -lean_dec_ref(x_58); -x_59 = lean_array_size(x_55); -x_60 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__2(x_59, x_10, x_55); -x_61 = l_Lean_IR_compile___closed__14; -x_62 = l_Lean_IR_compile___closed__15; -lean_inc_ref(x_60); -x_63 = l___private_Lean_Compiler_IR_CompilerM_0__Lean_IR_logDeclsAux(x_62, x_61, x_60, x_51, x_52); -if (lean_obj_tag(x_63) == 0) -{ -size_t x_64; lean_object* x_65; lean_object* x_66; -lean_dec_ref(x_63); -x_64 = lean_array_size(x_60); -x_65 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__3(x_64, x_10, x_60); -x_66 = l_Lean_IR_inferBorrow___redArg(x_65, x_52); -if (lean_obj_tag(x_66) == 0) -{ -lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_67 = lean_ctor_get(x_66, 0); -lean_inc(x_67); -lean_dec_ref(x_66); -x_68 = l_Lean_IR_compile___closed__17; -x_69 = l_Lean_IR_compile___closed__18; -lean_inc(x_67); -x_70 = l___private_Lean_Compiler_IR_CompilerM_0__Lean_IR_logDeclsAux(x_69, x_68, x_67, x_51, x_52); -if (lean_obj_tag(x_70) == 0) -{ -lean_object* x_71; -lean_dec_ref(x_70); -x_71 = l_Lean_IR_explicitBoxing___redArg(x_67, x_52); -if (lean_obj_tag(x_71) == 0) -{ -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_72 = lean_ctor_get(x_71, 0); -lean_inc(x_72); -lean_dec_ref(x_71); -x_73 = l_Lean_IR_compile___closed__20; -x_74 = l_Lean_IR_compile___closed__21; -lean_inc(x_72); -x_75 = l___private_Lean_Compiler_IR_CompilerM_0__Lean_IR_logDeclsAux(x_74, x_73, x_72, x_51, x_52); -if (lean_obj_tag(x_75) == 0) -{ -lean_object* x_76; -lean_dec_ref(x_75); -x_76 = l_Lean_IR_explicitRC___redArg(x_72, x_52); -if (lean_obj_tag(x_76) == 0) -{ -lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_77 = lean_ctor_get(x_76, 0); -lean_inc(x_77); -lean_dec_ref(x_76); -x_78 = l_Lean_IR_compile___closed__23; -x_79 = l_Lean_IR_compile___closed__24; -lean_inc(x_77); -x_80 = l___private_Lean_Compiler_IR_CompilerM_0__Lean_IR_logDeclsAux(x_79, x_78, x_77, x_51, x_52); -if (lean_obj_tag(x_80) == 0) -{ -lean_object* x_81; lean_object* x_82; uint8_t x_83; -lean_dec_ref(x_80); -x_81 = lean_ctor_get(x_51, 2); -x_82 = l_Lean_IR_compile___closed__25; -x_83 = l_Lean_Option_get___at___00Lean_IR_compile_spec__4(x_81, x_82); -if (x_83 == 0) -{ -x_13 = x_77; -x_14 = x_51; -x_15 = x_52; -x_16 = lean_box(0); -goto block_49; +uint8_t x_72; +lean_dec_ref(x_51); +lean_dec(x_48); +lean_dec_ref(x_47); +x_72 = !lean_is_exclusive(x_53); +if (x_72 == 0) +{ +return x_53; } else { -size_t x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_84 = lean_array_size(x_77); -x_85 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__5(x_84, x_10, x_77); -x_86 = l_Lean_IR_compile___closed__27; -x_87 = l_Lean_IR_compile___closed__28; -lean_inc_ref(x_85); -x_88 = l___private_Lean_Compiler_IR_CompilerM_0__Lean_IR_logDeclsAux(x_87, x_86, x_85, x_51, x_52); -if (lean_obj_tag(x_88) == 0) +lean_object* x_73; lean_object* x_74; +x_73 = lean_ctor_get(x_53, 0); +lean_inc(x_73); +lean_dec(x_53); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_73); +return x_74; +} +} +} +block_134: +{ +size_t x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_80 = lean_array_size(x_76); +x_81 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__2(x_80, x_29, x_76); +x_82 = l_Lean_IR_compile___closed__14; +x_83 = l_Lean_IR_compile___closed__15; +lean_inc_ref(x_81); +x_84 = l___private_Lean_Compiler_IR_CompilerM_0__Lean_IR_logDeclsAux(x_83, x_82, x_81, x_77, x_78); +if (lean_obj_tag(x_84) == 0) +{ +size_t x_85; lean_object* x_86; size_t x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +lean_dec_ref(x_84); +x_85 = lean_array_size(x_81); +x_86 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__3(x_85, x_29, x_81); +x_87 = lean_array_size(x_86); +x_88 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__4(x_87, x_29, x_86); +x_89 = l_Lean_IR_compile___closed__17; +x_90 = l_Lean_IR_compile___closed__18; +lean_inc_ref(x_88); +x_91 = l___private_Lean_Compiler_IR_CompilerM_0__Lean_IR_logDeclsAux(x_90, x_89, x_88, x_77, x_78); +if (lean_obj_tag(x_91) == 0) +{ +lean_object* x_92; +lean_dec_ref(x_91); +x_92 = l_Lean_IR_inferBorrow___redArg(x_88, x_78); +if (lean_obj_tag(x_92) == 0) +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_93 = lean_ctor_get(x_92, 0); +lean_inc(x_93); +lean_dec_ref(x_92); +x_94 = l_Lean_IR_compile___closed__20; +x_95 = l_Lean_IR_compile___closed__21; +lean_inc(x_93); +x_96 = l___private_Lean_Compiler_IR_CompilerM_0__Lean_IR_logDeclsAux(x_95, x_94, x_93, x_77, x_78); +if (lean_obj_tag(x_96) == 0) +{ +lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; +lean_dec_ref(x_96); +x_97 = lean_st_ref_get(x_78); +x_98 = lean_ctor_get(x_97, 0); +lean_inc_ref(x_98); +lean_dec(x_97); +x_99 = l_Lean_IR_ExplicitBoxing_addBoxedVersions(x_98, x_93); +x_100 = l_Lean_IR_compile___closed__23; +x_101 = l_Lean_IR_compile___closed__24; +lean_inc_ref(x_99); +x_102 = l___private_Lean_Compiler_IR_CompilerM_0__Lean_IR_logDeclsAux(x_101, x_100, x_99, x_77, x_78); +if (lean_obj_tag(x_102) == 0) +{ +lean_object* x_103; +lean_dec_ref(x_102); +x_103 = l_Lean_IR_explicitRC___redArg(x_99, x_78); +if (lean_obj_tag(x_103) == 0) +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_104 = lean_ctor_get(x_103, 0); +lean_inc(x_104); +lean_dec_ref(x_103); +x_105 = l_Lean_IR_compile___closed__26; +x_106 = l_Lean_IR_compile___closed__27; +lean_inc(x_104); +x_107 = l___private_Lean_Compiler_IR_CompilerM_0__Lean_IR_logDeclsAux(x_106, x_105, x_104, x_77, x_78); +if (lean_obj_tag(x_107) == 0) +{ +lean_object* x_108; lean_object* x_109; uint8_t x_110; +lean_dec_ref(x_107); +x_108 = lean_ctor_get(x_77, 2); +x_109 = l_Lean_IR_compile___closed__28; +x_110 = l_Lean_Option_get___at___00Lean_IR_compile_spec__5(x_108, x_109); +if (x_110 == 0) +{ +x_46 = x_104; +x_47 = x_77; +x_48 = x_78; +x_49 = lean_box(0); +goto block_75; +} +else { -lean_dec_ref(x_88); -x_13 = x_85; -x_14 = x_51; -x_15 = x_52; -x_16 = lean_box(0); -goto block_49; +size_t x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; +x_111 = lean_array_size(x_104); +x_112 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__6(x_111, x_29, x_104); +x_113 = l_Lean_IR_compile___closed__30; +x_114 = l_Lean_IR_compile___closed__31; +lean_inc_ref(x_112); +x_115 = l___private_Lean_Compiler_IR_CompilerM_0__Lean_IR_logDeclsAux(x_114, x_113, x_112, x_77, x_78); +if (lean_obj_tag(x_115) == 0) +{ +lean_dec_ref(x_115); +x_46 = x_112; +x_47 = x_77; +x_48 = x_78; +x_49 = lean_box(0); +goto block_75; } else { -uint8_t x_89; -lean_dec_ref(x_85); -lean_dec(x_52); -lean_dec_ref(x_51); -x_89 = !lean_is_exclusive(x_88); -if (x_89 == 0) +uint8_t x_116; +lean_dec_ref(x_112); +lean_dec(x_78); +lean_dec_ref(x_77); +x_116 = !lean_is_exclusive(x_115); +if (x_116 == 0) { -return x_88; +return x_115; } else { -lean_object* x_90; lean_object* x_91; -x_90 = lean_ctor_get(x_88, 0); -lean_inc(x_90); -lean_dec(x_88); -x_91 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_91, 0, x_90); -return x_91; +lean_object* x_117; lean_object* x_118; +x_117 = lean_ctor_get(x_115, 0); +lean_inc(x_117); +lean_dec(x_115); +x_118 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_118, 0, x_117); +return x_118; } } } } else { -uint8_t x_92; -lean_dec(x_77); -lean_dec(x_52); -lean_dec_ref(x_51); -x_92 = !lean_is_exclusive(x_80); -if (x_92 == 0) +uint8_t x_119; +lean_dec(x_104); +lean_dec(x_78); +lean_dec_ref(x_77); +x_119 = !lean_is_exclusive(x_107); +if (x_119 == 0) { -return x_80; +return x_107; } else { -lean_object* x_93; lean_object* x_94; -x_93 = lean_ctor_get(x_80, 0); -lean_inc(x_93); -lean_dec(x_80); -x_94 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_94, 0, x_93); -return x_94; +lean_object* x_120; lean_object* x_121; +x_120 = lean_ctor_get(x_107, 0); +lean_inc(x_120); +lean_dec(x_107); +x_121 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_121, 0, x_120); +return x_121; } } } else { -lean_dec(x_52); -lean_dec_ref(x_51); -return x_76; +lean_dec(x_78); +lean_dec_ref(x_77); +return x_103; } } else { -uint8_t x_95; -lean_dec(x_72); -lean_dec(x_52); -lean_dec_ref(x_51); -x_95 = !lean_is_exclusive(x_75); -if (x_95 == 0) +uint8_t x_122; +lean_dec_ref(x_99); +lean_dec(x_78); +lean_dec_ref(x_77); +x_122 = !lean_is_exclusive(x_102); +if (x_122 == 0) { -return x_75; +return x_102; } else { -lean_object* x_96; lean_object* x_97; -x_96 = lean_ctor_get(x_75, 0); -lean_inc(x_96); -lean_dec(x_75); -x_97 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_97, 0, x_96); -return x_97; -} +lean_object* x_123; lean_object* x_124; +x_123 = lean_ctor_get(x_102, 0); +lean_inc(x_123); +lean_dec(x_102); +x_124 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_124, 0, x_123); +return x_124; } } -else -{ -lean_dec(x_52); -lean_dec_ref(x_51); -return x_71; -} } else { -uint8_t x_98; -lean_dec(x_67); -lean_dec(x_52); -lean_dec_ref(x_51); -x_98 = !lean_is_exclusive(x_70); -if (x_98 == 0) +uint8_t x_125; +lean_dec(x_93); +lean_dec(x_78); +lean_dec_ref(x_77); +x_125 = !lean_is_exclusive(x_96); +if (x_125 == 0) { -return x_70; +return x_96; } else { -lean_object* x_99; lean_object* x_100; -x_99 = lean_ctor_get(x_70, 0); -lean_inc(x_99); -lean_dec(x_70); -x_100 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_100, 0, x_99); -return x_100; +lean_object* x_126; lean_object* x_127; +x_126 = lean_ctor_get(x_96, 0); +lean_inc(x_126); +lean_dec(x_96); +x_127 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_127, 0, x_126); +return x_127; } } } else { -lean_dec(x_52); -lean_dec_ref(x_51); -return x_66; +lean_dec(x_78); +lean_dec_ref(x_77); +return x_92; } } else { -uint8_t x_101; -lean_dec_ref(x_60); -lean_dec(x_52); -lean_dec_ref(x_51); -x_101 = !lean_is_exclusive(x_63); -if (x_101 == 0) +uint8_t x_128; +lean_dec_ref(x_88); +lean_dec(x_78); +lean_dec_ref(x_77); +x_128 = !lean_is_exclusive(x_91); +if (x_128 == 0) { -return x_63; +return x_91; } else { -lean_object* x_102; lean_object* x_103; -x_102 = lean_ctor_get(x_63, 0); -lean_inc(x_102); -lean_dec(x_63); -x_103 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_103, 0, x_102); -return x_103; +lean_object* x_129; lean_object* x_130; +x_129 = lean_ctor_get(x_91, 0); +lean_inc(x_129); +lean_dec(x_91); +x_130 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_130, 0, x_129); +return x_130; } } } else { -uint8_t x_104; -lean_dec_ref(x_55); -lean_dec(x_52); -lean_dec_ref(x_51); -x_104 = !lean_is_exclusive(x_58); -if (x_104 == 0) +uint8_t x_131; +lean_dec_ref(x_81); +lean_dec(x_78); +lean_dec_ref(x_77); +x_131 = !lean_is_exclusive(x_84); +if (x_131 == 0) { -return x_58; +return x_84; } else { -lean_object* x_105; lean_object* x_106; -x_105 = lean_ctor_get(x_58, 0); -lean_inc(x_105); -lean_dec(x_58); -x_106 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_106, 0, x_105); -return x_106; +lean_object* x_132; lean_object* x_133; +x_132 = lean_ctor_get(x_84, 0); +lean_inc(x_132); +lean_dec(x_84); +x_133 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_133, 0, x_132); +return x_133; } } } } else { -uint8_t x_126; +uint8_t x_170; lean_dec(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_126 = !lean_is_exclusive(x_8); -if (x_126 == 0) +x_170 = !lean_is_exclusive(x_27); +if (x_170 == 0) { -return x_8; +return x_27; } else { -lean_object* x_127; lean_object* x_128; -x_127 = lean_ctor_get(x_8, 0); -lean_inc(x_127); -lean_dec(x_8); -x_128 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_128, 0, x_127); -return x_128; +lean_object* x_171; lean_object* x_172; +x_171 = lean_ctor_get(x_27, 0); +lean_inc(x_171); +lean_dec(x_27); +x_172 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_172, 0, x_171); +return x_172; } } } else { -uint8_t x_129; +uint8_t x_173; lean_dec(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_129 = !lean_is_exclusive(x_7); -if (x_129 == 0) +x_173 = !lean_is_exclusive(x_26); +if (x_173 == 0) { -return x_7; +return x_26; +} +else +{ +lean_object* x_174; lean_object* x_175; +x_174 = lean_ctor_get(x_26, 0); +lean_inc(x_174); +lean_dec(x_26); +x_175 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_175, 0, x_174); +return x_175; +} +} +block_22: +{ +lean_object* x_9; +lean_inc(x_7); +lean_inc_ref(x_6); +x_9 = l_Lean_IR_toposortDecls(x_5, x_6, x_7); +lean_dec_ref(x_5); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +lean_dec_ref(x_9); +x_11 = l_Lean_IR_addDecls(x_10, x_6, x_7); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; +lean_dec_ref(x_11); +x_12 = l_Lean_IR_inferMeta(x_10, x_6, x_7); +lean_dec(x_7); +lean_dec_ref(x_6); +if (lean_obj_tag(x_12) == 0) +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; +x_14 = lean_ctor_get(x_12, 0); +lean_dec(x_14); +lean_ctor_set(x_12, 0, x_10); +return x_12; +} +else +{ +lean_object* x_15; +lean_dec(x_12); +x_15 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_15, 0, x_10); +return x_15; +} +} +else +{ +uint8_t x_16; +lean_dec(x_10); +x_16 = !lean_is_exclusive(x_12); +if (x_16 == 0) +{ +return x_12; +} +else +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_12, 0); +lean_inc(x_17); +lean_dec(x_12); +x_18 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_18, 0, x_17); +return x_18; +} +} +} +else +{ +uint8_t x_19; +lean_dec(x_10); +lean_dec(x_7); +lean_dec_ref(x_6); +x_19 = !lean_is_exclusive(x_11); +if (x_19 == 0) +{ +return x_11; +} +else +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_11, 0); +lean_inc(x_20); +lean_dec(x_11); +x_21 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_21, 0, x_20); +return x_21; +} +} } else { -lean_object* x_130; lean_object* x_131; -x_130 = lean_ctor_get(x_7, 0); -lean_inc(x_130); lean_dec(x_7); -x_131 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_131, 0, x_130); -return x_131; +lean_dec_ref(x_6); +return x_9; } } } } -LEAN_EXPORT lean_object* l_Lean_Option_get___at___00Lean_IR_compile_spec__4___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Option_get___at___00Lean_IR_compile_spec__5___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l_Lean_Option_get___at___00Lean_IR_compile_spec__4(x_1, x_2); +x_3 = l_Lean_Option_get___at___00Lean_IR_compile_spec__5(x_1, x_2); lean_dec_ref(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -1517,7 +1797,19 @@ x_6 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR return x_6; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +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___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__4(x_4, x_5, x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -1525,11 +1817,11 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__5(x_4, x_5, x_3); +x_6 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__6(x_4, x_5, x_3); return x_6; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { size_t x_5; size_t x_6; lean_object* x_7; @@ -1537,7 +1829,7 @@ x_5 = lean_unbox_usize(x_2); lean_dec(x_2); x_6 = lean_unbox_usize(x_3); lean_dec(x_3); -x_7 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__6(x_1, x_5, x_6, x_4); +x_7 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_compile_spec__7(x_1, x_5, x_6, x_4); return x_7; } } @@ -2021,6 +2313,24 @@ l_Lean_IR_compile___closed__30 = _init_l_Lean_IR_compile___closed__30(); lean_mark_persistent(l_Lean_IR_compile___closed__30); l_Lean_IR_compile___closed__31 = _init_l_Lean_IR_compile___closed__31(); lean_mark_persistent(l_Lean_IR_compile___closed__31); +l_Lean_IR_compile___closed__32 = _init_l_Lean_IR_compile___closed__32(); +lean_mark_persistent(l_Lean_IR_compile___closed__32); +l_Lean_IR_compile___closed__33 = _init_l_Lean_IR_compile___closed__33(); +lean_mark_persistent(l_Lean_IR_compile___closed__33); +l_Lean_IR_compile___closed__34 = _init_l_Lean_IR_compile___closed__34(); +lean_mark_persistent(l_Lean_IR_compile___closed__34); +l_Lean_IR_compile___closed__35 = _init_l_Lean_IR_compile___closed__35(); +lean_mark_persistent(l_Lean_IR_compile___closed__35); +l_Lean_IR_compile___closed__36 = _init_l_Lean_IR_compile___closed__36(); +lean_mark_persistent(l_Lean_IR_compile___closed__36); +l_Lean_IR_compile___closed__37 = _init_l_Lean_IR_compile___closed__37(); +lean_mark_persistent(l_Lean_IR_compile___closed__37); +l_Lean_IR_compile___closed__38 = _init_l_Lean_IR_compile___closed__38(); +lean_mark_persistent(l_Lean_IR_compile___closed__38); +l_Lean_IR_compile___closed__39 = _init_l_Lean_IR_compile___closed__39(); +lean_mark_persistent(l_Lean_IR_compile___closed__39); +l_Lean_IR_compile___closed__40 = _init_l_Lean_IR_compile___closed__40(); +lean_mark_persistent(l_Lean_IR_compile___closed__40); l___private_Lean_Compiler_IR_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_640659120____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_640659120____hygCtx___hyg_2_(); lean_mark_persistent(l___private_Lean_Compiler_IR_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_640659120____hygCtx___hyg_2_); l___private_Lean_Compiler_IR_0__Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_640659120____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_0__Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_640659120____hygCtx___hyg_2_(); diff --git a/stage0/stdlib/Lean/Compiler/IR/Basic.c b/stage0/stdlib/Lean/Compiler/IR/Basic.c index f3c50f9aa6c1..9e7a93ace9d7 100644 --- a/stage0/stdlib/Lean/Compiler/IR/Basic.c +++ b/stage0/stdlib/Lean/Compiler/IR/Basic.c @@ -14,6 +14,7 @@ extern "C" { #endif LEAN_EXPORT lean_object* l_Lean_IR_instInhabitedArg_default; +LEAN_EXPORT uint8_t l_Lean_IR_IRType_isScalarOrStruct(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instBEqCtorInfo_beq___boxed(lean_object*, lean_object*); static lean_object* l_Array_Array_repr___at___00Lean_IR_instReprIRType_repr_spec__1___closed__5; LEAN_EXPORT lean_object* l_Lean_IR_LitVal_num_elim(lean_object*, lean_object*, lean_object*, lean_object*); @@ -24,6 +25,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_Expr_ctorElim(lean_object*, lean_object*, lea static lean_object* l_Lean_IR_instReprVarId___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_FnBody_uset_elim___redArg(lean_object*, lean_object*); static lean_object* l_Lean_IR_instReprCtorInfo_repr___redArg___closed__6; +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_IRType_normalizeObject_match__1_splitter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_instReprVarId_repr___redArg___closed__2; static lean_object* l_Array_Array_repr___at___00Lean_IR_instReprIRType_repr_spec__1___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_instInhabitedAlt; @@ -35,6 +37,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_FnBody_setTag_elim(lean_object*, lean_object* LEAN_EXPORT lean_object* l_Lean_IR_Alt_default_elim___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_IRType_uint32_elim(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Expr_reset_elim___redArg(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Array_isEqvAux___at___00Lean_IR_IRType_compatibleWith_spec__0(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_FnBody_jmp_elim___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_FnBody_ctorElim(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Expr_isShared_elim___redArg(lean_object*, lean_object*); @@ -72,8 +75,10 @@ lean_object* l_Std_Format_fill(lean_object*); LEAN_EXPORT uint8_t l_Lean_IR_instBEqVarId_beq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_push(lean_object*, lean_object*); static lean_object* l_Lean_IR_instReprIRType_repr___closed__12; +LEAN_EXPORT lean_object* l_Lean_IR_IRType_compatibleWith___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_instToStringVarId___lam__0___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_IRType_isScalar___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Array_isEqvAux___at___00Lean_IR_IRType_compatibleWith_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_instReprArg_repr___closed__4; LEAN_EXPORT lean_object* l_Lean_IR_LocalContext_eraseJoinPointDecl(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_CtorInfo_type(lean_object*); @@ -118,6 +123,7 @@ LEAN_EXPORT uint8_t l_Lean_IR_FnBody_isTerminal(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instAlphaEqvArrayArg; LEAN_EXPORT lean_object* l_Lean_IR_instReprVarId; LEAN_EXPORT lean_object* l_Lean_IR_CtorInfo_isRef___boxed(lean_object*); +LEAN_EXPORT uint8_t l_Lean_IR_IRType_isObjOrStruct(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instInhabitedJoinPointId_default; static lean_object* l_Lean_IR_instReprIRType_repr___closed__19; lean_object* lean_array_fget_borrowed(lean_object*, lean_object*); @@ -126,6 +132,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_IRType_uint64_elim(lean_object*, lean_object* LEAN_EXPORT lean_object* l_Lean_IR_IRType_usize_elim(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instBEqJoinPointId_beq___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instReprIRType_repr___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Basic_0__Array_map__unattach_match__1_splitter___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_IRType_uint8_elim___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_addParamsRename_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Arg_var_elim___redArg(lean_object*, lean_object*); @@ -136,9 +143,11 @@ static lean_object* l_Lean_IR_instReprArg_repr___closed__0; static lean_object* l_Lean_IR_instInhabitedExpr_default___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_IRType_isVoid___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_FnBody_jdecl_elim___redArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_IRType_match__on__same__ctor_het___redArg_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_102_(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instInhabitedCtorInfo_default; LEAN_EXPORT lean_object* l_Lean_IR_reshapeAux(lean_object*, lean_object*, lean_object*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_IRType_isScalarOrStruct___boxed(lean_object*); LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_insert___at___00Lean_IR_mkIndexSet_spec__1___redArg(lean_object*, lean_object*, lean_object*); lean_object* lean_string_length(lean_object*); LEAN_EXPORT lean_object* l_Array_isEqvAux___at___00Lean_IR_args_alphaEqv_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -155,6 +164,7 @@ static lean_object* l_Lean_IR_instInhabitedDecl_default___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_IRType_float_elim(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Decl_extern_elim___redArg(lean_object*, lean_object*); static lean_object* l_Array_Array_repr___at___00Lean_IR_instReprIRType_repr_spec__1___closed__10; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_IRType_isDefiniteRef_spec__0___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_IR_LocalContext_isLocalVar(lean_object*, lean_object*); static lean_object* l_Option_repr___at___00Lean_IR_instReprIRType_repr_spec__0___closed__2; LEAN_EXPORT lean_object* l_Lean_IR_Alt_ctor_elim___redArg(lean_object*, lean_object*); @@ -178,6 +188,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_Alt_default_elim(lean_object*, lean_object*, LEAN_EXPORT lean_object* l_Lean_IR_IRType_object_elim(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_modifyJPsM___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_instInhabitedFnBody_default__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_IR_IRType_match__on__same__ctor_het_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_102_(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instReprParam_repr(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instBEqLitVal; size_t lean_usize_of_nat(lean_object*); @@ -200,6 +211,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_FnBody_vdecl_elim___redArg(lean_object*, lean LEAN_EXPORT lean_object* l_Lean_IR_Expr_fap_elim___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_FnBody_setTag_elim___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_IRType_isDefiniteRef___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_IRType_normalizeObject_match__1_splitter___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_IR_Index_lt(lean_object*, lean_object*); static lean_object* l_Lean_IR_instReprParam_repr___redArg___closed__3; static lean_object* l_Lean_IR_instBEqJoinPointId___closed__0; @@ -214,6 +226,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_IRType_union_elim(lean_object*, lean_object*, LEAN_EXPORT lean_object* l_Lean_IR_LocalContext_getType(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_FnBody_set_elim(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instInhabitedParam_default; +LEAN_EXPORT uint8_t l_Lean_IR_IRType_compatibleWith(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_IR_IRType_void_elim(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Decl_params(lean_object*); lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); @@ -240,8 +253,9 @@ LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_addParam LEAN_EXPORT lean_object* l_Lean_IR_instReprIRType_repr(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instBEqJoinPointId; LEAN_EXPORT lean_object* l_Lean_IR_Expr_alphaEqv___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_isEqvAux___at___00Lean_IR_IRType_compatibleWith_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_instInhabitedCtorInfo_default___closed__0; -LEAN_EXPORT uint8_t l_Lean_IR_instBEqIRType_beq___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_IR_instBEqIRType_beq___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_instReprCtorInfo_repr___redArg___closed__5; static lean_object* l_Lean_IR_instReprIRType_repr___closed__20; LEAN_EXPORT lean_object* l_Lean_IR_Decl_ctorIdx___boxed(lean_object*); @@ -261,6 +275,7 @@ lean_object* l_Id_instMonad___lam__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_IRType_union_elim___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_args_alphaEqv___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_FnBody_body___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_IRType_isObjOrStruct___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Decl_name(lean_object*); static lean_object* l_Lean_IR_instReprIRType_repr___closed__28; static lean_object* l_Lean_IR_getUnboxOpName___closed__1; @@ -277,7 +292,6 @@ LEAN_EXPORT lean_object* l_Lean_IR_IRType_struct_elim(lean_object*, lean_object* static lean_object* l_Lean_IR_instReprParam_repr___redArg___closed__5; LEAN_EXPORT lean_object* l_Lean_IR_Decl_name___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_FnBody_body(lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_IRType_match__on__same__ctor_het___redArg_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_84_(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_mkIf___closed__3; lean_object* lean_array_to_list(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instReprCtorInfo_repr(lean_object*, lean_object*); @@ -303,7 +317,9 @@ LEAN_EXPORT lean_object* l_Lean_IR_FnBody_unreachable_elim___redArg(lean_object* LEAN_EXPORT lean_object* l_Lean_IR_IRType_float_elim___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_IRType_uint8_elim(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Arg_var_elim(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_IRType_match__on__same__ctor_het_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_102____boxed(lean_object**); static lean_object* l_Lean_IR_instReprIRType_repr___closed__16; +LEAN_EXPORT uint8_t l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_IRType_isDefiniteRef_spec__0(lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_IR_addParamsRename___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_instInhabitedParam_default___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_instBEqFnBody; @@ -315,7 +331,6 @@ static lean_object* l_Lean_IR_Decl_updateBody_x21___closed__2; static lean_object* l_Lean_IR_instReprIRType_repr___closed__18; LEAN_EXPORT lean_object* l_Lean_IR_FnBody_case_elim___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_LitVal_ctorIdx(lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_IRType_match__on__same__ctor_het_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_84____boxed(lean_object**); LEAN_EXPORT uint8_t l_Option_instBEq_beq___at___00Lean_IR_instBEqIRType_beq_spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instReprJoinPointId; static lean_object* l_Lean_IR_instInhabitedDecl_default___closed__1; @@ -344,6 +359,7 @@ LEAN_EXPORT lean_object* l_Array_Array_repr___at___00Lean_IR_instReprIRType_repr static lean_object* l_Lean_IR_instReprIRType_repr___closed__1; LEAN_EXPORT uint8_t l_Lean_IR_FnBody_alphaEqv(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_modifyJPs___closed__7; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_IRType_normalizeObject_spec__0(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Expr_sproj_elim___redArg(lean_object*, lean_object*); static lean_object* l_Lean_IR_instReprIRType_repr___closed__25; LEAN_EXPORT uint8_t l_Lean_IR_Decl_isExtern(lean_object*); @@ -391,6 +407,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_IRType_uint32_elim___redArg(lean_object*, lea LEAN_EXPORT lean_object* l_Lean_IR_Alt_isDefault___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Expr_box_elim(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_modifyJPs(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_isEqvAux___at___00Lean_IR_IRType_compatibleWith_spec__1___redArg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Expr_ap_elim___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_mkIf(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_reshapeAux_spec__0(lean_object*, lean_object*); @@ -413,6 +430,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_FnBody_sset_elim___redArg(lean_object*, lean_ static lean_object* l_Lean_IR_instReprParam_repr___redArg___closed__7; LEAN_EXPORT lean_object* l_Lean_IR_FnBody_setBody(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_modifyJPsM___redArg___lam__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Array_isEqvAux___at___00Lean_IR_IRType_compatibleWith_spec__0___redArg(uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instReprParam; LEAN_EXPORT lean_object* l_Lean_IR_instBEqArg; LEAN_EXPORT lean_object* l_Lean_IR_Decl_updateBody_x21(lean_object*, lean_object*); @@ -424,7 +442,6 @@ LEAN_EXPORT lean_object* l_Lean_IR_Decl_params___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_FnBody_case_elim(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_LocalContext_isParam___boxed(lean_object*, lean_object*); static lean_object* l_Lean_IR_getUnboxOpName___closed__2; -LEAN_EXPORT lean_object* l_Lean_IR_IRType_match__on__same__ctor_het_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_84_(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_instReprVarId_repr___redArg___closed__9; static lean_object* l_Lean_IR_instReprCtorInfo_repr___redArg___closed__4; LEAN_EXPORT lean_object* l_Lean_IR_Expr_reuse_elim___redArg(lean_object*, lean_object*); @@ -435,6 +452,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_Expr_box_elim___redArg(lean_object*, lean_obj LEAN_EXPORT lean_object* l_Lean_IR_modifyJPsM(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Option_repr___at___00Lean_IR_instReprIRType_repr_spec__0___closed__3; +LEAN_EXPORT lean_object* l_Lean_IR_IRType_isStruct___boxed(lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); static lean_object* l_Lean_IR_modifyJPs___closed__9; LEAN_EXPORT lean_object* l_Lean_IR_IRType_uint16_elim___redArg(lean_object*, lean_object*); @@ -448,6 +466,7 @@ static lean_object* l_Lean_IR_instBEqCtorInfo___closed__0; 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*); LEAN_EXPORT lean_object* l_Lean_IR_Decl_ctorElim___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Basic_0__Array_map__unattach_match__1_splitter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_modifyJPs___closed__3; LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___00Array_Array_repr___at___00Lean_IR_instReprIRType_repr_spec__1_spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Arg_ctorElim(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -465,7 +484,9 @@ LEAN_EXPORT lean_object* l_Lean_IR_FnBody_inc_elim___redArg(lean_object*, lean_o lean_object* lean_panic_fn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Decl_ctorElim(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_IRType_usize_elim___redArg(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Array_isEqvAux___at___00Lean_IR_IRType_compatibleWith_spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_IRType_tagged_elim(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_IRType_match__on__same__ctor_het___redArg_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_102____boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_IRType_ctorElim(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Alt_ctor_elim(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); @@ -498,6 +519,7 @@ static lean_object* l_Lean_IR_modifyJPs___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_instReprVarId_repr(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Expr_uproj_elim(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_LocalContext_eraseJoinPointDecl___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_IR_IRType_isStruct(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_IRType_isObj___boxed(lean_object*); LEAN_EXPORT uint8_t l_Lean_IR_LocalContext_contains(lean_object*, lean_object*); LEAN_EXPORT uint64_t l_Lean_IR_instHashableJoinPointId_hash(lean_object*); @@ -532,13 +554,13 @@ LEAN_EXPORT lean_object* l_Lean_IR_instReprParam_repr___boxed(lean_object*, lean static lean_object* l_Lean_IR_instReprParam_repr___redArg___closed__6; LEAN_EXPORT lean_object* l_Lean_IR_LocalContext_addJP(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_Const_get_x3f___at___00Lean_IR_LocalContext_isJP_spec__0___redArg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_IRType_normalizeObject(lean_object*); static lean_object* l_Lean_IR_instReprCtorInfo_repr___redArg___closed__3; static lean_object* l_Lean_IR_modifyJPs___closed__6; lean_object* lean_string_append(lean_object*, lean_object*); static lean_object* l_Lean_IR_instInhabitedLitVal_default___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_instReprCtorInfo_repr___redArg(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Expr_ctorElim___redArg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_IRType_match__on__same__ctor_het___redArg_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_84____boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_instReprVarId_repr___redArg___closed__3; LEAN_EXPORT lean_object* l_Lean_IR_FnBody_del_elim(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_VarId_alphaEqv___boxed(lean_object*, lean_object*, lean_object*); @@ -555,6 +577,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_instToStringJoinPointId___lam__0(lean_object* LEAN_EXPORT lean_object* l_Lean_IR_reshape(lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instInhabitedVarId_default; +uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT uint64_t l_Lean_IR_instHashableVarId_hash(lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_LitVal_ctorIdx___boxed(lean_object*); @@ -562,13 +585,16 @@ LEAN_EXPORT lean_object* l_Lean_IR_Alt_modifyBodyM___redArg___lam__1(lean_object static lean_object* l_Lean_IR_instReprIRType_repr___closed__3; static lean_object* l_Lean_IR_getUnboxOpName___closed__3; static lean_object* l_Lean_IR_instReprVarId_repr___redArg___closed__1; +LEAN_EXPORT uint8_t l_Array_isEqvAux___at___00Lean_IR_IRType_compatibleWith_spec__1___redArg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instReprParam_repr___redArg(lean_object*); LEAN_EXPORT uint8_t l_Array_isEqvAux___at___00Lean_IR_instBEqIRType_beq_spec__0___redArg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Arg_ctorIdx___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_IRType_normalizeObject_spec__0___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Expr_proj_elim___redArg(lean_object*, lean_object*); static lean_object* l_Lean_IR_modifyJPs___closed__1; static lean_object* l_Lean_IR_Decl_updateBody_x21___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_Decl_fdecl_elim(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT uint8_t l_Std_DTreeMap_Internal_Impl_contains___at___00Lean_IR_mkIndexSet_spec__0(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instReprCtorInfo; LEAN_EXPORT lean_object* l_Lean_IR_IRType_erased_elim___redArg(lean_object*, lean_object*); @@ -586,6 +612,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_FnBody_ret_elim(lean_object*, lean_object*, l LEAN_EXPORT lean_object* l_Lean_IR_Expr_lit_elim___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Expr_unbox_elim___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_isEqvAux___at___00Lean_IR_FnBody_alphaEqv_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_isEqvAux___at___00Lean_IR_IRType_compatibleWith_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_instReprIRType_repr___closed__23; LEAN_EXPORT lean_object* l_Lean_IR_FnBody_inc_elim(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_FnBody_dec_elim(lean_object*, lean_object*, lean_object*, lean_object*); @@ -595,7 +622,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_IRType_float32_elim(lean_object*, lean_object static lean_object* l_Lean_IR_instReprIRType_repr___closed__0; static lean_object* l_Array_Array_repr___at___00Lean_IR_instReprIRType_repr_spec__1___closed__6; LEAN_EXPORT lean_object* l_Lean_IR_IRType_erased_elim(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_instBEqIRType_beq___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_instBEqIRType_beq___lam__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_Option_repr___at___00Lean_IR_instReprIRType_repr_spec__0___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_LocalContext_getValue___boxed(lean_object*, lean_object*); static lean_object* l_Lean_IR_instReprArg_repr___closed__2; @@ -1223,25 +1250,29 @@ LEAN_EXPORT lean_object* l_Lean_IR_IRType_ctorElim___redArg(lean_object* x_1, le switch (lean_obj_tag(x_1)) { case 10: { -lean_object* x_3; lean_object* x_4; lean_object* x_5; +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; x_3 = lean_ctor_get(x_1, 0); lean_inc(x_3); x_4 = lean_ctor_get(x_1, 1); lean_inc_ref(x_4); +x_5 = lean_ctor_get(x_1, 2); +lean_inc(x_5); +x_6 = lean_ctor_get(x_1, 3); +lean_inc(x_6); lean_dec_ref(x_1); -x_5 = lean_apply_2(x_2, x_3, x_4); -return x_5; +x_7 = lean_apply_4(x_2, x_3, x_4, x_5, x_6); +return x_7; } case 11: { -lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_6 = lean_ctor_get(x_1, 0); -lean_inc(x_6); -x_7 = lean_ctor_get(x_1, 1); -lean_inc_ref(x_7); +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_9); lean_dec_ref(x_1); -x_8 = lean_apply_2(x_2, x_6, x_7); -return x_8; +x_10 = lean_apply_2(x_2, x_8, x_9); +return x_10; } default: { @@ -1508,7 +1539,7 @@ x_1 = lean_box(0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_IRType_match__on__same__ctor_het___redArg_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_84_(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +LEAN_EXPORT lean_object* l_Lean_IR_IRType_match__on__same__ctor_het___redArg_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_102_(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { switch (lean_obj_tag(x_1)) { @@ -1594,37 +1625,45 @@ return x_12; } case 10: { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_dec(x_14); x_17 = lean_ctor_get(x_1, 0); lean_inc(x_17); x_18 = lean_ctor_get(x_1, 1); lean_inc_ref(x_18); -lean_dec_ref(x_1); -x_19 = lean_ctor_get(x_2, 0); +x_19 = lean_ctor_get(x_1, 2); lean_inc(x_19); -x_20 = lean_ctor_get(x_2, 1); -lean_inc_ref(x_20); +x_20 = lean_ctor_get(x_1, 3); +lean_inc(x_20); +lean_dec_ref(x_1); +x_21 = lean_ctor_get(x_2, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_2, 1); +lean_inc_ref(x_22); +x_23 = lean_ctor_get(x_2, 2); +lean_inc(x_23); +x_24 = lean_ctor_get(x_2, 3); +lean_inc(x_24); lean_dec(x_2); -x_21 = lean_apply_4(x_13, x_17, x_18, x_19, x_20); -return x_21; +x_25 = lean_apply_8(x_13, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24); +return x_25; } case 11: { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_dec(x_13); -x_22 = lean_ctor_get(x_1, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_1, 1); -lean_inc_ref(x_23); +x_26 = lean_ctor_get(x_1, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_27); lean_dec_ref(x_1); -x_24 = lean_ctor_get(x_2, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_2, 1); -lean_inc_ref(x_25); +x_28 = lean_ctor_get(x_2, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_2, 1); +lean_inc_ref(x_29); lean_dec(x_2); -x_26 = lean_apply_4(x_14, x_22, x_23, x_24, x_25); -return x_26; +x_30 = lean_apply_4(x_14, x_26, x_27, x_28, x_29); +return x_30; } case 12: { @@ -1645,15 +1684,15 @@ return x_16; } } } -LEAN_EXPORT lean_object* l_Lean_IR_IRType_match__on__same__ctor_het_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_84_(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +LEAN_EXPORT lean_object* l_Lean_IR_IRType_match__on__same__ctor_het_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_102_(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { _start: { lean_object* x_19; -x_19 = l_Lean_IR_IRType_match__on__same__ctor_het___redArg_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_84_(x_2, x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +x_19 = l_Lean_IR_IRType_match__on__same__ctor_het___redArg_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_102_(x_2, x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); return x_19; } } -LEAN_EXPORT lean_object* l_Lean_IR_IRType_match__on__same__ctor_het_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_84____boxed(lean_object** _args) { +LEAN_EXPORT lean_object* l_Lean_IR_IRType_match__on__same__ctor_het_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_102____boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -1675,7 +1714,7 @@ lean_object* x_18 = _args[17]; _start: { lean_object* x_19; -x_19 = l_Lean_IR_IRType_match__on__same__ctor_het_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_84_(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +x_19 = l_Lean_IR_IRType_match__on__same__ctor_het_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_102_(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); lean_dec(x_18); lean_dec(x_17); lean_dec(x_14); @@ -1691,11 +1730,11 @@ lean_dec(x_5); return x_19; } } -LEAN_EXPORT lean_object* l_Lean_IR_IRType_match__on__same__ctor_het___redArg_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_84____boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +LEAN_EXPORT lean_object* l_Lean_IR_IRType_match__on__same__ctor_het___redArg_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_102____boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { lean_object* x_17; -x_17 = l_Lean_IR_IRType_match__on__same__ctor_het___redArg_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_84_(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +x_17 = l_Lean_IR_IRType_match__on__same__ctor_het___redArg_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_102_(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); lean_dec(x_16); lean_dec(x_15); lean_dec(x_12); @@ -1805,45 +1844,67 @@ return x_8; } } } -LEAN_EXPORT uint8_t l_Lean_IR_instBEqIRType_beq___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT uint8_t l_Lean_IR_instBEqIRType_beq___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -uint8_t x_7; -x_7 = l_Option_instBEq_beq___at___00Lean_IR_instBEqIRType_beq_spec__1(x_1, x_3); -if (x_7 == 0) +uint8_t x_11; +x_11 = l_Option_instBEq_beq___at___00Lean_IR_instBEqIRType_beq_spec__1(x_1, x_5); +if (x_11 == 0) { -return x_7; +return x_11; } else { -lean_object* x_8; lean_object* x_9; uint8_t x_10; -x_8 = lean_array_get_size(x_2); -x_9 = lean_array_get_size(x_4); -x_10 = lean_nat_dec_eq(x_8, x_9); -if (x_10 == 0) +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = lean_array_get_size(x_2); +x_13 = lean_array_get_size(x_6); +x_14 = lean_nat_dec_eq(x_12, x_13); +if (x_14 == 0) { -return x_10; +return x_14; } else { -uint8_t x_11; -x_11 = l_Array_isEqvAux___at___00Lean_IR_instBEqIRType_beq_spec__0___redArg(x_2, x_4, x_8); -return x_11; +uint8_t x_15; +x_15 = l_Array_isEqvAux___at___00Lean_IR_instBEqIRType_beq_spec__0___redArg(x_2, x_6, x_12); +if (x_15 == 0) +{ +return x_15; +} +else +{ +uint8_t x_16; +x_16 = lean_nat_dec_eq(x_3, x_7); +if (x_16 == 0) +{ +return x_16; +} +else +{ +uint8_t x_17; +x_17 = lean_nat_dec_eq(x_4, x_8); +return x_17; } } } } -LEAN_EXPORT lean_object* l_Lean_IR_instBEqIRType_beq___lam__1___boxed(lean_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_IR_instBEqIRType_beq___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -uint8_t x_7; lean_object* x_8; -x_7 = l_Lean_IR_instBEqIRType_beq___lam__1(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec_ref(x_4); +uint8_t x_11; lean_object* x_12; +x_11 = l_Lean_IR_instBEqIRType_beq___lam__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); lean_dec_ref(x_2); lean_dec(x_1); -x_8 = lean_box(x_7); -return x_8; +x_12 = lean_box(x_11); +return x_12; } } LEAN_EXPORT uint8_t l_Lean_IR_instBEqIRType_beq(lean_object* x_1, lean_object* x_2) { @@ -1868,8 +1929,8 @@ x_6 = lean_box(x_5); x_7 = lean_alloc_closure((void*)(l_Lean_IR_instBEqIRType_beq___lam__0___boxed), 3, 1); lean_closure_set(x_7, 0, x_6); x_8 = lean_alloc_closure((void*)(l_Lean_IR_instBEqIRType_beq___lam__12___boxed), 6, 0); -x_9 = lean_alloc_closure((void*)(l_Lean_IR_instBEqIRType_beq___lam__1___boxed), 6, 0); -x_10 = l_Lean_IR_IRType_match__on__same__ctor_het___redArg_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_84_(x_1, x_2, x_7, x_7, x_7, x_7, x_7, x_7, x_7, x_7, x_7, x_7, x_9, x_8, x_7, x_7); +x_9 = lean_alloc_closure((void*)(l_Lean_IR_instBEqIRType_beq___lam__1___boxed), 10, 0); +x_10 = l_Lean_IR_IRType_match__on__same__ctor_het___redArg_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_102_(x_1, x_2, x_7, x_7, x_7, x_7, x_7, x_7, x_7, x_7, x_7, x_7, x_9, x_8, x_7, x_7); lean_dec_ref(x_7); x_11 = lean_apply_2(x_10, lean_box(0), lean_box(0)); x_12 = lean_unbox(x_11); @@ -2824,169 +2885,179 @@ goto block_72; } case 10: { -lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_144; uint8_t x_145; +lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_153; uint8_t x_154; x_127 = lean_ctor_get(x_1, 0); lean_inc(x_127); x_128 = lean_ctor_get(x_1, 1); lean_inc_ref(x_128); -if (lean_is_exclusive(x_1)) { - lean_ctor_release(x_1, 0); - lean_ctor_release(x_1, 1); - x_129 = x_1; -} else { - lean_dec_ref(x_1); - x_129 = lean_box(0); -} -x_144 = lean_unsigned_to_nat(1024u); -x_145 = lean_nat_dec_le(x_144, x_2); -if (x_145 == 0) +x_129 = lean_ctor_get(x_1, 2); +lean_inc(x_129); +x_130 = lean_ctor_get(x_1, 3); +lean_inc(x_130); +lean_dec_ref(x_1); +x_153 = lean_unsigned_to_nat(1024u); +x_154 = lean_nat_dec_le(x_153, x_2); +if (x_154 == 0) { -lean_object* x_146; -x_146 = l_Lean_IR_instReprIRType_repr___closed__24; -x_130 = x_146; -goto block_143; +lean_object* x_155; +x_155 = l_Lean_IR_instReprIRType_repr___closed__24; +x_131 = x_155; +goto block_152; } else { -lean_object* x_147; -x_147 = l_Lean_IR_instReprIRType_repr___closed__25; -x_130 = x_147; -goto block_143; +lean_object* x_156; +x_156 = l_Lean_IR_instReprIRType_repr___closed__25; +x_131 = x_156; +goto block_152; } -block_143: +block_152: { -lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; uint8_t x_140; lean_object* x_141; lean_object* x_142; -x_131 = lean_box(1); -x_132 = l_Lean_IR_instReprIRType_repr___closed__28; -x_133 = lean_unsigned_to_nat(1024u); -x_134 = l_Option_repr___at___00Lean_IR_instReprIRType_repr_spec__0(x_127, x_133); -if (lean_is_scalar(x_129)) { - x_135 = lean_alloc_ctor(5, 2, 0); -} else { - x_135 = x_129; - lean_ctor_set_tag(x_135, 5); -} -lean_ctor_set(x_135, 0, x_132); -lean_ctor_set(x_135, 1, x_134); +lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; uint8_t x_149; lean_object* x_150; lean_object* x_151; +x_132 = lean_box(1); +x_133 = l_Lean_IR_instReprIRType_repr___closed__28; +x_134 = lean_unsigned_to_nat(1024u); +x_135 = l_Option_repr___at___00Lean_IR_instReprIRType_repr_spec__0(x_127, x_134); x_136 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_136, 0, x_135); -lean_ctor_set(x_136, 1, x_131); -x_137 = l_Array_Array_repr___at___00Lean_IR_instReprIRType_repr_spec__1(x_128); -x_138 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_138, 0, x_136); -lean_ctor_set(x_138, 1, x_137); -x_139 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_139, 0, x_130); +lean_ctor_set(x_136, 0, x_133); +lean_ctor_set(x_136, 1, x_135); +x_137 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_137, 0, x_136); +lean_ctor_set(x_137, 1, x_132); +x_138 = l_Array_Array_repr___at___00Lean_IR_instReprIRType_repr_spec__1(x_128); +x_139 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_139, 0, x_137); lean_ctor_set(x_139, 1, x_138); -x_140 = 0; -x_141 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_141, 0, x_139); -lean_ctor_set_uint8(x_141, sizeof(void*)*1, x_140); -x_142 = l_Repr_addAppParen(x_141, x_2); -return x_142; +x_140 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_140, 0, x_139); +lean_ctor_set(x_140, 1, x_132); +x_141 = l_Nat_reprFast(x_129); +x_142 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_142, 0, x_141); +x_143 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_143, 0, x_140); +lean_ctor_set(x_143, 1, x_142); +x_144 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_144, 0, x_143); +lean_ctor_set(x_144, 1, x_132); +x_145 = l_Nat_reprFast(x_130); +x_146 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_146, 0, x_145); +x_147 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_147, 0, x_144); +lean_ctor_set(x_147, 1, x_146); +x_148 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_148, 0, x_131); +lean_ctor_set(x_148, 1, x_147); +x_149 = 0; +x_150 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_150, 0, x_148); +lean_ctor_set_uint8(x_150, sizeof(void*)*1, x_149); +x_151 = l_Repr_addAppParen(x_150, x_2); +return x_151; } } case 11: { -lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_165; uint8_t x_166; -x_148 = lean_ctor_get(x_1, 0); -lean_inc(x_148); -x_149 = lean_ctor_get(x_1, 1); -lean_inc_ref(x_149); +lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_174; uint8_t x_175; +x_157 = lean_ctor_get(x_1, 0); +lean_inc(x_157); +x_158 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_158); if (lean_is_exclusive(x_1)) { lean_ctor_release(x_1, 0); lean_ctor_release(x_1, 1); - x_150 = x_1; + x_159 = x_1; } else { lean_dec_ref(x_1); - x_150 = lean_box(0); + x_159 = lean_box(0); } -x_165 = lean_unsigned_to_nat(1024u); -x_166 = lean_nat_dec_le(x_165, x_2); -if (x_166 == 0) +x_174 = lean_unsigned_to_nat(1024u); +x_175 = lean_nat_dec_le(x_174, x_2); +if (x_175 == 0) { -lean_object* x_167; -x_167 = l_Lean_IR_instReprIRType_repr___closed__24; -x_151 = x_167; -goto block_164; +lean_object* x_176; +x_176 = l_Lean_IR_instReprIRType_repr___closed__24; +x_160 = x_176; +goto block_173; } else { -lean_object* x_168; -x_168 = l_Lean_IR_instReprIRType_repr___closed__25; -x_151 = x_168; -goto block_164; +lean_object* x_177; +x_177 = l_Lean_IR_instReprIRType_repr___closed__25; +x_160 = x_177; +goto block_173; } -block_164: +block_173: { -lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; uint8_t x_161; lean_object* x_162; lean_object* x_163; -x_152 = lean_box(1); -x_153 = l_Lean_IR_instReprIRType_repr___closed__31; -x_154 = lean_unsigned_to_nat(1024u); -x_155 = l_Lean_Name_reprPrec(x_148, x_154); -if (lean_is_scalar(x_150)) { - x_156 = lean_alloc_ctor(5, 2, 0); +lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; uint8_t x_170; lean_object* x_171; lean_object* x_172; +x_161 = lean_box(1); +x_162 = l_Lean_IR_instReprIRType_repr___closed__31; +x_163 = lean_unsigned_to_nat(1024u); +x_164 = l_Lean_Name_reprPrec(x_157, x_163); +if (lean_is_scalar(x_159)) { + x_165 = lean_alloc_ctor(5, 2, 0); } else { - x_156 = x_150; - lean_ctor_set_tag(x_156, 5); -} -lean_ctor_set(x_156, 0, x_153); -lean_ctor_set(x_156, 1, x_155); -x_157 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_157, 0, x_156); -lean_ctor_set(x_157, 1, x_152); -x_158 = l_Array_Array_repr___at___00Lean_IR_instReprIRType_repr_spec__1(x_149); -x_159 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_159, 0, x_157); -lean_ctor_set(x_159, 1, x_158); -x_160 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_160, 0, x_151); -lean_ctor_set(x_160, 1, x_159); -x_161 = 0; -x_162 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_162, 0, x_160); -lean_ctor_set_uint8(x_162, sizeof(void*)*1, x_161); -x_163 = l_Repr_addAppParen(x_162, x_2); -return x_163; + x_165 = x_159; + lean_ctor_set_tag(x_165, 5); +} +lean_ctor_set(x_165, 0, x_162); +lean_ctor_set(x_165, 1, x_164); +x_166 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_166, 0, x_165); +lean_ctor_set(x_166, 1, x_161); +x_167 = l_Array_Array_repr___at___00Lean_IR_instReprIRType_repr_spec__1(x_158); +x_168 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_168, 0, x_166); +lean_ctor_set(x_168, 1, x_167); +x_169 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_169, 0, x_160); +lean_ctor_set(x_169, 1, x_168); +x_170 = 0; +x_171 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_171, 0, x_169); +lean_ctor_set_uint8(x_171, sizeof(void*)*1, x_170); +x_172 = l_Repr_addAppParen(x_171, x_2); +return x_172; } } case 12: { -lean_object* x_169; uint8_t x_170; -x_169 = lean_unsigned_to_nat(1024u); -x_170 = lean_nat_dec_le(x_169, x_2); -if (x_170 == 0) +lean_object* x_178; uint8_t x_179; +x_178 = lean_unsigned_to_nat(1024u); +x_179 = lean_nat_dec_le(x_178, x_2); +if (x_179 == 0) { -lean_object* x_171; -x_171 = l_Lean_IR_instReprIRType_repr___closed__24; -x_73 = x_171; +lean_object* x_180; +x_180 = l_Lean_IR_instReprIRType_repr___closed__24; +x_73 = x_180; goto block_79; } else { -lean_object* x_172; -x_172 = l_Lean_IR_instReprIRType_repr___closed__25; -x_73 = x_172; +lean_object* x_181; +x_181 = l_Lean_IR_instReprIRType_repr___closed__25; +x_73 = x_181; goto block_79; } } default: { -lean_object* x_173; uint8_t x_174; -x_173 = lean_unsigned_to_nat(1024u); -x_174 = lean_nat_dec_le(x_173, x_2); -if (x_174 == 0) +lean_object* x_182; uint8_t x_183; +x_182 = lean_unsigned_to_nat(1024u); +x_183 = lean_nat_dec_le(x_182, x_2); +if (x_183 == 0) { -lean_object* x_175; -x_175 = l_Lean_IR_instReprIRType_repr___closed__24; -x_80 = x_175; +lean_object* x_184; +x_184 = l_Lean_IR_instReprIRType_repr___closed__24; +x_80 = x_184; goto block_86; } else { -lean_object* x_176; -x_176 = l_Lean_IR_instReprIRType_repr___closed__25; -x_80 = x_176; +lean_object* x_185; +x_185 = l_Lean_IR_instReprIRType_repr___closed__25; +x_80 = x_185; goto block_86; } } @@ -3316,17 +3387,17 @@ x_3 = lean_box(x_2); return x_3; } } -LEAN_EXPORT uint8_t l_Lean_IR_IRType_isPossibleRef(lean_object* x_1) { +LEAN_EXPORT uint8_t l_Lean_IR_IRType_isStruct(lean_object* x_1) { _start: { switch (lean_obj_tag(x_1)) { -case 7: +case 10: { uint8_t x_2; x_2 = 1; return x_2; } -case 8: +case 11: { uint8_t x_3; x_3 = 1; @@ -3341,155 +3412,1394 @@ return x_4; } } } -LEAN_EXPORT lean_object* l_Lean_IR_IRType_isPossibleRef___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_IR_IRType_isStruct___boxed(lean_object* x_1) { _start: { uint8_t x_2; lean_object* x_3; -x_2 = l_Lean_IR_IRType_isPossibleRef(x_1); +x_2 = l_Lean_IR_IRType_isStruct(x_1); lean_dec(x_1); x_3 = lean_box(x_2); return x_3; } } -LEAN_EXPORT uint8_t l_Lean_IR_IRType_isDefiniteRef(lean_object* x_1) { +LEAN_EXPORT uint8_t l_Lean_IR_IRType_isObjOrStruct(lean_object* x_1) { _start: { -if (lean_obj_tag(x_1) == 7) +switch (lean_obj_tag(x_1)) { +case 7: { uint8_t x_2; x_2 = 1; return x_2; } -else +case 12: { uint8_t x_3; -x_3 = 0; +x_3 = 1; return x_3; } +case 8: +{ +uint8_t x_4; +x_4 = 1; +return x_4; } +case 13: +{ +uint8_t x_5; +x_5 = 1; +return x_5; } -LEAN_EXPORT lean_object* l_Lean_IR_IRType_isDefiniteRef___boxed(lean_object* x_1) { +case 10: +{ +uint8_t x_6; +x_6 = 1; +return x_6; +} +case 11: +{ +uint8_t x_7; +x_7 = 1; +return x_7; +} +default: +{ +uint8_t x_8; +x_8 = 0; +return x_8; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_IRType_isObjOrStruct___boxed(lean_object* x_1) { _start: { uint8_t x_2; lean_object* x_3; -x_2 = l_Lean_IR_IRType_isDefiniteRef(x_1); +x_2 = l_Lean_IR_IRType_isObjOrStruct(x_1); lean_dec(x_1); x_3 = lean_box(x_2); return x_3; } } -LEAN_EXPORT uint8_t l_Lean_IR_IRType_isErased(lean_object* x_1) { +LEAN_EXPORT uint8_t l_Lean_IR_IRType_isScalarOrStruct(lean_object* x_1) { _start: { -if (lean_obj_tag(x_1) == 6) +switch (lean_obj_tag(x_1)) { +case 0: { uint8_t x_2; x_2 = 1; return x_2; } -else +case 9: { uint8_t x_3; -x_3 = 0; +x_3 = 1; return x_3; } +case 1: +{ +uint8_t x_4; +x_4 = 1; +return x_4; +} +case 2: +{ +uint8_t x_5; +x_5 = 1; +return x_5; } +case 3: +{ +uint8_t x_6; +x_6 = 1; +return x_6; } -LEAN_EXPORT lean_object* l_Lean_IR_IRType_isErased___boxed(lean_object* x_1) { +case 4: +{ +uint8_t x_7; +x_7 = 1; +return x_7; +} +case 5: +{ +uint8_t x_8; +x_8 = 1; +return x_8; +} +case 10: +{ +uint8_t x_9; +x_9 = 1; +return x_9; +} +case 11: +{ +uint8_t x_10; +x_10 = 1; +return x_10; +} +default: +{ +uint8_t x_11; +x_11 = 0; +return x_11; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_IRType_isScalarOrStruct___boxed(lean_object* x_1) { _start: { uint8_t x_2; lean_object* x_3; -x_2 = l_Lean_IR_IRType_isErased(x_1); +x_2 = l_Lean_IR_IRType_isScalarOrStruct(x_1); lean_dec(x_1); x_3 = lean_box(x_2); return x_3; } } -LEAN_EXPORT uint8_t l_Lean_IR_IRType_isVoid(lean_object* x_1) { -_start: +LEAN_EXPORT uint8_t l_Lean_IR_IRType_isPossibleRef(lean_object* x_1) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 7: +{ +uint8_t x_2; +x_2 = 1; +return x_2; +} +case 8: +{ +uint8_t x_3; +x_3 = 1; +return x_3; +} +case 10: +{ +uint8_t x_4; +x_4 = 1; +return x_4; +} +case 11: +{ +uint8_t x_5; +x_5 = 1; +return x_5; +} +default: +{ +uint8_t x_6; +x_6 = 0; +return x_6; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_IRType_isPossibleRef___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_IR_IRType_isPossibleRef(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT uint8_t l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_IRType_isDefiniteRef_spec__0(lean_object* x_1, size_t x_2, size_t x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_eq(x_2, x_3); +if (x_4 == 0) +{ +uint8_t x_5; uint8_t x_6; lean_object* x_11; +x_5 = 1; +x_11 = lean_array_uget(x_1, x_2); +if (lean_obj_tag(x_11) == 10) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_12 = lean_ctor_get(x_11, 1); +lean_inc_ref(x_12); +x_13 = lean_ctor_get(x_11, 2); +lean_inc(x_13); +x_14 = lean_ctor_get(x_11, 3); +lean_inc(x_14); +lean_dec_ref(x_11); +x_15 = lean_array_get_size(x_12); +lean_dec_ref(x_12); +x_16 = lean_unsigned_to_nat(0u); +x_17 = lean_nat_dec_eq(x_15, x_16); +if (x_17 == 0) +{ +lean_dec(x_14); +lean_dec(x_13); +x_6 = x_4; +goto block_10; +} +else +{ +uint8_t x_18; +x_18 = lean_nat_dec_eq(x_13, x_16); +lean_dec(x_13); +if (x_18 == 0) +{ +lean_dec(x_14); +x_6 = x_4; +goto block_10; +} +else +{ +uint8_t x_19; +x_19 = lean_nat_dec_eq(x_14, x_16); +lean_dec(x_14); +if (x_19 == 0) +{ +x_6 = x_4; +goto block_10; +} +else +{ +return x_5; +} +} +} +} +else +{ +lean_dec(x_11); +x_6 = x_4; +goto block_10; +} +block_10: +{ +if (x_6 == 0) +{ +size_t x_7; size_t x_8; +x_7 = 1; +x_8 = lean_usize_add(x_2, x_7); +x_2 = x_8; +goto _start; +} +else +{ +return x_5; +} +} +} +else +{ +uint8_t x_20; +x_20 = 0; +return x_20; +} +} +} +LEAN_EXPORT uint8_t l_Lean_IR_IRType_isDefiniteRef(lean_object* x_1) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 7: +{ +uint8_t x_2; +x_2 = 1; +return x_2; +} +case 11: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_3 = lean_ctor_get(x_1, 1); +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_array_get_size(x_3); +x_6 = lean_nat_dec_lt(x_4, x_5); +if (x_6 == 0) +{ +uint8_t x_7; +x_7 = 1; +return x_7; +} +else +{ +if (x_6 == 0) +{ +return x_6; +} +else +{ +size_t x_8; size_t x_9; uint8_t x_10; +x_8 = 0; +x_9 = lean_usize_of_nat(x_5); +x_10 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_IRType_isDefiniteRef_spec__0(x_3, x_8, x_9); +if (x_10 == 0) +{ +return x_6; +} +else +{ +uint8_t x_11; +x_11 = 0; +return x_11; +} +} +} +} +case 10: +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_12 = lean_ctor_get(x_1, 1); +x_13 = lean_ctor_get(x_1, 2); +x_14 = lean_ctor_get(x_1, 3); +x_15 = lean_array_get_size(x_12); +x_16 = lean_unsigned_to_nat(0u); +x_17 = lean_nat_dec_eq(x_15, x_16); +if (x_17 == 0) +{ +uint8_t x_18; +x_18 = 1; +return x_18; +} +else +{ +uint8_t x_19; +x_19 = lean_nat_dec_eq(x_13, x_16); +if (x_19 == 0) +{ +return x_17; +} +else +{ +uint8_t x_20; +x_20 = lean_nat_dec_eq(x_14, x_16); +if (x_20 == 0) +{ +return x_19; +} +else +{ +uint8_t x_21; +x_21 = 0; +return x_21; +} +} +} +} +default: +{ +uint8_t x_22; +x_22 = 0; +return x_22; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_IRType_isDefiniteRef_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; uint8_t x_6; lean_object* x_7; +x_4 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_5 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_6 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_IRType_isDefiniteRef_spec__0(x_1, x_4, x_5); +lean_dec_ref(x_1); +x_7 = lean_box(x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_IRType_isDefiniteRef___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_IR_IRType_isDefiniteRef(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT uint8_t l_Lean_IR_IRType_isErased(lean_object* x_1) { +_start: +{ +if (lean_obj_tag(x_1) == 6) +{ +uint8_t x_2; +x_2 = 1; +return x_2; +} +else +{ +uint8_t x_3; +x_3 = 0; +return x_3; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_IRType_isErased___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_IR_IRType_isErased(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT uint8_t l_Lean_IR_IRType_isVoid(lean_object* x_1) { +_start: +{ +if (lean_obj_tag(x_1) == 13) +{ +uint8_t x_2; +x_2 = 1; +return x_2; +} +else +{ +uint8_t x_3; +x_3 = 0; +return x_3; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_IRType_isVoid___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_IR_IRType_isVoid(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_IRType_boxed(lean_object* x_1) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 0: +{ +lean_object* x_2; +x_2 = lean_box(7); +return x_2; +} +case 3: +{ +lean_object* x_3; +x_3 = lean_box(8); +return x_3; +} +case 4: +{ +lean_object* x_4; +x_4 = lean_box(8); +return x_4; +} +case 5: +{ +lean_object* x_5; +x_5 = lean_box(8); +return x_5; +} +case 7: +{ +return x_1; +} +case 8: +{ +return x_1; +} +case 9: +{ +lean_object* x_6; +x_6 = lean_box(7); +return x_6; +} +case 10: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_7 = lean_ctor_get(x_1, 1); +x_8 = lean_ctor_get(x_1, 2); +x_9 = lean_ctor_get(x_1, 3); +x_10 = lean_array_get_size(x_7); +x_11 = lean_unsigned_to_nat(0u); +x_12 = lean_nat_dec_eq(x_10, x_11); +if (x_12 == 0) +{ +lean_object* x_13; +x_13 = lean_box(7); +return x_13; +} +else +{ +uint8_t x_14; +x_14 = lean_nat_dec_eq(x_8, x_11); +if (x_14 == 0) +{ +lean_object* x_15; +x_15 = lean_box(7); +return x_15; +} +else +{ +uint8_t x_16; +x_16 = lean_nat_dec_eq(x_9, x_11); +if (x_16 == 0) +{ +lean_object* x_17; +x_17 = lean_box(7); +return x_17; +} +else +{ +lean_object* x_18; +x_18 = lean_box(12); +return x_18; +} +} +} +} +case 11: +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_19 = lean_ctor_get(x_1, 1); +x_20 = lean_unsigned_to_nat(0u); +x_21 = lean_array_get_size(x_19); +x_22 = lean_nat_dec_lt(x_20, x_21); +if (x_22 == 0) +{ +lean_object* x_23; +x_23 = lean_box(7); +return x_23; +} +else +{ +if (x_22 == 0) +{ +lean_object* x_24; +x_24 = lean_box(7); +return x_24; +} +else +{ +size_t x_25; size_t x_26; uint8_t x_27; +x_25 = 0; +x_26 = lean_usize_of_nat(x_21); +x_27 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_IRType_isDefiniteRef_spec__0(x_19, x_25, x_26); +if (x_27 == 0) +{ +lean_object* x_28; +x_28 = lean_box(7); +return x_28; +} +else +{ +lean_object* x_29; +x_29 = lean_box(8); +return x_29; +} +} +} +} +case 12: +{ +return x_1; +} +default: +{ +lean_object* x_30; +x_30 = lean_box(12); +return x_30; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_IRType_boxed___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_IR_IRType_boxed(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_IRType_normalizeObject(lean_object* x_1) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 7: +{ +lean_object* x_2; +x_2 = lean_box(8); +return x_2; +} +case 12: +{ +lean_object* x_3; +x_3 = lean_box(8); +return x_3; +} +case 10: +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) +{ +lean_object* x_5; size_t x_6; size_t x_7; lean_object* x_8; +x_5 = lean_ctor_get(x_1, 1); +x_6 = lean_array_size(x_5); +x_7 = 0; +x_8 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_IRType_normalizeObject_spec__0(x_6, x_7, x_5); +lean_ctor_set(x_1, 1, x_8); +return x_1; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; size_t x_13; size_t x_14; lean_object* x_15; lean_object* x_16; +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); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_1); +x_13 = lean_array_size(x_10); +x_14 = 0; +x_15 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_IRType_normalizeObject_spec__0(x_13, x_14, x_10); +x_16 = lean_alloc_ctor(10, 4, 0); +lean_ctor_set(x_16, 0, x_9); +lean_ctor_set(x_16, 1, x_15); +lean_ctor_set(x_16, 2, x_11); +lean_ctor_set(x_16, 3, x_12); +return x_16; +} +} +case 11: +{ +uint8_t x_17; +x_17 = !lean_is_exclusive(x_1); +if (x_17 == 0) +{ +lean_object* x_18; size_t x_19; size_t x_20; lean_object* x_21; +x_18 = lean_ctor_get(x_1, 1); +x_19 = lean_array_size(x_18); +x_20 = 0; +x_21 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_IRType_normalizeObject_spec__0(x_19, x_20, x_18); +lean_ctor_set(x_1, 1, x_21); +return x_1; +} +else +{ +lean_object* x_22; lean_object* x_23; size_t x_24; size_t x_25; lean_object* x_26; lean_object* x_27; +x_22 = lean_ctor_get(x_1, 0); +x_23 = lean_ctor_get(x_1, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_1); +x_24 = lean_array_size(x_23); +x_25 = 0; +x_26 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_IRType_normalizeObject_spec__0(x_24, x_25, x_23); +x_27 = lean_alloc_ctor(11, 2, 0); +lean_ctor_set(x_27, 0, x_22); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +case 13: +{ +lean_object* x_28; +x_28 = lean_box(6); +return x_28; +} +default: +{ +return x_1; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_IRType_normalizeObject_spec__0(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 = l_Lean_IR_IRType_normalizeObject(x_5); +x_9 = 1; +x_10 = lean_usize_add(x_2, x_9); +x_11 = lean_array_uset(x_7, x_2, x_8); +x_2 = x_10; +x_3 = x_11; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_IRType_normalizeObject_spec__0___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___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_IRType_normalizeObject_spec__0(x_4, x_5, x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_IRType_normalizeObject_match__1_splitter(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +switch (lean_obj_tag(x_2)) { +case 7: +{ +lean_object* x_9; lean_object* x_10; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_9 = lean_box(0); +x_10 = lean_apply_1(x_3, x_9); +return x_10; +} +case 12: +{ +lean_object* x_11; lean_object* x_12; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_11 = lean_box(0); +x_12 = lean_apply_1(x_4, x_11); +return x_12; +} +case 10: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +x_13 = lean_ctor_get(x_2, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_2, 1); +lean_inc_ref(x_14); +x_15 = lean_ctor_get(x_2, 2); +lean_inc(x_15); +x_16 = lean_ctor_get(x_2, 3); +lean_inc(x_16); +lean_dec_ref(x_2); +x_17 = lean_apply_4(x_5, x_13, x_14, x_15, x_16); +return x_17; +} +case 11: +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_18 = lean_ctor_get(x_2, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_2, 1); +lean_inc_ref(x_19); +lean_dec_ref(x_2); +x_20 = lean_apply_2(x_6, x_18, x_19); +return x_20; +} +case 13: +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_21 = lean_box(0); +x_22 = lean_apply_1(x_7, x_21); +return x_22; +} +default: +{ +lean_object* x_23; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_23 = lean_apply_6(x_8, x_2, lean_box(0), lean_box(0), lean_box(0), lean_box(0), lean_box(0)); +return x_23; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_IRType_normalizeObject_match__1_splitter___redArg(lean_object* x_1, lean_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 7: +{ +lean_object* x_8; lean_object* x_9; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_8 = lean_box(0); +x_9 = lean_apply_1(x_2, x_8); +return x_9; +} +case 12: +{ +lean_object* x_10; lean_object* x_11; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_10 = lean_box(0); +x_11 = lean_apply_1(x_3, x_10); +return x_11; +} +case 10: +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_12 = lean_ctor_get(x_1, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_13); +x_14 = lean_ctor_get(x_1, 2); +lean_inc(x_14); +x_15 = lean_ctor_get(x_1, 3); +lean_inc(x_15); +lean_dec_ref(x_1); +x_16 = lean_apply_4(x_4, x_12, x_13, x_14, x_15); +return x_16; +} +case 11: +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_17 = lean_ctor_get(x_1, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_18); +lean_dec_ref(x_1); +x_19 = lean_apply_2(x_5, x_17, x_18); +return x_19; +} +case 13: +{ +lean_object* x_20; lean_object* x_21; +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_20 = lean_box(0); +x_21 = lean_apply_1(x_6, x_20); +return x_21; +} +default: +{ +lean_object* x_22; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_22 = lean_apply_6(x_7, x_1, lean_box(0), lean_box(0), lean_box(0), lean_box(0), lean_box(0)); +return x_22; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Basic_0__Array_map__unattach_match__1_splitter(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = lean_apply_2(x_5, x_4, lean_box(0)); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Basic_0__Array_map__unattach_match__1_splitter___redArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_apply_2(x_2, x_1, lean_box(0)); +return x_3; +} +} +LEAN_EXPORT uint8_t l_Array_isEqvAux___at___00Lean_IR_IRType_compatibleWith_spec__1___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_nat_dec_eq(x_3, x_4); +if (x_5 == 1) +{ +lean_dec(x_3); +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_6 = lean_unsigned_to_nat(1u); +x_7 = lean_nat_sub(x_3, x_6); +lean_dec(x_3); +x_8 = lean_array_fget_borrowed(x_1, x_7); +x_9 = lean_array_fget_borrowed(x_2, x_7); +x_10 = l_Lean_IR_IRType_compatibleWith(x_8, x_9, x_5); +if (x_10 == 0) +{ +lean_dec(x_7); +return x_10; +} +else +{ +x_3 = x_7; +goto _start; +} +} +} +} +LEAN_EXPORT uint8_t l_Lean_IR_IRType_compatibleWith(lean_object* x_1, lean_object* x_2, uint8_t x_3) { +_start: +{ +lean_object* x_4; +switch (lean_obj_tag(x_1)) { +case 0: +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_12; +x_12 = 1; +return x_12; +} +else +{ +uint8_t x_13; +x_13 = 0; +return x_13; +} +} +case 1: +{ +if (lean_obj_tag(x_2) == 1) +{ +uint8_t x_14; +x_14 = 1; +return x_14; +} +else +{ +uint8_t x_15; +x_15 = 0; +return x_15; +} +} +case 2: +{ +if (lean_obj_tag(x_2) == 2) +{ +uint8_t x_16; +x_16 = 1; +return x_16; +} +else +{ +uint8_t x_17; +x_17 = 0; +return x_17; +} +} +case 3: +{ +if (lean_obj_tag(x_2) == 3) +{ +uint8_t x_18; +x_18 = 1; +return x_18; +} +else +{ +uint8_t x_19; +x_19 = 0; +return x_19; +} +} +case 4: +{ +if (lean_obj_tag(x_2) == 4) +{ +uint8_t x_20; +x_20 = 1; +return x_20; +} +else +{ +uint8_t x_21; +x_21 = 0; +return x_21; +} +} +case 5: +{ +if (lean_obj_tag(x_2) == 5) +{ +uint8_t x_22; +x_22 = 1; +return x_22; +} +else +{ +uint8_t x_23; +x_23 = 0; +return x_23; +} +} +case 6: +{ +x_4 = x_2; +goto block_11; +} +case 9: +{ +if (lean_obj_tag(x_2) == 9) +{ +uint8_t x_24; +x_24 = 1; +return x_24; +} +else +{ +uint8_t x_25; +x_25 = 0; +return x_25; +} +} +case 10: +{ +if (lean_obj_tag(x_2) == 10) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; uint8_t x_38; +x_26 = lean_ctor_get(x_1, 1); +x_27 = lean_ctor_get(x_1, 2); +x_28 = lean_ctor_get(x_1, 3); +x_29 = lean_ctor_get(x_2, 1); +x_30 = lean_ctor_get(x_2, 2); +x_31 = lean_ctor_get(x_2, 3); +x_38 = lean_nat_dec_eq(x_27, x_30); +if (x_38 == 0) +{ +x_32 = x_38; +goto block_37; +} +else +{ +uint8_t x_39; +x_39 = lean_nat_dec_eq(x_28, x_31); +x_32 = x_39; +goto block_37; +} +block_37: +{ +if (x_32 == 0) +{ +return x_32; +} +else +{ +lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_33 = lean_array_get_size(x_26); +x_34 = lean_array_get_size(x_29); +x_35 = lean_nat_dec_eq(x_33, x_34); +if (x_35 == 0) +{ +return x_35; +} +else +{ +uint8_t x_36; +x_36 = l_Array_isEqvAux___at___00Lean_IR_IRType_compatibleWith_spec__0___redArg(x_32, x_26, x_29, x_33); +return x_36; +} +} +} +} +else +{ +uint8_t x_40; +x_40 = 0; +return x_40; +} +} +case 11: +{ +if (lean_obj_tag(x_2) == 11) +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; +x_41 = lean_ctor_get(x_1, 1); +x_42 = lean_ctor_get(x_2, 1); +x_43 = lean_array_get_size(x_41); +x_44 = lean_array_get_size(x_42); +x_45 = lean_nat_dec_eq(x_43, x_44); +if (x_45 == 0) +{ +return x_45; +} +else +{ +uint8_t x_46; +x_46 = l_Array_isEqvAux___at___00Lean_IR_IRType_compatibleWith_spec__1___redArg(x_41, x_42, x_43); +return x_46; +} +} +else +{ +uint8_t x_47; +x_47 = 0; +return x_47; +} +} +case 13: +{ +x_4 = x_2; +goto block_11; +} +default: +{ +switch (lean_obj_tag(x_2)) { +case 7: +{ +uint8_t x_48; +x_48 = 1; +return x_48; +} +case 8: +{ +uint8_t x_49; +x_49 = 1; +return x_49; +} +case 12: +{ +uint8_t x_50; +x_50 = 1; +return x_50; +} +default: +{ +if (x_3 == 0) +{ +switch (lean_obj_tag(x_2)) { +case 6: +{ +uint8_t x_51; +x_51 = 1; +return x_51; +} +case 13: +{ +uint8_t x_52; +x_52 = 1; +return x_52; +} +default: +{ +return x_3; +} +} +} +else +{ +uint8_t x_53; +x_53 = 0; +return x_53; +} +} +} +} +} +block_11: +{ +switch (lean_obj_tag(x_4)) { +case 6: +{ +uint8_t x_5; +x_5 = 1; +return x_5; +} +case 13: +{ +uint8_t x_6; +x_6 = 1; +return x_6; +} +default: +{ +if (x_3 == 0) +{ +switch (lean_obj_tag(x_4)) { +case 7: +{ +uint8_t x_7; +x_7 = 1; +return x_7; +} +case 8: +{ +uint8_t x_8; +x_8 = 1; +return x_8; +} +case 12: +{ +uint8_t x_9; +x_9 = 1; +return x_9; +} +default: +{ +return x_3; +} +} +} +else +{ +uint8_t x_10; +x_10 = 0; +return x_10; +} +} +} +} +} +} +LEAN_EXPORT uint8_t l_Array_isEqvAux___at___00Lean_IR_IRType_compatibleWith_spec__0___redArg(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; uint8_t x_6; +x_5 = lean_unsigned_to_nat(0u); +x_6 = lean_nat_dec_eq(x_4, x_5); +if (x_6 == 1) +{ +lean_dec(x_4); +return x_6; +} +else { -if (lean_obj_tag(x_1) == 13) +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_7 = lean_unsigned_to_nat(1u); +x_8 = lean_nat_sub(x_4, x_7); +lean_dec(x_4); +x_9 = lean_array_fget_borrowed(x_2, x_8); +x_10 = lean_array_fget_borrowed(x_3, x_8); +x_11 = l_Lean_IR_IRType_compatibleWith(x_9, x_10, x_1); +if (x_11 == 0) { -uint8_t x_2; -x_2 = 1; -return x_2; +lean_dec(x_8); +return x_11; } else { -uint8_t x_3; -x_3 = 0; -return x_3; +x_4 = x_8; +goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_IR_IRType_isVoid___boxed(lean_object* x_1) { +} +LEAN_EXPORT uint8_t l_Array_isEqvAux___at___00Lean_IR_IRType_compatibleWith_spec__0(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -uint8_t x_2; lean_object* x_3; -x_2 = l_Lean_IR_IRType_isVoid(x_1); -lean_dec(x_1); -x_3 = lean_box(x_2); -return x_3; +uint8_t x_7; +x_7 = l_Array_isEqvAux___at___00Lean_IR_IRType_compatibleWith_spec__0___redArg(x_1, x_2, x_3, x_5); +return x_7; } } -LEAN_EXPORT lean_object* l_Lean_IR_IRType_boxed(lean_object* x_1) { +LEAN_EXPORT uint8_t l_Array_isEqvAux___at___00Lean_IR_IRType_compatibleWith_spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -switch (lean_obj_tag(x_1)) { -case 7: -{ -return x_1; -} -case 0: -{ -lean_object* x_2; -x_2 = lean_box(7); -return x_2; +uint8_t x_6; +x_6 = l_Array_isEqvAux___at___00Lean_IR_IRType_compatibleWith_spec__1___redArg(x_1, x_2, x_4); +return x_6; } -case 9: -{ -lean_object* x_3; -x_3 = lean_box(7); -return x_3; } -case 13: +LEAN_EXPORT lean_object* l_Array_isEqvAux___at___00Lean_IR_IRType_compatibleWith_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -lean_object* x_4; -x_4 = lean_box(12); -return x_4; +uint8_t x_7; uint8_t x_8; lean_object* x_9; +x_7 = lean_unbox(x_1); +x_8 = l_Array_isEqvAux___at___00Lean_IR_IRType_compatibleWith_spec__0(x_7, x_2, x_3, x_4, x_5, x_6); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +x_9 = lean_box(x_8); +return x_9; } -case 12: -{ -return x_1; } -case 1: +LEAN_EXPORT lean_object* l_Array_isEqvAux___at___00Lean_IR_IRType_compatibleWith_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -lean_object* x_5; -x_5 = lean_box(12); -return x_5; +uint8_t x_6; lean_object* x_7; +x_6 = l_Array_isEqvAux___at___00Lean_IR_IRType_compatibleWith_spec__1(x_1, x_2, x_3, x_4, x_5); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_7 = lean_box(x_6); +return x_7; } -case 2: -{ -lean_object* x_6; -x_6 = lean_box(12); -return x_6; } -default: +LEAN_EXPORT lean_object* l_Array_isEqvAux___at___00Lean_IR_IRType_compatibleWith_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_7; -x_7 = lean_box(8); +uint8_t x_5; uint8_t x_6; lean_object* x_7; +x_5 = lean_unbox(x_1); +x_6 = l_Array_isEqvAux___at___00Lean_IR_IRType_compatibleWith_spec__0___redArg(x_5, x_2, x_3, x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +x_7 = lean_box(x_6); return x_7; } } +LEAN_EXPORT lean_object* l_Array_isEqvAux___at___00Lean_IR_IRType_compatibleWith_spec__1___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = l_Array_isEqvAux___at___00Lean_IR_IRType_compatibleWith_spec__1___redArg(x_1, x_2, x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_5 = lean_box(x_4); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_IR_IRType_boxed___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_IR_IRType_compatibleWith___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_2; -x_2 = l_Lean_IR_IRType_boxed(x_1); +uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_4 = lean_unbox(x_3); +x_5 = l_Lean_IR_IRType_compatibleWith(x_1, x_2, x_4); +lean_dec(x_2); lean_dec(x_1); -return x_2; +x_6 = lean_box(x_5); +return x_6; } } LEAN_EXPORT lean_object* l_Lean_IR_Arg_ctorIdx(lean_object* x_1) { @@ -4708,104 +6018,121 @@ lean_dec_ref(x_1); x_5 = lean_apply_2(x_2, x_3, x_4); return x_5; } -case 2: +case 1: { -lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = lean_ctor_get(x_1, 0); lean_inc(x_6); x_7 = lean_ctor_get(x_1, 1); -lean_inc_ref(x_7); -x_8 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); -x_9 = lean_ctor_get(x_1, 2); -lean_inc_ref(x_9); +lean_inc(x_7); lean_dec_ref(x_1); -x_10 = lean_box(x_8); -x_11 = lean_apply_4(x_2, x_6, x_7, x_10, x_9); -return x_11; +x_8 = lean_apply_2(x_2, x_6, x_7); +return x_8; } -case 5: +case 2: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_12 = lean_ctor_get(x_1, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_1, 1); -lean_inc(x_13); -x_14 = lean_ctor_get(x_1, 2); -lean_inc(x_14); +lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_9 = lean_ctor_get(x_1, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_10); +x_11 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); +x_12 = lean_ctor_get(x_1, 2); +lean_inc_ref(x_12); lean_dec_ref(x_1); -x_15 = lean_apply_3(x_2, x_12, x_13, x_14); -return x_15; +x_13 = lean_box(x_11); +x_14 = lean_apply_4(x_2, x_9, x_10, x_13, x_12); +return x_14; } -case 6: +case 3: { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_1, 0); +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_15 = lean_ctor_get(x_1, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_1, 1); lean_inc(x_16); -x_17 = lean_ctor_get(x_1, 1); -lean_inc_ref(x_17); +x_17 = lean_ctor_get(x_1, 2); +lean_inc(x_17); lean_dec_ref(x_1); -x_18 = lean_apply_2(x_2, x_16, x_17); +x_18 = lean_apply_3(x_2, x_15, x_16, x_17); return x_18; } -case 7: +case 4: { -lean_object* x_19; lean_object* x_20; lean_object* x_21; +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_1, 0); lean_inc(x_19); x_20 = lean_ctor_get(x_1, 1); -lean_inc_ref(x_20); +lean_inc(x_20); +x_21 = lean_ctor_get(x_1, 2); +lean_inc(x_21); lean_dec_ref(x_1); -x_21 = lean_apply_2(x_2, x_19, x_20); -return x_21; +x_22 = lean_apply_3(x_2, x_19, x_20, x_21); +return x_22; } -case 8: +case 5: { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_1, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_1, 1); -lean_inc_ref(x_23); +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_23 = lean_ctor_get(x_1, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_1, 1); +lean_inc(x_24); +x_25 = lean_ctor_get(x_1, 2); +lean_inc(x_25); +x_26 = lean_ctor_get(x_1, 3); +lean_inc(x_26); lean_dec_ref(x_1); -x_24 = lean_apply_2(x_2, x_22, x_23); -return x_24; +x_27 = lean_apply_4(x_2, x_23, x_24, x_25, x_26); +return x_27; +} +case 9: +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_1, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_1, 1); +lean_inc(x_29); +lean_dec_ref(x_1); +x_30 = lean_apply_2(x_2, x_28, x_29); +return x_30; } case 10: { -lean_object* x_25; lean_object* x_26; -x_25 = lean_ctor_get(x_1, 0); -lean_inc(x_25); +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_1, 0); +lean_inc(x_31); lean_dec_ref(x_1); -x_26 = lean_apply_1(x_2, x_25); -return x_26; +x_32 = lean_apply_1(x_2, x_31); +return x_32; } case 11: { -lean_object* x_27; lean_object* x_28; -x_27 = lean_ctor_get(x_1, 0); -lean_inc_ref(x_27); +lean_object* x_33; lean_object* x_34; +x_33 = lean_ctor_get(x_1, 0); +lean_inc_ref(x_33); lean_dec_ref(x_1); -x_28 = lean_apply_1(x_2, x_27); -return x_28; +x_34 = lean_apply_1(x_2, x_33); +return x_34; } case 12: { -lean_object* x_29; lean_object* x_30; -x_29 = lean_ctor_get(x_1, 0); -lean_inc(x_29); +lean_object* x_35; lean_object* x_36; +x_35 = lean_ctor_get(x_1, 0); +lean_inc(x_35); lean_dec_ref(x_1); -x_30 = lean_apply_1(x_2, x_29); -return x_30; +x_36 = lean_apply_1(x_2, x_35); +return x_36; } default: { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_1, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_1, 1); -lean_inc(x_32); +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_1, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_38); lean_dec_ref(x_1); -x_33 = lean_apply_2(x_2, x_31, x_32); -return x_33; +x_39 = lean_apply_2(x_2, x_37, x_38); +return x_39; } } } @@ -5560,117 +6887,117 @@ lean_dec_ref(x_1); x_12 = lean_apply_4(x_2, x_8, x_9, x_10, x_11); return x_12; } -case 3: +case 2: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_13 = lean_ctor_get(x_1, 0); lean_inc(x_13); x_14 = lean_ctor_get(x_1, 1); lean_inc(x_14); x_15 = lean_ctor_get(x_1, 2); lean_inc(x_15); +x_16 = lean_ctor_get(x_1, 3); +lean_inc(x_16); lean_dec_ref(x_1); -x_16 = lean_apply_3(x_2, x_13, x_14, x_15); -return x_16; +x_17 = lean_apply_4(x_2, x_13, x_14, x_15, x_16); +return x_17; } -case 5: +case 3: { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_17 = lean_ctor_get(x_1, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_1, 1); +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_18 = lean_ctor_get(x_1, 0); lean_inc(x_18); -x_19 = lean_ctor_get(x_1, 2); +x_19 = lean_ctor_get(x_1, 1); lean_inc(x_19); -x_20 = lean_ctor_get(x_1, 3); +x_20 = lean_ctor_get(x_1, 2); lean_inc(x_20); -x_21 = lean_ctor_get(x_1, 4); -lean_inc(x_21); -x_22 = lean_ctor_get(x_1, 5); -lean_inc(x_22); lean_dec_ref(x_1); -x_23 = lean_apply_6(x_2, x_17, x_18, x_19, x_20, x_21, x_22); -return x_23; +x_21 = lean_apply_3(x_2, x_18, x_19, x_20); +return x_21; } -case 6: +case 4: { -lean_object* x_24; lean_object* x_25; uint8_t x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_24 = lean_ctor_get(x_1, 0); +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_22 = lean_ctor_get(x_1, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_1, 1); +lean_inc(x_23); +x_24 = lean_ctor_get(x_1, 2); lean_inc(x_24); -x_25 = lean_ctor_get(x_1, 1); +x_25 = lean_ctor_get(x_1, 3); lean_inc(x_25); -x_26 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); -x_27 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 1); -x_28 = lean_ctor_get(x_1, 2); -lean_inc(x_28); +x_26 = lean_ctor_get(x_1, 4); +lean_inc(x_26); lean_dec_ref(x_1); -x_29 = lean_box(x_26); -x_30 = lean_box(x_27); -x_31 = lean_apply_5(x_2, x_24, x_25, x_29, x_30, x_28); -return x_31; +x_27 = lean_apply_5(x_2, x_22, x_23, x_24, x_25, x_26); +return x_27; } -case 7: +case 5: { -lean_object* x_32; lean_object* x_33; uint8_t x_34; uint8_t x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_32 = lean_ctor_get(x_1, 0); +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_28 = lean_ctor_get(x_1, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_1, 1); +lean_inc(x_29); +x_30 = lean_ctor_get(x_1, 2); +lean_inc(x_30); +x_31 = lean_ctor_get(x_1, 3); +lean_inc(x_31); +x_32 = lean_ctor_get(x_1, 4); lean_inc(x_32); -x_33 = lean_ctor_get(x_1, 1); +x_33 = lean_ctor_get(x_1, 5); lean_inc(x_33); -x_34 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); -x_35 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 1); -x_36 = lean_ctor_get(x_1, 2); -lean_inc(x_36); +x_34 = lean_ctor_get(x_1, 6); +lean_inc(x_34); lean_dec_ref(x_1); -x_37 = lean_box(x_34); -x_38 = lean_box(x_35); -x_39 = lean_apply_5(x_2, x_32, x_33, x_37, x_38, x_36); -return x_39; +x_35 = lean_apply_7(x_2, x_28, x_29, x_30, x_31, x_32, x_33, x_34); +return x_35; } case 8: { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_1, 0); -lean_inc(x_40); -x_41 = lean_ctor_get(x_1, 1); -lean_inc(x_41); +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_1, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_1, 1); +lean_inc(x_37); lean_dec_ref(x_1); -x_42 = lean_apply_2(x_2, x_40, x_41); -return x_42; +x_38 = lean_apply_2(x_2, x_36, x_37); +return x_38; } case 9: { -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_43 = lean_ctor_get(x_1, 0); -lean_inc(x_43); -x_44 = lean_ctor_get(x_1, 1); -lean_inc(x_44); -x_45 = lean_ctor_get(x_1, 2); -lean_inc(x_45); -x_46 = lean_ctor_get(x_1, 3); -lean_inc_ref(x_46); +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_39 = lean_ctor_get(x_1, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_1, 1); +lean_inc(x_40); +x_41 = lean_ctor_get(x_1, 2); +lean_inc(x_41); +x_42 = lean_ctor_get(x_1, 3); +lean_inc_ref(x_42); lean_dec_ref(x_1); -x_47 = lean_apply_4(x_2, x_43, x_44, x_45, x_46); -return x_47; +x_43 = lean_apply_4(x_2, x_39, x_40, x_41, x_42); +return x_43; } case 10: { -lean_object* x_48; lean_object* x_49; -x_48 = lean_ctor_get(x_1, 0); -lean_inc(x_48); +lean_object* x_44; lean_object* x_45; +x_44 = lean_ctor_get(x_1, 0); +lean_inc(x_44); lean_dec_ref(x_1); -x_49 = lean_apply_1(x_2, x_48); -return x_49; +x_45 = lean_apply_1(x_2, x_44); +return x_45; } case 11: { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_1, 0); -lean_inc(x_50); -x_51 = lean_ctor_get(x_1, 1); -lean_inc_ref(x_51); +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_1, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_47); lean_dec_ref(x_1); -x_52 = lean_apply_2(x_2, x_50, x_51); -return x_52; +x_48 = lean_apply_2(x_2, x_46, x_47); +return x_48; } case 12: { @@ -5678,18 +7005,20 @@ return x_2; } default: { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_53 = lean_ctor_get(x_1, 0); +lean_object* x_49; lean_object* x_50; uint8_t x_51; uint8_t x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_49 = lean_ctor_get(x_1, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_1, 1); +lean_inc(x_50); +x_51 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); +x_52 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 1); +x_53 = lean_ctor_get(x_1, 2); lean_inc(x_53); -x_54 = lean_ctor_get(x_1, 1); -lean_inc(x_54); -x_55 = lean_ctor_get(x_1, 2); -lean_inc(x_55); -x_56 = lean_ctor_get(x_1, 3); -lean_inc(x_56); lean_dec(x_1); -x_57 = lean_apply_4(x_2, x_53, x_54, x_55, x_56); -return x_57; +x_54 = lean_box(x_51); +x_55 = lean_box(x_52); +x_56 = lean_apply_5(x_2, x_49, x_50, x_54, x_55, x_53); +return x_56; } } } @@ -6070,14 +7399,14 @@ return x_4; case 4: { lean_object* x_5; -x_5 = lean_ctor_get(x_1, 3); +x_5 = lean_ctor_get(x_1, 4); lean_inc(x_5); return x_5; } case 5: { lean_object* x_6; -x_6 = lean_ctor_get(x_1, 5); +x_6 = lean_ctor_get(x_1, 6); lean_inc(x_6); return x_6; } @@ -6227,176 +7556,182 @@ x_21 = !lean_is_exclusive(x_1); if (x_21 == 0) { lean_object* x_22; -x_22 = lean_ctor_get(x_1, 3); +x_22 = lean_ctor_get(x_1, 4); lean_dec(x_22); -lean_ctor_set(x_1, 3, x_2); +lean_ctor_set(x_1, 4, x_2); return x_1; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; x_23 = lean_ctor_get(x_1, 0); x_24 = lean_ctor_get(x_1, 1); x_25 = lean_ctor_get(x_1, 2); +x_26 = lean_ctor_get(x_1, 3); +lean_inc(x_26); lean_inc(x_25); lean_inc(x_24); lean_inc(x_23); lean_dec(x_1); -x_26 = lean_alloc_ctor(4, 4, 0); -lean_ctor_set(x_26, 0, x_23); -lean_ctor_set(x_26, 1, x_24); -lean_ctor_set(x_26, 2, x_25); -lean_ctor_set(x_26, 3, x_2); -return x_26; +x_27 = lean_alloc_ctor(4, 5, 0); +lean_ctor_set(x_27, 0, x_23); +lean_ctor_set(x_27, 1, x_24); +lean_ctor_set(x_27, 2, x_25); +lean_ctor_set(x_27, 3, x_26); +lean_ctor_set(x_27, 4, x_2); +return x_27; } } case 5: { -uint8_t x_27; -x_27 = !lean_is_exclusive(x_1); -if (x_27 == 0) +uint8_t x_28; +x_28 = !lean_is_exclusive(x_1); +if (x_28 == 0) { -lean_object* x_28; -x_28 = lean_ctor_get(x_1, 5); -lean_dec(x_28); -lean_ctor_set(x_1, 5, x_2); +lean_object* x_29; +x_29 = lean_ctor_get(x_1, 6); +lean_dec(x_29); +lean_ctor_set(x_1, 6, x_2); return x_1; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_29 = lean_ctor_get(x_1, 0); -x_30 = lean_ctor_get(x_1, 1); -x_31 = lean_ctor_get(x_1, 2); -x_32 = lean_ctor_get(x_1, 3); -x_33 = lean_ctor_get(x_1, 4); +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_30 = lean_ctor_get(x_1, 0); +x_31 = lean_ctor_get(x_1, 1); +x_32 = lean_ctor_get(x_1, 2); +x_33 = lean_ctor_get(x_1, 3); +x_34 = lean_ctor_get(x_1, 4); +x_35 = lean_ctor_get(x_1, 5); +lean_inc(x_35); +lean_inc(x_34); lean_inc(x_33); lean_inc(x_32); lean_inc(x_31); lean_inc(x_30); -lean_inc(x_29); lean_dec(x_1); -x_34 = lean_alloc_ctor(5, 6, 0); -lean_ctor_set(x_34, 0, x_29); -lean_ctor_set(x_34, 1, x_30); -lean_ctor_set(x_34, 2, x_31); -lean_ctor_set(x_34, 3, x_32); -lean_ctor_set(x_34, 4, x_33); -lean_ctor_set(x_34, 5, x_2); -return x_34; +x_36 = lean_alloc_ctor(5, 7, 0); +lean_ctor_set(x_36, 0, x_30); +lean_ctor_set(x_36, 1, x_31); +lean_ctor_set(x_36, 2, x_32); +lean_ctor_set(x_36, 3, x_33); +lean_ctor_set(x_36, 4, x_34); +lean_ctor_set(x_36, 5, x_35); +lean_ctor_set(x_36, 6, x_2); +return x_36; } } case 3: { -uint8_t x_35; -x_35 = !lean_is_exclusive(x_1); -if (x_35 == 0) +uint8_t x_37; +x_37 = !lean_is_exclusive(x_1); +if (x_37 == 0) { -lean_object* x_36; -x_36 = lean_ctor_get(x_1, 2); -lean_dec(x_36); +lean_object* x_38; +x_38 = lean_ctor_get(x_1, 2); +lean_dec(x_38); lean_ctor_set(x_1, 2, x_2); return x_1; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_1, 0); -x_38 = lean_ctor_get(x_1, 1); -lean_inc(x_38); -lean_inc(x_37); +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_1, 0); +x_40 = lean_ctor_get(x_1, 1); +lean_inc(x_40); +lean_inc(x_39); lean_dec(x_1); -x_39 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_39, 0, x_37); -lean_ctor_set(x_39, 1, x_38); -lean_ctor_set(x_39, 2, x_2); -return x_39; +x_41 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +lean_ctor_set(x_41, 2, x_2); +return x_41; } } case 6: { -uint8_t x_40; -x_40 = !lean_is_exclusive(x_1); -if (x_40 == 0) +uint8_t x_42; +x_42 = !lean_is_exclusive(x_1); +if (x_42 == 0) { -lean_object* x_41; -x_41 = lean_ctor_get(x_1, 2); -lean_dec(x_41); +lean_object* x_43; +x_43 = lean_ctor_get(x_1, 2); +lean_dec(x_43); lean_ctor_set(x_1, 2, x_2); return x_1; } else { -lean_object* x_42; lean_object* x_43; uint8_t x_44; uint8_t x_45; lean_object* x_46; -x_42 = lean_ctor_get(x_1, 0); -x_43 = lean_ctor_get(x_1, 1); -x_44 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); -x_45 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 1); -lean_inc(x_43); -lean_inc(x_42); +lean_object* x_44; lean_object* x_45; uint8_t x_46; uint8_t x_47; lean_object* x_48; +x_44 = lean_ctor_get(x_1, 0); +x_45 = lean_ctor_get(x_1, 1); +x_46 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); +x_47 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 1); +lean_inc(x_45); +lean_inc(x_44); lean_dec(x_1); -x_46 = lean_alloc_ctor(6, 3, 2); -lean_ctor_set(x_46, 0, x_42); -lean_ctor_set(x_46, 1, x_43); -lean_ctor_set(x_46, 2, x_2); -lean_ctor_set_uint8(x_46, sizeof(void*)*3, x_44); -lean_ctor_set_uint8(x_46, sizeof(void*)*3 + 1, x_45); -return x_46; +x_48 = lean_alloc_ctor(6, 3, 2); +lean_ctor_set(x_48, 0, x_44); +lean_ctor_set(x_48, 1, x_45); +lean_ctor_set(x_48, 2, x_2); +lean_ctor_set_uint8(x_48, sizeof(void*)*3, x_46); +lean_ctor_set_uint8(x_48, sizeof(void*)*3 + 1, x_47); +return x_48; } } case 7: { -uint8_t x_47; -x_47 = !lean_is_exclusive(x_1); -if (x_47 == 0) +uint8_t x_49; +x_49 = !lean_is_exclusive(x_1); +if (x_49 == 0) { -lean_object* x_48; -x_48 = lean_ctor_get(x_1, 2); -lean_dec(x_48); +lean_object* x_50; +x_50 = lean_ctor_get(x_1, 2); +lean_dec(x_50); lean_ctor_set(x_1, 2, x_2); return x_1; } else { -lean_object* x_49; lean_object* x_50; uint8_t x_51; uint8_t x_52; lean_object* x_53; -x_49 = lean_ctor_get(x_1, 0); -x_50 = lean_ctor_get(x_1, 1); -x_51 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); -x_52 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 1); -lean_inc(x_50); -lean_inc(x_49); +lean_object* x_51; lean_object* x_52; uint8_t x_53; uint8_t x_54; lean_object* x_55; +x_51 = lean_ctor_get(x_1, 0); +x_52 = lean_ctor_get(x_1, 1); +x_53 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); +x_54 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 1); +lean_inc(x_52); +lean_inc(x_51); lean_dec(x_1); -x_53 = lean_alloc_ctor(7, 3, 2); -lean_ctor_set(x_53, 0, x_49); -lean_ctor_set(x_53, 1, x_50); -lean_ctor_set(x_53, 2, x_2); -lean_ctor_set_uint8(x_53, sizeof(void*)*3, x_51); -lean_ctor_set_uint8(x_53, sizeof(void*)*3 + 1, x_52); -return x_53; +x_55 = lean_alloc_ctor(7, 3, 2); +lean_ctor_set(x_55, 0, x_51); +lean_ctor_set(x_55, 1, x_52); +lean_ctor_set(x_55, 2, x_2); +lean_ctor_set_uint8(x_55, sizeof(void*)*3, x_53); +lean_ctor_set_uint8(x_55, sizeof(void*)*3 + 1, x_54); +return x_55; } } case 8: { -uint8_t x_54; -x_54 = !lean_is_exclusive(x_1); -if (x_54 == 0) +uint8_t x_56; +x_56 = !lean_is_exclusive(x_1); +if (x_56 == 0) { -lean_object* x_55; -x_55 = lean_ctor_get(x_1, 1); -lean_dec(x_55); +lean_object* x_57; +x_57 = lean_ctor_get(x_1, 1); +lean_dec(x_57); lean_ctor_set(x_1, 1, x_2); return x_1; } else { -lean_object* x_56; lean_object* x_57; -x_56 = lean_ctor_get(x_1, 0); -lean_inc(x_56); +lean_object* x_58; lean_object* x_59; +x_58 = lean_ctor_get(x_1, 0); +lean_inc(x_58); lean_dec(x_1); -x_57 = lean_alloc_ctor(8, 2, 0); -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_2); -return x_57; +x_59 = lean_alloc_ctor(8, 2, 0); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_2); +return x_59; } } default: @@ -6448,7 +7783,7 @@ goto block_6; case 4: { lean_object* x_10; -x_10 = lean_ctor_get(x_1, 3); +x_10 = lean_ctor_get(x_1, 4); lean_inc(x_10); x_2 = x_10; goto block_6; @@ -6456,7 +7791,7 @@ goto block_6; case 5: { lean_object* x_11; -x_11 = lean_ctor_get(x_1, 5); +x_11 = lean_ctor_get(x_1, 6); lean_inc(x_11); x_2 = x_11; goto block_6; @@ -6829,7 +8164,7 @@ goto block_6; case 4: { lean_object* x_11; -x_11 = lean_ctor_get(x_1, 3); +x_11 = lean_ctor_get(x_1, 4); lean_inc(x_11); x_3 = x_11; goto block_6; @@ -6837,7 +8172,7 @@ goto block_6; case 5: { lean_object* x_12; -x_12 = lean_ctor_get(x_1, 5); +x_12 = lean_ctor_get(x_1, 6); lean_inc(x_12); x_3 = x_12; goto block_6; @@ -7584,7 +8919,7 @@ static lean_object* _init_l_Lean_IR_Decl_updateBody_x21___closed__3() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_IR_Decl_updateBody_x21___closed__2; x_2 = lean_unsigned_to_nat(9u); -x_3 = lean_unsigned_to_nat(381u); +x_3 = lean_unsigned_to_nat(474u); x_4 = l_Lean_IR_Decl_updateBody_x21___closed__1; x_5 = l_Lean_IR_Decl_updateBody_x21___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -7841,9 +9176,9 @@ goto block_52; block_45: { lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_nat_add(x_39, x_41); +x_42 = lean_nat_add(x_40, x_41); lean_dec(x_41); -lean_dec(x_39); +lean_dec(x_40); if (lean_is_scalar(x_36)) { x_43 = lean_alloc_ctor(0, 5, 0); } else { @@ -7862,7 +9197,7 @@ if (lean_is_scalar(x_26)) { lean_ctor_set(x_44, 0, x_38); lean_ctor_set(x_44, 1, x_28); lean_ctor_set(x_44, 2, x_29); -lean_ctor_set(x_44, 3, x_40); +lean_ctor_set(x_44, 3, x_39); lean_ctor_set(x_44, 4, x_43); return x_44; } @@ -7888,8 +9223,8 @@ if (lean_obj_tag(x_31) == 0) lean_object* x_50; x_50 = lean_ctor_get(x_31, 0); lean_inc(x_50); -x_39 = x_49; -x_40 = x_48; +x_39 = x_48; +x_40 = x_49; x_41 = x_50; goto block_45; } @@ -7897,8 +9232,8 @@ else { lean_object* x_51; x_51 = lean_unsigned_to_nat(0u); -x_39 = x_49; -x_40 = x_48; +x_39 = x_48; +x_40 = x_49; x_41 = x_51; goto block_45; } @@ -10106,9 +11441,9 @@ goto block_184; block_177: { lean_object* x_174; lean_object* x_175; lean_object* x_176; -x_174 = lean_nat_add(x_172, x_173); +x_174 = lean_nat_add(x_171, x_173); lean_dec(x_173); -lean_dec(x_172); +lean_dec(x_171); if (lean_is_scalar(x_168)) { x_175 = lean_alloc_ctor(0, 5, 0); } else { @@ -10127,7 +11462,7 @@ if (lean_is_scalar(x_158)) { lean_ctor_set(x_176, 0, x_170); lean_ctor_set(x_176, 1, x_160); lean_ctor_set(x_176, 2, x_161); -lean_ctor_set(x_176, 3, x_171); +lean_ctor_set(x_176, 3, x_172); lean_ctor_set(x_176, 4, x_175); return x_176; } @@ -10153,8 +11488,8 @@ if (lean_obj_tag(x_163) == 0) lean_object* x_182; x_182 = lean_ctor_get(x_163, 0); lean_inc(x_182); -x_171 = x_180; -x_172 = x_181; +x_171 = x_181; +x_172 = x_180; x_173 = x_182; goto block_177; } @@ -10162,8 +11497,8 @@ else { lean_object* x_183; x_183 = lean_unsigned_to_nat(0u); -x_171 = x_180; -x_172 = x_181; +x_171 = x_181; +x_172 = x_180; x_173 = x_183; goto block_177; } @@ -10652,9 +11987,9 @@ goto block_303; block_295: { lean_object* x_286; lean_object* x_287; uint8_t x_288; -x_286 = lean_nat_add(x_284, x_285); +x_286 = lean_nat_add(x_283, x_285); lean_dec(x_285); -lean_dec(x_284); +lean_dec(x_283); lean_inc_ref(x_255); if (lean_is_scalar(x_280)) { x_287 = lean_alloc_ctor(0, 5, 0); @@ -10681,7 +12016,7 @@ lean_dec(x_292); x_293 = lean_ctor_get(x_255, 0); lean_dec(x_293); lean_ctor_set(x_255, 4, x_287); -lean_ctor_set(x_255, 3, x_283); +lean_ctor_set(x_255, 3, x_284); lean_ctor_set(x_255, 2, x_274); lean_ctor_set(x_255, 1, x_273); lean_ctor_set(x_255, 0, x_282); @@ -10695,7 +12030,7 @@ x_294 = lean_alloc_ctor(0, 5, 0); lean_ctor_set(x_294, 0, x_282); lean_ctor_set(x_294, 1, x_273); lean_ctor_set(x_294, 2, x_274); -lean_ctor_set(x_294, 3, x_283); +lean_ctor_set(x_294, 3, x_284); lean_ctor_set(x_294, 4, x_287); return x_294; } @@ -10722,8 +12057,8 @@ if (lean_obj_tag(x_276) == 0) lean_object* x_301; x_301 = lean_ctor_get(x_276, 0); lean_inc(x_301); -x_283 = x_299; -x_284 = x_300; +x_283 = x_300; +x_284 = x_299; x_285 = x_301; goto block_295; } @@ -10731,8 +12066,8 @@ else { lean_object* x_302; x_302 = lean_unsigned_to_nat(0u); -x_283 = x_299; -x_284 = x_300; +x_283 = x_300; +x_284 = x_299; x_285 = x_302; goto block_295; } @@ -10818,9 +12153,9 @@ goto block_338; block_330: { lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; -x_326 = lean_nat_add(x_324, x_325); +x_326 = lean_nat_add(x_323, x_325); lean_dec(x_325); -lean_dec(x_324); +lean_dec(x_323); lean_inc_ref(x_255); if (lean_is_scalar(x_320)) { x_327 = lean_alloc_ctor(0, 5, 0); @@ -10851,7 +12186,7 @@ if (lean_is_scalar(x_328)) { lean_ctor_set(x_329, 0, x_322); lean_ctor_set(x_329, 1, x_313); lean_ctor_set(x_329, 2, x_314); -lean_ctor_set(x_329, 3, x_323); +lean_ctor_set(x_329, 3, x_324); lean_ctor_set(x_329, 4, x_327); return x_329; } @@ -10877,8 +12212,8 @@ if (lean_obj_tag(x_316) == 0) lean_object* x_336; x_336 = lean_ctor_get(x_316, 0); lean_inc(x_336); -x_323 = x_334; -x_324 = x_335; +x_323 = x_335; +x_324 = x_334; x_325 = x_336; goto block_330; } @@ -10886,8 +12221,8 @@ else { lean_object* x_337; x_337 = lean_unsigned_to_nat(0u); -x_323 = x_334; -x_324 = x_335; +x_323 = x_335; +x_324 = x_334; x_325 = x_337; goto block_330; } @@ -12346,7 +13681,7 @@ case 1: { if (lean_obj_tag(x_3) == 1) { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; x_25 = lean_ctor_get(x_2, 0); lean_inc(x_25); x_26 = lean_ctor_get(x_2, 1); @@ -12357,103 +13692,115 @@ lean_inc(x_27); x_28 = lean_ctor_get(x_3, 1); lean_inc(x_28); lean_dec_ref(x_3); -x_4 = x_25; -x_5 = x_26; -x_6 = x_27; -x_7 = x_28; -goto block_10; +x_29 = lean_nat_dec_eq(x_25, x_27); +lean_dec(x_27); +lean_dec(x_25); +if (x_29 == 0) +{ +lean_dec(x_28); +lean_dec(x_26); +return x_29; +} +else +{ +uint8_t x_30; +x_30 = l_Lean_IR_VarId_alphaEqv(x_1, x_26, x_28); +lean_dec(x_28); +lean_dec(x_26); +return x_30; +} } else { -uint8_t x_29; +uint8_t x_31; lean_dec_ref(x_3); lean_dec_ref(x_2); -x_29 = 0; -return x_29; +x_31 = 0; +return x_31; } } case 2: { if (lean_obj_tag(x_3) == 2) { -lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; lean_object* x_37; uint8_t x_38; uint8_t x_42; -x_30 = lean_ctor_get(x_2, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_2, 1); -lean_inc_ref(x_31); -x_32 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); -x_33 = lean_ctor_get(x_2, 2); +lean_object* x_32; lean_object* x_33; uint8_t x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; lean_object* x_39; uint8_t x_40; uint8_t x_44; +x_32 = lean_ctor_get(x_2, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_2, 1); lean_inc_ref(x_33); -lean_dec_ref(x_2); -x_34 = lean_ctor_get(x_3, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_3, 1); +x_34 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); +x_35 = lean_ctor_get(x_2, 2); lean_inc_ref(x_35); -x_36 = lean_ctor_get_uint8(x_3, sizeof(void*)*3); -x_37 = lean_ctor_get(x_3, 2); +lean_dec_ref(x_2); +x_36 = lean_ctor_get(x_3, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_3, 1); lean_inc_ref(x_37); +x_38 = lean_ctor_get_uint8(x_3, sizeof(void*)*3); +x_39 = lean_ctor_get(x_3, 2); +lean_inc_ref(x_39); lean_dec_ref(x_3); -x_42 = l_Lean_IR_VarId_alphaEqv(x_1, x_30, x_34); -lean_dec(x_34); -lean_dec(x_30); -if (x_42 == 0) +x_44 = l_Lean_IR_VarId_alphaEqv(x_1, x_32, x_36); +lean_dec(x_36); +lean_dec(x_32); +if (x_44 == 0) { -lean_dec_ref(x_35); -lean_dec_ref(x_31); -x_38 = x_42; -goto block_41; +lean_dec_ref(x_37); +lean_dec_ref(x_33); +x_40 = x_44; +goto block_43; } else { -uint8_t x_43; -x_43 = l_Lean_IR_instBEqCtorInfo_beq(x_31, x_35); -lean_dec_ref(x_35); -lean_dec_ref(x_31); -x_38 = x_43; -goto block_41; +uint8_t x_45; +x_45 = l_Lean_IR_instBEqCtorInfo_beq(x_33, x_37); +lean_dec_ref(x_37); +lean_dec_ref(x_33); +x_40 = x_45; +goto block_43; } -block_41: +block_43: { -if (x_38 == 0) +if (x_40 == 0) { -lean_dec_ref(x_37); -lean_dec_ref(x_33); -return x_38; +lean_dec_ref(x_39); +lean_dec_ref(x_35); +return x_40; } else { -if (x_32 == 0) +if (x_34 == 0) { -if (x_36 == 0) +if (x_38 == 0) { -uint8_t x_39; -x_39 = l_Lean_IR_args_alphaEqv(x_1, x_33, x_37); -lean_dec_ref(x_37); -lean_dec_ref(x_33); -return x_39; +uint8_t x_41; +x_41 = l_Lean_IR_args_alphaEqv(x_1, x_35, x_39); +lean_dec_ref(x_39); +lean_dec_ref(x_35); +return x_41; } else { -lean_dec_ref(x_37); -lean_dec_ref(x_33); -return x_32; +lean_dec_ref(x_39); +lean_dec_ref(x_35); +return x_34; } } else { -if (x_36 == 0) +if (x_38 == 0) { -lean_dec_ref(x_37); -lean_dec_ref(x_33); -return x_36; +lean_dec_ref(x_39); +lean_dec_ref(x_35); +return x_38; } else { -uint8_t x_40; -x_40 = l_Lean_IR_args_alphaEqv(x_1, x_33, x_37); -lean_dec_ref(x_37); -lean_dec_ref(x_33); -return x_40; +uint8_t x_42; +x_42 = l_Lean_IR_args_alphaEqv(x_1, x_35, x_39); +lean_dec_ref(x_39); +lean_dec_ref(x_35); +return x_42; } } } @@ -12461,353 +13808,353 @@ return x_40; } else { -uint8_t x_44; +uint8_t x_46; lean_dec_ref(x_3); lean_dec_ref(x_2); -x_44 = 0; -return x_44; +x_46 = 0; +return x_46; } } case 3: { if (lean_obj_tag(x_3) == 3) { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_45 = lean_ctor_get(x_2, 0); -lean_inc(x_45); -x_46 = lean_ctor_get(x_2, 1); -lean_inc(x_46); -lean_dec_ref(x_2); -x_47 = lean_ctor_get(x_3, 0); +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_47 = lean_ctor_get(x_2, 1); lean_inc(x_47); -x_48 = lean_ctor_get(x_3, 1); +x_48 = lean_ctor_get(x_2, 2); lean_inc(x_48); +lean_dec_ref(x_2); +x_49 = lean_ctor_get(x_3, 1); +lean_inc(x_49); +x_50 = lean_ctor_get(x_3, 2); +lean_inc(x_50); lean_dec_ref(x_3); -x_4 = x_45; -x_5 = x_46; -x_6 = x_47; -x_7 = x_48; +x_4 = x_47; +x_5 = x_48; +x_6 = x_49; +x_7 = x_50; goto block_10; } else { -uint8_t x_49; +uint8_t x_51; lean_dec_ref(x_3); lean_dec_ref(x_2); -x_49 = 0; -return x_49; +x_51 = 0; +return x_51; } } case 4: { if (lean_obj_tag(x_3) == 4) { -lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_50 = lean_ctor_get(x_2, 0); -lean_inc(x_50); -x_51 = lean_ctor_get(x_2, 1); -lean_inc(x_51); -lean_dec_ref(x_2); -x_52 = lean_ctor_get(x_3, 0); +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_52 = lean_ctor_get(x_2, 1); lean_inc(x_52); -x_53 = lean_ctor_get(x_3, 1); +x_53 = lean_ctor_get(x_2, 2); lean_inc(x_53); +lean_dec_ref(x_2); +x_54 = lean_ctor_get(x_3, 1); +lean_inc(x_54); +x_55 = lean_ctor_get(x_3, 2); +lean_inc(x_55); lean_dec_ref(x_3); -x_4 = x_50; -x_5 = x_51; -x_6 = x_52; -x_7 = x_53; +x_4 = x_52; +x_5 = x_53; +x_6 = x_54; +x_7 = x_55; goto block_10; } else { -uint8_t x_54; +uint8_t x_56; lean_dec_ref(x_3); lean_dec_ref(x_2); -x_54 = 0; -return x_54; +x_56 = 0; +return x_56; } } case 5: { if (lean_obj_tag(x_3) == 5) { -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; uint8_t x_64; -x_55 = lean_ctor_get(x_2, 0); -lean_inc(x_55); -x_56 = lean_ctor_get(x_2, 1); -lean_inc(x_56); -x_57 = lean_ctor_get(x_2, 2); +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; uint8_t x_66; +x_57 = lean_ctor_get(x_2, 1); lean_inc(x_57); -lean_dec_ref(x_2); -x_58 = lean_ctor_get(x_3, 0); +x_58 = lean_ctor_get(x_2, 2); lean_inc(x_58); -x_59 = lean_ctor_get(x_3, 1); +x_59 = lean_ctor_get(x_2, 3); lean_inc(x_59); -x_60 = lean_ctor_get(x_3, 2); +lean_dec_ref(x_2); +x_60 = lean_ctor_get(x_3, 1); lean_inc(x_60); +x_61 = lean_ctor_get(x_3, 2); +lean_inc(x_61); +x_62 = lean_ctor_get(x_3, 3); +lean_inc(x_62); lean_dec_ref(x_3); -x_64 = lean_nat_dec_eq(x_55, x_58); -lean_dec(x_58); -lean_dec(x_55); -if (x_64 == 0) +x_66 = lean_nat_dec_eq(x_57, x_60); +lean_dec(x_60); +lean_dec(x_57); +if (x_66 == 0) { -lean_dec(x_59); -lean_dec(x_56); -x_61 = x_64; -goto block_63; +lean_dec(x_61); +lean_dec(x_58); +x_63 = x_66; +goto block_65; } else { -uint8_t x_65; -x_65 = lean_nat_dec_eq(x_56, x_59); -lean_dec(x_59); -lean_dec(x_56); -x_61 = x_65; -goto block_63; +uint8_t x_67; +x_67 = lean_nat_dec_eq(x_58, x_61); +lean_dec(x_61); +lean_dec(x_58); +x_63 = x_67; +goto block_65; } -block_63: +block_65: { -if (x_61 == 0) +if (x_63 == 0) { -lean_dec(x_60); -lean_dec(x_57); -return x_61; +lean_dec(x_62); +lean_dec(x_59); +return x_63; } else { -uint8_t x_62; -x_62 = l_Lean_IR_VarId_alphaEqv(x_1, x_57, x_60); -lean_dec(x_60); -lean_dec(x_57); -return x_62; +uint8_t x_64; +x_64 = l_Lean_IR_VarId_alphaEqv(x_1, x_59, x_62); +lean_dec(x_62); +lean_dec(x_59); +return x_64; } } } else { -uint8_t x_66; +uint8_t x_68; lean_dec_ref(x_3); lean_dec_ref(x_2); -x_66 = 0; -return x_66; +x_68 = 0; +return x_68; } } case 6: { if (lean_obj_tag(x_3) == 6) { -lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_67 = lean_ctor_get(x_2, 0); -lean_inc(x_67); -x_68 = lean_ctor_get(x_2, 1); -lean_inc_ref(x_68); -lean_dec_ref(x_2); -x_69 = lean_ctor_get(x_3, 0); +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_69 = lean_ctor_get(x_2, 0); lean_inc(x_69); -x_70 = lean_ctor_get(x_3, 1); +x_70 = lean_ctor_get(x_2, 1); lean_inc_ref(x_70); +lean_dec_ref(x_2); +x_71 = lean_ctor_get(x_3, 0); +lean_inc(x_71); +x_72 = lean_ctor_get(x_3, 1); +lean_inc_ref(x_72); lean_dec_ref(x_3); -x_11 = x_67; -x_12 = x_68; -x_13 = x_69; -x_14 = x_70; +x_11 = x_69; +x_12 = x_70; +x_13 = x_71; +x_14 = x_72; goto block_17; } else { -uint8_t x_71; +uint8_t x_73; lean_dec_ref(x_3); lean_dec_ref(x_2); -x_71 = 0; -return x_71; +x_73 = 0; +return x_73; } } case 7: { if (lean_obj_tag(x_3) == 7) { -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_72 = lean_ctor_get(x_2, 0); -lean_inc(x_72); -x_73 = lean_ctor_get(x_2, 1); -lean_inc_ref(x_73); -lean_dec_ref(x_2); -x_74 = lean_ctor_get(x_3, 0); +lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_74 = lean_ctor_get(x_2, 0); lean_inc(x_74); -x_75 = lean_ctor_get(x_3, 1); +x_75 = lean_ctor_get(x_2, 1); lean_inc_ref(x_75); +lean_dec_ref(x_2); +x_76 = lean_ctor_get(x_3, 0); +lean_inc(x_76); +x_77 = lean_ctor_get(x_3, 1); +lean_inc_ref(x_77); lean_dec_ref(x_3); -x_11 = x_72; -x_12 = x_73; -x_13 = x_74; -x_14 = x_75; +x_11 = x_74; +x_12 = x_75; +x_13 = x_76; +x_14 = x_77; goto block_17; } else { -uint8_t x_76; +uint8_t x_78; lean_dec_ref(x_3); lean_dec_ref(x_2); -x_76 = 0; -return x_76; +x_78 = 0; +return x_78; } } case 8: { if (lean_obj_tag(x_3) == 8) { -lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; uint8_t x_81; -x_77 = lean_ctor_get(x_2, 0); -lean_inc(x_77); -x_78 = lean_ctor_get(x_2, 1); -lean_inc_ref(x_78); -lean_dec_ref(x_2); -x_79 = lean_ctor_get(x_3, 0); +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; uint8_t x_83; +x_79 = lean_ctor_get(x_2, 0); lean_inc(x_79); -x_80 = lean_ctor_get(x_3, 1); +x_80 = lean_ctor_get(x_2, 1); lean_inc_ref(x_80); +lean_dec_ref(x_2); +x_81 = lean_ctor_get(x_3, 0); +lean_inc(x_81); +x_82 = lean_ctor_get(x_3, 1); +lean_inc_ref(x_82); lean_dec_ref(x_3); -x_81 = l_Lean_IR_VarId_alphaEqv(x_1, x_77, x_79); +x_83 = l_Lean_IR_VarId_alphaEqv(x_1, x_79, x_81); +lean_dec(x_81); lean_dec(x_79); -lean_dec(x_77); -if (x_81 == 0) +if (x_83 == 0) { +lean_dec_ref(x_82); lean_dec_ref(x_80); -lean_dec_ref(x_78); -return x_81; +return x_83; } else { -uint8_t x_82; -x_82 = l_Lean_IR_args_alphaEqv(x_1, x_78, x_80); +uint8_t x_84; +x_84 = l_Lean_IR_args_alphaEqv(x_1, x_80, x_82); +lean_dec_ref(x_82); lean_dec_ref(x_80); -lean_dec_ref(x_78); -return x_82; +return x_84; } } else { -uint8_t x_83; +uint8_t x_85; lean_dec_ref(x_3); lean_dec_ref(x_2); -x_83 = 0; -return x_83; +x_85 = 0; +return x_85; } } case 9: { if (lean_obj_tag(x_3) == 9) { -lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; uint8_t x_88; -x_84 = lean_ctor_get(x_2, 0); -lean_inc(x_84); -x_85 = lean_ctor_get(x_2, 1); -lean_inc(x_85); -lean_dec_ref(x_2); -x_86 = lean_ctor_get(x_3, 0); +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; uint8_t x_90; +x_86 = lean_ctor_get(x_2, 0); lean_inc(x_86); -x_87 = lean_ctor_get(x_3, 1); +x_87 = lean_ctor_get(x_2, 1); lean_inc(x_87); +lean_dec_ref(x_2); +x_88 = lean_ctor_get(x_3, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_3, 1); +lean_inc(x_89); lean_dec_ref(x_3); -x_88 = l_Lean_IR_instBEqIRType_beq(x_84, x_86); -if (x_88 == 0) +x_90 = l_Lean_IR_instBEqIRType_beq(x_86, x_88); +if (x_90 == 0) { +lean_dec(x_89); lean_dec(x_87); -lean_dec(x_85); -return x_88; +return x_90; } else { -uint8_t x_89; -x_89 = l_Lean_IR_VarId_alphaEqv(x_1, x_85, x_87); +uint8_t x_91; +x_91 = l_Lean_IR_VarId_alphaEqv(x_1, x_87, x_89); +lean_dec(x_89); lean_dec(x_87); -lean_dec(x_85); -return x_89; +return x_91; } } else { -uint8_t x_90; +uint8_t x_92; lean_dec_ref(x_3); lean_dec_ref(x_2); -x_90 = 0; -return x_90; +x_92 = 0; +return x_92; } } case 10: { if (lean_obj_tag(x_3) == 10) { -lean_object* x_91; lean_object* x_92; uint8_t x_93; -x_91 = lean_ctor_get(x_2, 0); -lean_inc(x_91); +lean_object* x_93; lean_object* x_94; uint8_t x_95; +x_93 = lean_ctor_get(x_2, 0); +lean_inc(x_93); lean_dec_ref(x_2); -x_92 = lean_ctor_get(x_3, 0); -lean_inc(x_92); +x_94 = lean_ctor_get(x_3, 0); +lean_inc(x_94); lean_dec_ref(x_3); -x_93 = l_Lean_IR_VarId_alphaEqv(x_1, x_91, x_92); -lean_dec(x_92); -lean_dec(x_91); -return x_93; +x_95 = l_Lean_IR_VarId_alphaEqv(x_1, x_93, x_94); +lean_dec(x_94); +lean_dec(x_93); +return x_95; } else { -uint8_t x_94; +uint8_t x_96; lean_dec_ref(x_3); lean_dec_ref(x_2); -x_94 = 0; -return x_94; +x_96 = 0; +return x_96; } } case 11: { if (lean_obj_tag(x_3) == 11) { -lean_object* x_95; lean_object* x_96; uint8_t x_97; -x_95 = lean_ctor_get(x_2, 0); -lean_inc_ref(x_95); +lean_object* x_97; lean_object* x_98; uint8_t x_99; +x_97 = lean_ctor_get(x_2, 0); +lean_inc_ref(x_97); lean_dec_ref(x_2); -x_96 = lean_ctor_get(x_3, 0); -lean_inc_ref(x_96); +x_98 = lean_ctor_get(x_3, 0); +lean_inc_ref(x_98); lean_dec_ref(x_3); -x_97 = l_Lean_IR_instBEqLitVal_beq(x_95, x_96); -lean_dec_ref(x_96); -lean_dec_ref(x_95); -return x_97; +x_99 = l_Lean_IR_instBEqLitVal_beq(x_97, x_98); +lean_dec_ref(x_98); +lean_dec_ref(x_97); +return x_99; } else { -uint8_t x_98; +uint8_t x_100; lean_dec_ref(x_3); lean_dec_ref(x_2); -x_98 = 0; -return x_98; +x_100 = 0; +return x_100; } } default: { if (lean_obj_tag(x_3) == 12) { -lean_object* x_99; lean_object* x_100; uint8_t x_101; -x_99 = lean_ctor_get(x_2, 0); -lean_inc(x_99); +lean_object* x_101; lean_object* x_102; uint8_t x_103; +x_101 = lean_ctor_get(x_2, 0); +lean_inc(x_101); lean_dec_ref(x_2); -x_100 = lean_ctor_get(x_3, 0); -lean_inc(x_100); +x_102 = lean_ctor_get(x_3, 0); +lean_inc(x_102); lean_dec_ref(x_3); -x_101 = l_Lean_IR_VarId_alphaEqv(x_1, x_99, x_100); -lean_dec(x_100); -lean_dec(x_99); -return x_101; +x_103 = l_Lean_IR_VarId_alphaEqv(x_1, x_101, x_102); +lean_dec(x_102); +lean_dec(x_101); +return x_103; } else { -uint8_t x_102; +uint8_t x_104; lean_dec_ref(x_3); lean_dec_ref(x_2); -x_102 = 0; -return x_102; +x_104 = 0; +return x_104; } } } @@ -13078,7 +14425,7 @@ return x_6; LEAN_EXPORT uint8_t l_Lean_IR_FnBody_alphaEqv(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -uint8_t x_4; lean_object* x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; uint8_t x_12; uint8_t x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; uint8_t x_19; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; uint8_t x_30; lean_object* x_31; +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; uint8_t x_8; lean_object* x_12; uint8_t x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; uint8_t x_17; uint8_t x_18; uint8_t x_19; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; uint8_t x_30; lean_object* x_31; switch (lean_obj_tag(x_2)) { case 0: { @@ -13384,20 +14731,20 @@ if (lean_obj_tag(x_3) == 4) lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; uint8_t x_100; uint8_t x_104; x_92 = lean_ctor_get(x_2, 0); lean_inc(x_92); -x_93 = lean_ctor_get(x_2, 1); +x_93 = lean_ctor_get(x_2, 2); lean_inc(x_93); -x_94 = lean_ctor_get(x_2, 2); +x_94 = lean_ctor_get(x_2, 3); lean_inc(x_94); -x_95 = lean_ctor_get(x_2, 3); +x_95 = lean_ctor_get(x_2, 4); lean_inc(x_95); lean_dec_ref(x_2); x_96 = lean_ctor_get(x_3, 0); lean_inc(x_96); -x_97 = lean_ctor_get(x_3, 1); +x_97 = lean_ctor_get(x_3, 2); lean_inc(x_97); -x_98 = lean_ctor_get(x_3, 2); +x_98 = lean_ctor_get(x_3, 3); lean_inc(x_98); -x_99 = lean_ctor_get(x_3, 3); +x_99 = lean_ctor_get(x_3, 4); lean_inc(x_99); lean_dec_ref(x_3); x_104 = l_Lean_IR_VarId_alphaEqv(x_1, x_92, x_96); @@ -13469,28 +14816,28 @@ if (lean_obj_tag(x_3) == 5) lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; uint8_t x_119; uint8_t x_120; uint8_t x_125; x_107 = lean_ctor_get(x_2, 0); lean_inc(x_107); -x_108 = lean_ctor_get(x_2, 1); +x_108 = lean_ctor_get(x_2, 2); lean_inc(x_108); -x_109 = lean_ctor_get(x_2, 2); +x_109 = lean_ctor_get(x_2, 3); lean_inc(x_109); -x_110 = lean_ctor_get(x_2, 3); +x_110 = lean_ctor_get(x_2, 4); lean_inc(x_110); -x_111 = lean_ctor_get(x_2, 4); +x_111 = lean_ctor_get(x_2, 5); lean_inc(x_111); -x_112 = lean_ctor_get(x_2, 5); +x_112 = lean_ctor_get(x_2, 6); lean_inc(x_112); lean_dec_ref(x_2); x_113 = lean_ctor_get(x_3, 0); lean_inc(x_113); -x_114 = lean_ctor_get(x_3, 1); +x_114 = lean_ctor_get(x_3, 2); lean_inc(x_114); -x_115 = lean_ctor_get(x_3, 2); +x_115 = lean_ctor_get(x_3, 3); lean_inc(x_115); -x_116 = lean_ctor_get(x_3, 3); +x_116 = lean_ctor_get(x_3, 4); lean_inc(x_116); -x_117 = lean_ctor_get(x_3, 4); +x_117 = lean_ctor_get(x_3, 5); lean_inc(x_117); -x_118 = lean_ctor_get(x_3, 5); +x_118 = lean_ctor_get(x_3, 6); lean_inc(x_118); lean_dec_ref(x_3); x_119 = lean_nat_dec_eq(x_109, x_115); @@ -13895,37 +15242,37 @@ return x_184; } block_11: { -if (x_4 == 0) +if (x_8 == 0) { -if (x_6 == 0) +if (x_7 == 0) { -x_1 = x_8; +x_1 = x_4; x_2 = x_5; -x_3 = x_7; +x_3 = x_6; goto _start; } else { -lean_dec(x_8); -lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); -return x_4; +lean_dec(x_4); +return x_8; } } else { -if (x_6 == 0) +if (x_7 == 0) { -lean_dec(x_8); -lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); -return x_6; +lean_dec(x_4); +return x_7; } else { -x_1 = x_8; +x_1 = x_4; x_2 = x_5; -x_3 = x_7; +x_3 = x_6; goto _start; } } @@ -13934,47 +15281,47 @@ goto _start; { if (x_19 == 0) { -lean_dec(x_18); lean_dec(x_16); lean_dec(x_14); +lean_dec(x_12); return x_19; } else { -if (x_17 == 0) +if (x_13 == 0) { -if (x_15 == 0) +if (x_17 == 0) { x_4 = x_12; x_5 = x_14; -x_6 = x_13; -x_7 = x_16; +x_6 = x_16; +x_7 = x_15; x_8 = x_18; goto block_11; } else { -lean_dec(x_18); lean_dec(x_16); lean_dec(x_14); -return x_17; +lean_dec(x_12); +return x_13; } } else { -if (x_15 == 0) +if (x_17 == 0) { -lean_dec(x_18); lean_dec(x_16); lean_dec(x_14); -return x_15; +lean_dec(x_12); +return x_17; } else { x_4 = x_12; x_5 = x_14; -x_6 = x_13; -x_7 = x_16; +x_6 = x_16; +x_7 = x_15; x_8 = x_18; goto block_11; } @@ -13991,13 +15338,13 @@ if (x_32 == 0) { lean_dec(x_28); lean_dec(x_23); -x_12 = x_25; -x_13 = x_30; +x_12 = x_21; +x_13 = x_24; x_14 = x_26; -x_15 = x_29; +x_15 = x_30; x_16 = x_31; -x_17 = x_24; -x_18 = x_21; +x_17 = x_29; +x_18 = x_25; x_19 = x_32; goto block_20; } @@ -14007,13 +15354,13 @@ uint8_t x_33; x_33 = lean_nat_dec_eq(x_23, x_28); lean_dec(x_28); lean_dec(x_23); -x_12 = x_25; -x_13 = x_30; +x_12 = x_21; +x_13 = x_24; x_14 = x_26; -x_15 = x_29; +x_15 = x_30; x_16 = x_31; -x_17 = x_24; -x_18 = x_21; +x_17 = x_29; +x_18 = x_25; x_19 = x_33; goto block_20; } diff --git a/stage0/stdlib/Lean/Compiler/IR/Borrow.c b/stage0/stdlib/Lean/Compiler/IR/Borrow.c index 5e1dd4867950..ffc14c9ba579 100644 --- a/stage0/stdlib/Lean/Compiler/IR/Borrow.c +++ b/stage0/stdlib/Lean/Compiler/IR/Borrow.c @@ -127,6 +127,7 @@ uint64_t lean_uint64_shift_right(uint64_t, uint64_t); LEAN_EXPORT lean_object* l_Lean_IR_Borrow_ownParamsUsingArgs___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_Borrow_ApplyParamMap_visitFnBody_spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* lean_nat_to_int(lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_Borrow_collectExpr___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_contains___at___00Lean_IR_Borrow_ownArgsIfParam_spec__0___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_Borrow_InitParamMap_visitDecls_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_div(lean_object*, lean_object*); @@ -307,7 +308,7 @@ static lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___00Lean_IR_Bo lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_Borrow_collectFnBody_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Id_instMonad___lam__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_Borrow_collectExpr(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_Borrow_collectExpr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_IR_IRType_isScalar(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Borrow_ownArg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_Borrow_ApplyParamMap_visitFnBody___closed__0; @@ -2679,7 +2680,7 @@ goto block_9; case 4: { lean_object* x_50; -x_50 = lean_ctor_get(x_3, 3); +x_50 = lean_ctor_get(x_3, 4); lean_inc(x_50); x_4 = x_50; goto block_9; @@ -2687,7 +2688,7 @@ goto block_9; case 5: { lean_object* x_51; -x_51 = lean_ctor_get(x_3, 5); +x_51 = lean_ctor_get(x_3, 6); lean_inc(x_51); x_4 = x_51; goto block_9; @@ -4440,223 +4441,423 @@ lean_dec_ref(x_1); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_IR_Borrow_collectExpr(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_IR_Borrow_collectExpr(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -switch (lean_obj_tag(x_2)) { -case 1: +switch (lean_obj_tag(x_3)) { +case 0: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_5 = lean_ctor_get(x_2, 1); -lean_inc(x_5); -lean_dec_ref(x_2); -x_6 = l_Lean_IR_Borrow_ownVar(x_1, x_3, x_4); -x_7 = lean_ctor_get(x_6, 1); -lean_inc(x_7); -lean_dec_ref(x_6); -x_8 = l_Lean_IR_Borrow_ownVar(x_5, x_3, x_7); +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_6 = lean_ctor_get(x_3, 1); +lean_inc_ref(x_6); lean_dec_ref(x_3); -return x_8; +x_7 = l_Lean_IR_Borrow_ownVar(x_1, x_4, x_5); +x_8 = lean_ctor_get(x_7, 1); +lean_inc(x_8); +lean_dec_ref(x_7); +x_9 = l_Lean_IR_Borrow_ownArgsIfParam(x_6, x_4, x_8); +lean_dec_ref(x_4); +lean_dec_ref(x_6); +return x_9; } -case 2: +case 1: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_9 = lean_ctor_get(x_2, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_2, 2); -lean_inc_ref(x_10); -lean_dec_ref(x_2); -x_11 = l_Lean_IR_Borrow_ownVar(x_1, x_3, x_4); +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_ctor_get(x_3, 1); +lean_inc(x_10); +lean_dec_ref(x_3); +x_11 = l_Lean_IR_Borrow_ownVar(x_1, x_4, x_5); x_12 = lean_ctor_get(x_11, 1); lean_inc(x_12); lean_dec_ref(x_11); -x_13 = l_Lean_IR_Borrow_ownVar(x_9, x_3, x_12); -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec_ref(x_13); -x_15 = l_Lean_IR_Borrow_ownArgsIfParam(x_10, x_3, x_14); -lean_dec_ref(x_3); -lean_dec_ref(x_10); -return x_15; +x_13 = l_Lean_IR_Borrow_ownVar(x_10, x_4, x_12); +lean_dec_ref(x_4); +return x_13; } -case 0: +case 2: { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_16 = lean_ctor_get(x_2, 1); -lean_inc_ref(x_16); -lean_dec_ref(x_2); -x_17 = l_Lean_IR_Borrow_ownVar(x_1, x_3, x_4); -x_18 = lean_ctor_get(x_17, 1); -lean_inc(x_18); -lean_dec_ref(x_17); -x_19 = l_Lean_IR_Borrow_ownArgsIfParam(x_16, x_3, x_18); +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_14 = lean_ctor_get(x_3, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_3, 2); +lean_inc_ref(x_15); lean_dec_ref(x_3); +x_16 = l_Lean_IR_Borrow_ownVar(x_1, x_4, x_5); +x_17 = lean_ctor_get(x_16, 1); +lean_inc(x_17); lean_dec_ref(x_16); -return x_19; +x_18 = l_Lean_IR_Borrow_ownVar(x_14, x_4, x_17); +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +lean_dec_ref(x_18); +x_20 = l_Lean_IR_Borrow_ownArgsIfParam(x_15, x_4, x_19); +lean_dec_ref(x_4); +lean_dec_ref(x_15); +return x_20; } case 3: { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_20 = lean_ctor_get(x_2, 1); -lean_inc(x_20); -lean_dec_ref(x_2); -lean_inc(x_20); -x_35 = l_Lean_IR_Borrow_isOwned(x_20, x_3, x_4); -x_36 = lean_ctor_get(x_35, 0); -lean_inc(x_36); -x_37 = lean_unbox(x_36); -lean_dec(x_36); -if (x_37 == 0) +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_45; lean_object* x_46; uint8_t x_47; +x_21 = lean_ctor_get(x_3, 2); +lean_inc(x_21); +lean_dec_ref(x_3); +lean_inc(x_21); +x_45 = l_Lean_IR_Borrow_isOwned(x_21, x_4, x_5); +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_unbox(x_46); +lean_dec(x_46); +if (x_47 == 0) { -lean_object* x_38; -x_38 = lean_ctor_get(x_35, 1); -lean_inc(x_38); -lean_dec_ref(x_35); -x_21 = x_3; -x_22 = x_38; -goto block_34; +lean_object* x_48; +x_48 = lean_ctor_get(x_45, 1); +lean_inc(x_48); +lean_dec_ref(x_45); +x_22 = x_4; +x_23 = x_48; +goto block_44; } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_39 = lean_ctor_get(x_35, 1); -lean_inc(x_39); -lean_dec_ref(x_35); +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_45, 1); +lean_inc(x_49); +lean_dec_ref(x_45); lean_inc(x_1); -x_40 = l_Lean_IR_Borrow_ownVar(x_1, x_3, x_39); -x_41 = lean_ctor_get(x_40, 1); -lean_inc(x_41); -lean_dec_ref(x_40); -x_21 = x_3; -x_22 = x_41; -goto block_34; +x_50 = l_Lean_IR_Borrow_ownVar(x_1, x_4, x_49); +x_51 = lean_ctor_get(x_50, 1); +lean_inc(x_51); +lean_dec_ref(x_50); +x_22 = x_4; +x_23 = x_51; +goto block_44; } -block_34: +block_44: { -lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_23 = l_Lean_IR_Borrow_isOwned(x_1, x_21, x_22); -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_unbox(x_24); +lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_24 = l_Lean_IR_Borrow_isOwned(x_1, x_22, x_23); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_unbox(x_25); +lean_dec(x_25); +if (x_26 == 0) +{ +uint8_t x_27; +lean_dec_ref(x_22); +lean_dec(x_21); +x_27 = !lean_is_exclusive(x_24); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_24, 0); +lean_dec(x_28); +x_29 = lean_box(0); +lean_ctor_set(x_24, 0, x_29); +return x_24; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_24, 1); +lean_inc(x_30); lean_dec(x_24); -if (x_25 == 0) +x_31 = lean_box(0); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +return x_32; +} +} +else { -uint8_t x_26; -lean_dec_ref(x_21); -lean_dec(x_20); -x_26 = !lean_is_exclusive(x_23); -if (x_26 == 0) +uint8_t x_33; +x_33 = !lean_is_exclusive(x_24); +if (x_33 == 0) { -lean_object* x_27; lean_object* x_28; -x_27 = lean_ctor_get(x_23, 0); -lean_dec(x_27); -x_28 = lean_box(0); -lean_ctor_set(x_23, 0, x_28); -return x_23; +lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_34 = lean_ctor_get(x_24, 1); +x_35 = lean_ctor_get(x_24, 0); +lean_dec(x_35); +x_36 = l_Lean_IR_IRType_isScalar(x_2); +if (x_36 == 0) +{ +lean_object* x_37; +lean_free_object(x_24); +x_37 = l_Lean_IR_Borrow_ownVar(x_21, x_22, x_34); +lean_dec_ref(x_22); +return x_37; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_23, 1); -lean_inc(x_29); -lean_dec(x_23); -x_30 = lean_box(0); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_29); -return x_31; +lean_object* x_38; +lean_dec_ref(x_22); +lean_dec(x_21); +x_38 = lean_box(0); +lean_ctor_set(x_24, 0, x_38); +return x_24; } } else { -lean_object* x_32; lean_object* x_33; -x_32 = lean_ctor_get(x_23, 1); -lean_inc(x_32); -lean_dec_ref(x_23); -x_33 = l_Lean_IR_Borrow_ownVar(x_20, x_21, x_32); -lean_dec_ref(x_21); -return x_33; +lean_object* x_39; uint8_t x_40; +x_39 = lean_ctor_get(x_24, 1); +lean_inc(x_39); +lean_dec(x_24); +x_40 = l_Lean_IR_IRType_isScalar(x_2); +if (x_40 == 0) +{ +lean_object* x_41; +x_41 = l_Lean_IR_Borrow_ownVar(x_21, x_22, x_39); +lean_dec_ref(x_22); +return x_41; +} +else +{ +lean_object* x_42; lean_object* x_43; +lean_dec_ref(x_22); +lean_dec(x_21); +x_42 = lean_box(0); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_39); +return x_43; +} +} } } } case 6: { -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_42 = lean_ctor_get(x_2, 0); -lean_inc(x_42); -x_43 = lean_ctor_get(x_2, 1); -lean_inc_ref(x_43); -lean_dec_ref(x_2); -x_44 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_44, 0, x_42); -lean_inc_ref(x_3); -x_45 = l_Lean_IR_Borrow_getParamInfo(x_44, x_3, x_4); -x_46 = lean_ctor_get(x_45, 0); -lean_inc(x_46); -x_47 = lean_ctor_get(x_45, 1); -lean_inc(x_47); -lean_dec_ref(x_45); -x_48 = l_Lean_IR_Borrow_ownVar(x_1, x_3, x_47); -x_49 = lean_ctor_get(x_48, 1); -lean_inc(x_49); -lean_dec_ref(x_48); -x_50 = l_Lean_IR_Borrow_ownArgsUsingParams(x_43, x_46, x_3, x_49); +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_52 = lean_ctor_get(x_3, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_3, 1); +lean_inc_ref(x_53); lean_dec_ref(x_3); -lean_dec(x_46); -lean_dec_ref(x_43); -return x_50; +x_54 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_54, 0, x_52); +lean_inc_ref(x_4); +x_55 = l_Lean_IR_Borrow_getParamInfo(x_54, x_4, x_5); +x_56 = lean_ctor_get(x_55, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_55, 1); +lean_inc(x_57); +lean_dec_ref(x_55); +x_58 = l_Lean_IR_Borrow_ownVar(x_1, x_4, x_57); +x_59 = lean_ctor_get(x_58, 1); +lean_inc(x_59); +lean_dec_ref(x_58); +x_60 = l_Lean_IR_Borrow_ownArgsUsingParams(x_53, x_56, x_4, x_59); +lean_dec_ref(x_4); +lean_dec(x_56); +lean_dec_ref(x_53); +return x_60; +} +case 7: +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_61 = lean_ctor_get(x_3, 1); +lean_inc_ref(x_61); +lean_dec_ref(x_3); +x_62 = l_Lean_IR_Borrow_ownVar(x_1, x_4, x_5); +x_63 = lean_ctor_get(x_62, 1); +lean_inc(x_63); +lean_dec_ref(x_62); +x_64 = l_Lean_IR_Borrow_ownArgs(x_61, x_4, x_63); +lean_dec_ref(x_4); +lean_dec_ref(x_61); +return x_64; } case 8: { -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_51 = lean_ctor_get(x_2, 0); -lean_inc(x_51); -x_52 = lean_ctor_get(x_2, 1); -lean_inc_ref(x_52); -lean_dec_ref(x_2); -x_53 = l_Lean_IR_Borrow_ownVar(x_1, x_3, x_4); -x_54 = lean_ctor_get(x_53, 1); -lean_inc(x_54); -lean_dec_ref(x_53); -x_55 = l_Lean_IR_Borrow_ownVar(x_51, x_3, x_54); -x_56 = lean_ctor_get(x_55, 1); -lean_inc(x_56); -lean_dec_ref(x_55); -x_57 = l_Lean_IR_Borrow_ownArgs(x_52, x_3, x_56); +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_65 = lean_ctor_get(x_3, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_3, 1); +lean_inc_ref(x_66); lean_dec_ref(x_3); -lean_dec_ref(x_52); -return x_57; +x_67 = l_Lean_IR_Borrow_ownVar(x_1, x_4, x_5); +x_68 = lean_ctor_get(x_67, 1); +lean_inc(x_68); +lean_dec_ref(x_67); +x_69 = l_Lean_IR_Borrow_ownVar(x_65, x_4, x_68); +x_70 = lean_ctor_get(x_69, 1); +lean_inc(x_70); +lean_dec_ref(x_69); +x_71 = l_Lean_IR_Borrow_ownArgs(x_66, x_4, x_70); +lean_dec_ref(x_4); +lean_dec_ref(x_66); +return x_71; } -case 7: +case 9: { -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_58 = lean_ctor_get(x_2, 1); -lean_inc_ref(x_58); -lean_dec_ref(x_2); -x_59 = l_Lean_IR_Borrow_ownVar(x_1, x_3, x_4); -x_60 = lean_ctor_get(x_59, 1); -lean_inc(x_60); -lean_dec_ref(x_59); -x_61 = l_Lean_IR_Borrow_ownArgs(x_58, x_3, x_60); +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_72 = lean_ctor_get(x_3, 1); +lean_inc(x_72); lean_dec_ref(x_3); -lean_dec_ref(x_58); -return x_61; +x_73 = l_Lean_IR_Borrow_ownVar(x_1, x_4, x_5); +x_74 = lean_ctor_get(x_73, 1); +lean_inc(x_74); +lean_dec_ref(x_73); +x_75 = l_Lean_IR_Borrow_ownVar(x_72, x_4, x_74); +lean_dec_ref(x_4); +return x_75; +} +case 10: +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_100; lean_object* x_101; uint8_t x_102; +x_76 = lean_ctor_get(x_3, 0); +lean_inc(x_76); +lean_dec_ref(x_3); +lean_inc(x_76); +x_100 = l_Lean_IR_Borrow_isOwned(x_76, x_4, x_5); +x_101 = lean_ctor_get(x_100, 0); +lean_inc(x_101); +x_102 = lean_unbox(x_101); +lean_dec(x_101); +if (x_102 == 0) +{ +lean_object* x_103; +x_103 = lean_ctor_get(x_100, 1); +lean_inc(x_103); +lean_dec_ref(x_100); +x_77 = x_4; +x_78 = x_103; +goto block_99; +} +else +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_104 = lean_ctor_get(x_100, 1); +lean_inc(x_104); +lean_dec_ref(x_100); +lean_inc(x_1); +x_105 = l_Lean_IR_Borrow_ownVar(x_1, x_4, x_104); +x_106 = lean_ctor_get(x_105, 1); +lean_inc(x_106); +lean_dec_ref(x_105); +x_77 = x_4; +x_78 = x_106; +goto block_99; +} +block_99: +{ +lean_object* x_79; lean_object* x_80; uint8_t x_81; +x_79 = l_Lean_IR_Borrow_isOwned(x_1, x_77, x_78); +x_80 = lean_ctor_get(x_79, 0); +lean_inc(x_80); +x_81 = lean_unbox(x_80); +lean_dec(x_80); +if (x_81 == 0) +{ +uint8_t x_82; +lean_dec_ref(x_77); +lean_dec(x_76); +x_82 = !lean_is_exclusive(x_79); +if (x_82 == 0) +{ +lean_object* x_83; lean_object* x_84; +x_83 = lean_ctor_get(x_79, 0); +lean_dec(x_83); +x_84 = lean_box(0); +lean_ctor_set(x_79, 0, x_84); +return x_79; +} +else +{ +lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_85 = lean_ctor_get(x_79, 1); +lean_inc(x_85); +lean_dec(x_79); +x_86 = lean_box(0); +x_87 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_87, 0, x_86); +lean_ctor_set(x_87, 1, x_85); +return x_87; +} +} +else +{ +uint8_t x_88; +x_88 = !lean_is_exclusive(x_79); +if (x_88 == 0) +{ +lean_object* x_89; lean_object* x_90; uint8_t x_91; +x_89 = lean_ctor_get(x_79, 1); +x_90 = lean_ctor_get(x_79, 0); +lean_dec(x_90); +x_91 = l_Lean_IR_IRType_isScalar(x_2); +if (x_91 == 0) +{ +lean_object* x_92; +lean_free_object(x_79); +x_92 = l_Lean_IR_Borrow_ownVar(x_76, x_77, x_89); +lean_dec_ref(x_77); +return x_92; +} +else +{ +lean_object* x_93; +lean_dec_ref(x_77); +lean_dec(x_76); +x_93 = lean_box(0); +lean_ctor_set(x_79, 0, x_93); +return x_79; +} +} +else +{ +lean_object* x_94; uint8_t x_95; +x_94 = lean_ctor_get(x_79, 1); +lean_inc(x_94); +lean_dec(x_79); +x_95 = l_Lean_IR_IRType_isScalar(x_2); +if (x_95 == 0) +{ +lean_object* x_96; +x_96 = l_Lean_IR_Borrow_ownVar(x_76, x_77, x_94); +lean_dec_ref(x_77); +return x_96; +} +else +{ +lean_object* x_97; lean_object* x_98; +lean_dec_ref(x_77); +lean_dec(x_76); +x_97 = lean_box(0); +x_98 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_98, 0, x_97); +lean_ctor_set(x_98, 1, x_94); +return x_98; +} +} +} +} } default: { -lean_object* x_62; lean_object* x_63; +lean_object* x_107; lean_object* x_108; +lean_dec_ref(x_4); lean_dec_ref(x_3); -lean_dec_ref(x_2); lean_dec(x_1); -x_62 = lean_box(0); -x_63 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_63, 0, x_62); -lean_ctor_set(x_63, 1, x_4); -return x_63; +x_107 = lean_box(0); +x_108 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_108, 0, x_107); +lean_ctor_set(x_108, 1, x_5); +return x_108; } } } } +LEAN_EXPORT lean_object* l_Lean_IR_Borrow_collectExpr___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_IR_Borrow_collectExpr(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_2); +return x_6; +} +} LEAN_EXPORT lean_object* l_Lean_IR_Borrow_preserveTailCall(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { @@ -4928,9 +5129,9 @@ goto block_52; block_45: { lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_nat_add(x_39, x_41); +x_42 = lean_nat_add(x_40, x_41); lean_dec(x_41); -lean_dec(x_39); +lean_dec(x_40); if (lean_is_scalar(x_36)) { x_43 = lean_alloc_ctor(0, 5, 0); } else { @@ -4949,7 +5150,7 @@ if (lean_is_scalar(x_26)) { lean_ctor_set(x_44, 0, x_38); lean_ctor_set(x_44, 1, x_28); lean_ctor_set(x_44, 2, x_29); -lean_ctor_set(x_44, 3, x_40); +lean_ctor_set(x_44, 3, x_39); lean_ctor_set(x_44, 4, x_43); return x_44; } @@ -4975,8 +5176,8 @@ if (lean_obj_tag(x_31) == 0) lean_object* x_50; x_50 = lean_ctor_get(x_31, 0); lean_inc(x_50); -x_39 = x_49; -x_40 = x_48; +x_39 = x_48; +x_40 = x_49; x_41 = x_50; goto block_45; } @@ -4984,8 +5185,8 @@ else { lean_object* x_51; x_51 = lean_unsigned_to_nat(0u); -x_39 = x_49; -x_40 = x_48; +x_39 = x_48; +x_40 = x_49; x_41 = x_51; goto block_45; } @@ -6007,173 +6208,176 @@ goto _start; } case 0: { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; x_23 = lean_ctor_get(x_1, 0); lean_inc(x_23); -x_24 = lean_ctor_get(x_1, 2); -lean_inc_ref(x_24); -x_25 = lean_ctor_get(x_1, 3); -lean_inc(x_25); +x_24 = lean_ctor_get(x_1, 1); +lean_inc(x_24); +x_25 = lean_ctor_get(x_1, 2); +lean_inc_ref(x_25); +x_26 = lean_ctor_get(x_1, 3); +lean_inc(x_26); lean_dec_ref(x_1); lean_inc_ref(x_2); -lean_inc(x_25); -x_26 = l_Lean_IR_Borrow_collectFnBody(x_25, x_2, x_3); -x_27 = lean_ctor_get(x_26, 1); -lean_inc(x_27); -lean_dec_ref(x_26); +lean_inc(x_26); +x_27 = l_Lean_IR_Borrow_collectFnBody(x_26, x_2, x_3); +x_28 = lean_ctor_get(x_27, 1); +lean_inc(x_28); +lean_dec_ref(x_27); lean_inc_ref(x_2); -lean_inc_ref(x_24); +lean_inc_ref(x_25); lean_inc(x_23); -x_28 = l_Lean_IR_Borrow_collectExpr(x_23, x_24, x_2, x_27); -x_29 = lean_ctor_get(x_28, 1); -lean_inc(x_29); -lean_dec_ref(x_28); -x_30 = l_Lean_IR_Borrow_preserveTailCall(x_23, x_24, x_25, x_2, x_29); -lean_dec(x_25); +x_29 = l_Lean_IR_Borrow_collectExpr(x_23, x_24, x_25, x_2, x_28); +lean_dec(x_24); +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +lean_dec_ref(x_29); +x_31 = l_Lean_IR_Borrow_preserveTailCall(x_23, x_25, x_26, x_2, x_30); +lean_dec(x_26); lean_dec(x_23); -return x_30; +return x_31; } case 11: { -uint8_t x_31; -x_31 = !lean_is_exclusive(x_1); -if (x_31 == 0) +uint8_t x_32; +x_32 = !lean_is_exclusive(x_1); +if (x_32 == 0) { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_32 = lean_ctor_get(x_1, 0); -x_33 = lean_ctor_get(x_1, 1); -x_34 = lean_ctor_get(x_2, 2); -lean_inc(x_34); +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_33 = lean_ctor_get(x_1, 0); +x_34 = lean_ctor_get(x_1, 1); +x_35 = lean_ctor_get(x_2, 2); +lean_inc(x_35); lean_ctor_set_tag(x_1, 1); -lean_ctor_set(x_1, 1, x_32); -lean_ctor_set(x_1, 0, x_34); +lean_ctor_set(x_1, 1, x_33); +lean_ctor_set(x_1, 0, x_35); lean_inc_ref(x_2); -x_35 = l_Lean_IR_Borrow_getParamInfo(x_1, x_2, x_3); -x_36 = lean_ctor_get(x_35, 0); -lean_inc(x_36); -x_37 = lean_ctor_get(x_35, 1); +x_36 = l_Lean_IR_Borrow_getParamInfo(x_1, x_2, x_3); +x_37 = lean_ctor_get(x_36, 0); lean_inc(x_37); -lean_dec_ref(x_35); -x_38 = l_Lean_IR_Borrow_ownArgsUsingParams(x_33, x_36, x_2, x_37); -x_39 = lean_ctor_get(x_38, 1); -lean_inc(x_39); -lean_dec_ref(x_38); -x_40 = l_Lean_IR_Borrow_ownParamsUsingArgs(x_33, x_36, x_2, x_39); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +lean_dec_ref(x_36); +x_39 = l_Lean_IR_Borrow_ownArgsUsingParams(x_34, x_37, x_2, x_38); +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +lean_dec_ref(x_39); +x_41 = l_Lean_IR_Borrow_ownParamsUsingArgs(x_34, x_37, x_2, x_40); lean_dec_ref(x_2); -lean_dec(x_36); -lean_dec_ref(x_33); -return x_40; +lean_dec(x_37); +lean_dec_ref(x_34); +return x_41; } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_41 = lean_ctor_get(x_1, 0); -x_42 = lean_ctor_get(x_1, 1); +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_42 = lean_ctor_get(x_1, 0); +x_43 = lean_ctor_get(x_1, 1); +lean_inc(x_43); lean_inc(x_42); -lean_inc(x_41); lean_dec(x_1); -x_43 = lean_ctor_get(x_2, 2); -lean_inc(x_43); -x_44 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_41); +x_44 = lean_ctor_get(x_2, 2); +lean_inc(x_44); +x_45 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_42); lean_inc_ref(x_2); -x_45 = l_Lean_IR_Borrow_getParamInfo(x_44, x_2, x_3); -x_46 = lean_ctor_get(x_45, 0); -lean_inc(x_46); -x_47 = lean_ctor_get(x_45, 1); +x_46 = l_Lean_IR_Borrow_getParamInfo(x_45, x_2, x_3); +x_47 = lean_ctor_get(x_46, 0); lean_inc(x_47); -lean_dec_ref(x_45); -x_48 = l_Lean_IR_Borrow_ownArgsUsingParams(x_42, x_46, x_2, x_47); -x_49 = lean_ctor_get(x_48, 1); -lean_inc(x_49); -lean_dec_ref(x_48); -x_50 = l_Lean_IR_Borrow_ownParamsUsingArgs(x_42, x_46, x_2, x_49); +x_48 = lean_ctor_get(x_46, 1); +lean_inc(x_48); +lean_dec_ref(x_46); +x_49 = l_Lean_IR_Borrow_ownArgsUsingParams(x_43, x_47, x_2, x_48); +x_50 = lean_ctor_get(x_49, 1); +lean_inc(x_50); +lean_dec_ref(x_49); +x_51 = l_Lean_IR_Borrow_ownParamsUsingArgs(x_43, x_47, x_2, x_50); lean_dec_ref(x_2); -lean_dec(x_46); -lean_dec_ref(x_42); -return x_50; +lean_dec(x_47); +lean_dec_ref(x_43); +return x_51; } } case 9: { -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; -x_51 = lean_ctor_get(x_1, 3); -lean_inc_ref(x_51); +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; +x_52 = lean_ctor_get(x_1, 3); +lean_inc_ref(x_52); lean_dec_ref(x_1); -x_52 = lean_unsigned_to_nat(0u); -x_53 = lean_array_get_size(x_51); -x_54 = lean_box(0); -x_55 = lean_nat_dec_lt(x_52, x_53); -if (x_55 == 0) -{ -lean_object* x_56; -lean_dec_ref(x_51); +x_53 = lean_unsigned_to_nat(0u); +x_54 = lean_array_get_size(x_52); +x_55 = lean_box(0); +x_56 = lean_nat_dec_lt(x_53, x_54); +if (x_56 == 0) +{ +lean_object* x_57; +lean_dec_ref(x_52); lean_dec_ref(x_2); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_54); -lean_ctor_set(x_56, 1, x_3); -return x_56; +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_3); +return x_57; } else { -uint8_t x_57; -x_57 = lean_nat_dec_le(x_53, x_53); -if (x_57 == 0) +uint8_t x_58; +x_58 = lean_nat_dec_le(x_54, x_54); +if (x_58 == 0) { -if (x_55 == 0) +if (x_56 == 0) { -lean_object* x_58; -lean_dec_ref(x_51); +lean_object* x_59; +lean_dec_ref(x_52); lean_dec_ref(x_2); -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_54); -lean_ctor_set(x_58, 1, x_3); -return x_58; +x_59 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_59, 0, x_55); +lean_ctor_set(x_59, 1, x_3); +return x_59; } else { -size_t x_59; size_t x_60; lean_object* x_61; -x_59 = 0; -x_60 = lean_usize_of_nat(x_53); -x_61 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_Borrow_collectFnBody_spec__0(x_51, x_59, x_60, x_54, x_2, x_3); -lean_dec_ref(x_51); -return x_61; +size_t x_60; size_t x_61; lean_object* x_62; +x_60 = 0; +x_61 = lean_usize_of_nat(x_54); +x_62 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_Borrow_collectFnBody_spec__0(x_52, x_60, x_61, x_55, x_2, x_3); +lean_dec_ref(x_52); +return x_62; } } else { -size_t x_62; size_t x_63; lean_object* x_64; -x_62 = 0; -x_63 = lean_usize_of_nat(x_53); -x_64 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_Borrow_collectFnBody_spec__0(x_51, x_62, x_63, x_54, x_2, x_3); -lean_dec_ref(x_51); -return x_64; +size_t x_63; size_t x_64; lean_object* x_65; +x_63 = 0; +x_64 = lean_usize_of_nat(x_54); +x_65 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_Borrow_collectFnBody_spec__0(x_52, x_63, x_64, x_55, x_2, x_3); +lean_dec_ref(x_52); +return x_65; } } } default: { -uint8_t x_65; -x_65 = l_Lean_IR_FnBody_isTerminal(x_1); -if (x_65 == 0) +uint8_t x_66; +x_66 = l_Lean_IR_FnBody_isTerminal(x_1); +if (x_66 == 0) { -lean_object* x_66; -x_66 = l_Lean_IR_FnBody_body(x_1); +lean_object* x_67; +x_67 = l_Lean_IR_FnBody_body(x_1); lean_dec(x_1); -x_1 = x_66; +x_1 = x_67; goto _start; } else { -lean_object* x_68; lean_object* x_69; +lean_object* x_69; lean_object* x_70; lean_dec_ref(x_2); lean_dec(x_1); -x_68 = lean_box(0); -x_69 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_69, 0, x_68); -lean_ctor_set(x_69, 1, x_3); -return x_69; +x_69 = lean_box(0); +x_70 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_3); +return x_70; } } } diff --git a/stage0/stdlib/Lean/Compiler/IR/Boxing.c b/stage0/stdlib/Lean/Compiler/IR/Boxing.c index 6d7e3cce9f3b..ff308359cf7c 100644 --- a/stage0/stdlib/Lean/Compiler/IR/Boxing.c +++ b/stage0/stdlib/Lean/Compiler/IR/Boxing.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Compiler.IR.Boxing -// Imports: public import Lean.Runtime public import Lean.Compiler.ClosedTermCache public import Lean.Compiler.IR.CompilerM public import Lean.Compiler.IR.ElimDeadVars public import Lean.Compiler.IR.ToIRType public import Lean.Data.AssocList +// Imports: public import Lean.Runtime public import Lean.Compiler.ClosedTermCache public import Lean.Compiler.IR.CompilerM public import Lean.Compiler.IR.ElimDeadVars public import Lean.Compiler.IR.ToIRType public import Lean.Data.AssocList import Lean.Compiler.InitAttr #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -13,228 +13,314 @@ #ifdef __cplusplus extern "C" { #endif -static lean_object* l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__5; +uint8_t l_Lean_IR_IRType_isScalarOrStruct(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__5(lean_object*, lean_object*, lean_object*); +static lean_object* l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__1; +static lean_object* l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__3; lean_object* l_Lean_IR_LocalContext_addParams(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_mkBoxedVersionAux_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go_match__1_splitter___redArg(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__8_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_ExplicitBoxing_requiresBoxedVersion_spec__0___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_mkBoxedVersion(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_BasicAux_0__mapMonoMImp_go___at___00Lean_IR_ExplicitBoxing_BoxParams_boxParams_spec__0___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitVDeclExpr___lam__1(lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); -static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__14_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +LEAN_EXPORT uint8_t l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___lam__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_castArgsIfNeeded___lam__0(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__8_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_getEnv(lean_object*, lean_object*); lean_object* l_Lean_IR_Decl_elimDead(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_visitFnBody_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_requiresBoxedVersion___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__15_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; static lean_object* l_Lean_IR_ExplicitBoxing_requiresBoxedVersion___closed__0; uint8_t l_Lean_IR_IRType_isVoid(lean_object*); lean_object* l_StateT_instMonad___redArg___lam__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_withParams(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_IR_Decl_maxIndex(lean_object*); +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__28_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_getJPParams___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_castArgsIfNeeded___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__27_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +LEAN_EXPORT uint8_t l_Lean_IR_ExplicitBoxing_BoxParams_boxParams___lam__0(lean_object*, lean_object*, lean_object*); +uint8_t l_Array_isEmpty___redArg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_BoxParams_boxParams(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__3_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_; +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_insert___at___00__private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go_spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__16_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_castResultIfNeeded(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_ExplicitBoxing_mkBoxedVersionAux___closed__1; LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__3___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__11_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__10_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_getResultType(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_BasicAux_0__mapMonoMImp_go___at___00Lean_IR_ExplicitBoxing_BoxParams_boxParams_spec__0(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_ExplicitBoxing_transformBoxedPaps___closed__2; LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_boxArgsIfNeeded___lam__0___boxed(lean_object*); +static lean_object* l_Lean_IR_ExplicitBoxing_transformBoxedPaps___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_unboxResultIfNeeded___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux___closed__1; lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_getVarType(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitVDeclExpr___lam__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_ExplicitBoxing_BoxParams_boxParams_spec__2(lean_object*); extern lean_object* l_Lean_IR_instInhabitedDecl_default; uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_unboxResultIfNeeded___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__10_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; -uint8_t l_Lean_IR_instBEqIRType_beq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__19_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; uint8_t l_Lean_IR_CtorInfo_isScalar(lean_object*); +lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__0; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_visitFnBody_spec__0(size_t, size_t, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed___boxed(lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_IR_FnBody_isTerminal(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_withVDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_eqvTypes___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0(lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget_borrowed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__24_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__5_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_prepareBoxParams_spec__0(size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_withVDecl___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_getVarType___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__22_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_getLocalContext___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__4_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_boxArgsIfNeeded___lam__0(lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_BoxParams_isExpensiveScalar___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_mkBoxedVersionAux___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ExplicitBoxing_addBoxedVersions_spec__0(lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ExplicitBoxing_addBoxedVersions_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_ExplicitBoxing_mkCast___closed__1; +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_BoxParams_boxParams___lam__0___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_prepareBoxParams_spec__0___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_castArgsIfNeeded(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Id_instMonad___lam__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_BoxParams_boxParams_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__2_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_castArgIfNeeded___lam__0(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_IR_Decl_resultType(lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at___00Lean_IR_ExplicitBoxing_mkCast_spec__0___redArg(lean_object*, lean_object*); -static lean_object* l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__4; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_addBoxedVersions_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__7_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +lean_object* l_Nat_reprFast(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___at___00Lean_IR_ExplicitBoxing_mkBoxedVersionAux_spec__1___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +size_t lean_ptr_addr(lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; -static lean_object* l_Lean_IR_ExplicitBoxing_visitVDeclExpr___closed__0; +uint8_t l_Lean_isExport(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitCtorExpr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___lam__3___boxed(lean_object*); static lean_object* l_Lean_IR_ExplicitBoxing_addBoxedVersions___closed__0; +static lean_object* l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__4(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_IR_LocalContext_getType(lean_object*, lean_object*); extern lean_object* l_Lean_IR_instInhabitedParam_default; LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_withJDecl___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_IR_IRType_compatibleWith(lean_object*, lean_object*, uint8_t); lean_object* l_Lean_IR_Decl_params(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_BoxParams_boxParams_spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__14_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +lean_object* l_Array_empty(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_boxArgsIfNeeded(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitVDeclExpr___lam__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_ExplicitBoxing_transformBoxedPaps___closed__1; extern lean_object* l_Lean_closureMaxArgs; LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_mkBoxedVersion___boxed(lean_object*); lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*); static lean_object* l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__1; -static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__18_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr_spec__0(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_IR_LocalContext_getJPParams(lean_object*, lean_object*); static lean_object* l_Lean_IR_ExplicitBoxing_mkCast___closed__3; +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__18_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; extern lean_object* l_Lean_IR_instInhabitedFnBody_default__1; static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__2_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__5___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_isExpensiveConstantValueBoxing(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___at___00Lean_IR_ExplicitBoxing_mkBoxedVersionAux_spec__1___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Id_instMonad___lam__2___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__28_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; -static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__15_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; lean_object* l_Lean_IR_Decl_name(lean_object*); -static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__30_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___lam__1___boxed(lean_object*, lean_object*); lean_object* l_instInhabitedOfMonad___redArg(lean_object*, lean_object*); -static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__13_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; lean_object* lean_st_ref_get(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_explicitBoxing___redArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2____boxed(lean_object*); lean_object* l_instInhabitedForall___redArg___lam__0___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__29_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +uint8_t l_Lean_IR_IRType_isObj(lean_object*); +static lean_object* l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__2; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ExplicitBoxing_addBoxedVersions_spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*); +lean_object* l_Lean_IR_FnBody_body(lean_object*); +static lean_object* l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__4; +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__12_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__11_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__13_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__29_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_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_num___override(lean_object*, lean_object*); -lean_object* l_Lean_Name_append(lean_object*, lean_object*); lean_object* l_StateT_instMonad___redArg___lam__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___lam__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__5; +static lean_object* l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__1; +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_castResultIfNeeded___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_(); static lean_object* l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___at___00Lean_IR_ExplicitBoxing_mkBoxedVersionAux_spec__1___redArg___closed__0; +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_transformBoxedPaps___boxed(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_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__7_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_addBoxedVersions_spec__0(lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__17_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__21_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_insert___at___00__private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go_spec__1___redArg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_DTreeMap_Internal_Impl_contains___at___00__private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go_spec__0(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); lean_object* l_Id_instMonad___lam__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_isExpensiveConstantValueBoxing___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_mkCast(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_IR_LocalContext_addLocal(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_isExtern(lean_object*, lean_object*); lean_object* l_StateT_instMonad___redArg___lam__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__6; LEAN_EXPORT lean_object* l_Lean_IR_explicitBoxing___redArg___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_IR_ExplicitBoxing_mkBoxedName(lean_object*); LEAN_EXPORT uint8_t l_Lean_IR_ExplicitBoxing_requiresBoxedVersion(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_castVarIfNeeded(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__0; +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__25_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0(lean_object*, lean_object*, lean_object*); lean_object* l_StateT_bind(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_IR_ExplicitBoxing_BoxParams_isExpensiveScalar(lean_object*); static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__3_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; lean_object* l_Lean_IR_IRType_boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody(lean_object*, lean_object*, lean_object*); -static lean_object* l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__1; static lean_object* l_Lean_IR_ExplicitBoxing_mkCast___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_unboxResultIfNeeded(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___lam__3(lean_object*); +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__2_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_; lean_object* lean_array_get_borrowed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___at___00Lean_IR_ExplicitBoxing_mkBoxedVersionAux_spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux___closed__0; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_ExplicitBoxing_castArgsIfNeededAux_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_; static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__5_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; -static lean_object* l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__0; LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2____boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ExplicitBoxing_run_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_name_append_index_after(lean_object*, lean_object*); +static lean_object* l_Lean_IR_ExplicitBoxing_transformBoxedPaps___closed__3; uint8_t l_Lean_IR_FnBody_beq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_mkBoxedVersionAux(lean_object*, lean_object*); -static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__26_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(lean_object*); +LEAN_EXPORT uint8_t l_Std_DTreeMap_Internal_Impl_contains___at___00__private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go_spec__0___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_IR_FnBody_setBody(lean_object*, lean_object*); lean_object* l_Lean_IR_Decl_updateBody_x21(lean_object*, lean_object*); lean_object* l_Id_instMonad___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_explicitBoxing___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_IR_ExplicitBoxing_mkCast___closed__4; -static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__6_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +LEAN_EXPORT uint8_t l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___lam__0(lean_object*, uint8_t, uint8_t, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__5_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_N_mkFresh(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ExplicitBoxing_run_spec__0(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_getJPParams(lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__21_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_transformBoxedPaps_spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT uint8_t l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_ExplicitBoxing_requiresBoxedVersion_spec__0(lean_object*, size_t, size_t); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ExplicitBoxing_addBoxedVersions_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___closed__1; LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_castArgsIfNeeded___lam__0___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__4; LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___at___00Lean_IR_ExplicitBoxing_mkBoxedVersionAux_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_(); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__9_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_getLocalContext(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_withParams___redArg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_getResultType___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__24_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_boxArgsIfNeeded___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_IR_LocalContext_getValue(lean_object*, lean_object*); static lean_object* l_Lean_IR_ExplicitBoxing_mkCast___closed__2; +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; lean_object* lean_panic_fn(lean_object*, lean_object*); -static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__19_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__3_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_getEnv___boxed(lean_object*, lean_object*); lean_object* l_Array_append___redArg(lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__20_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +lean_object* lean_nat_mul(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_getDecl(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_IR_ExplicitBoxing_getDecl___closed__0; lean_object* l_Id_instMonad___lam__6(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__3; lean_object* l_Lean_IR_Decl_getInfo(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_withParams___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_addBoxedVersions(lean_object*, lean_object*); -static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__20_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +static lean_object* l_panic___at___00Lean_IR_ExplicitBoxing_BoxParams_boxParams_spec__2___closed__0; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_prepareBoxParams_spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_withJDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__25_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_IR_IRType_isStruct(lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_prepareBoxParams(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_ExplicitBoxing_transformBoxedPaps_spec__0(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_castArgIfNeeded(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); lean_object* l_mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__12_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitCtorExpr___lam__0___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_boxedConstDecl(lean_object*, lean_object*); +static lean_object* l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__3; static lean_object* l_Lean_IR_ExplicitBoxing_getJPParams___closed__0; lean_object* lean_array_uget(lean_object*, size_t); +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__22_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__4_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; size_t lean_array_size(lean_object*); -lean_object* l_Lean_Name_mkStr1(lean_object*); -static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__17_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__27_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +static lean_object* l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go_match__1_splitter(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__30_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_mkBoxedVersionAux_spec__0(size_t, size_t, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Lean_IR_ExplicitBoxing_eqvTypes(lean_object*, lean_object*); lean_object* l_Lean_IR_LocalContext_addJP(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__23_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitCtorExpr___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_string_append(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_transformBoxedPaps_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); -static lean_object* l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__2; lean_object* l_Id_instMonad___lam__0(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitVDeclExpr___lam__2(lean_object*, lean_object*, lean_object*); -static lean_object* l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__6; +lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_IR_reshape(lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); -static lean_object* l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__3; uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_ExplicitBoxing_castArgsIfNeededAux_spec__0(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); lean_object* l_StateT_map(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__4_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_; +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__0; static lean_object* l_Lean_IR_ExplicitBoxing_mkBoxedVersionAux___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_withParams___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__23_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +static lean_object* l_panic___at___00Lean_IR_ExplicitBoxing_transformBoxedPaps_spec__0___closed__0; +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitCtorExpr___lam__0(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_run(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at___00Lean_IR_ExplicitBoxing_mkCast_spec__0(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__26_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); -static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__16_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l_Lean_IR_explicitBoxing(lean_object*, lean_object*, lean_object*); lean_object* l_Id_instMonad___lam__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_prepareBoxParams_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2____boxed(lean_object*); uint8_t l_Lean_IR_IRType_isScalar(lean_object*); -static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__9_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; static lean_object* l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__2; -LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_contains___at___00__private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go_spec__0___redArg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___lam__1(uint8_t, lean_object*); lean_object* l_StateT_instMonad___redArg___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_transformBoxedPaps(lean_object*, lean_object*, lean_object*); lean_object* l_StateT_pure(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__6_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_contains___at___00__private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go_spec__0___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); static lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_N_mkFresh(lean_object* x_1) { _start: @@ -261,7 +347,7 @@ x_6 = lean_ctor_get_uint8(x_5, sizeof(void*)*2); x_7 = lean_ctor_get(x_5, 1); lean_inc(x_7); x_8 = 1; -x_16 = l_Lean_IR_IRType_isScalar(x_7); +x_16 = l_Lean_IR_IRType_isScalarOrStruct(x_7); lean_dec(x_7); if (x_16 == 0) { @@ -337,7 +423,7 @@ else { lean_object* x_16; uint8_t x_17; x_16 = l_Lean_IR_Decl_resultType(x_2); -x_17 = l_Lean_IR_IRType_isScalar(x_16); +x_17 = l_Lean_IR_IRType_isScalarOrStruct(x_16); lean_dec(x_16); if (x_17 == 0) { @@ -474,7 +560,7 @@ x_18 = lean_array_get_borrowed(x_1, x_2, x_17); x_19 = lean_ctor_get(x_18, 1); x_20 = lean_array_fget_borrowed(x_3, x_17); lean_dec(x_17); -x_21 = l_Lean_IR_IRType_isScalar(x_19); +x_21 = l_Lean_IR_IRType_isScalarOrStruct(x_19); if (x_21 == 0) { lean_object* x_22; lean_object* x_23; lean_object* x_24; @@ -574,7 +660,7 @@ x_55 = lean_array_get_borrowed(x_1, x_2, x_54); x_56 = lean_ctor_get(x_55, 1); x_57 = lean_array_fget_borrowed(x_3, x_54); lean_dec(x_54); -x_58 = l_Lean_IR_IRType_isScalar(x_56); +x_58 = l_Lean_IR_IRType_isScalarOrStruct(x_56); if (x_58 == 0) { lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; @@ -802,7 +888,7 @@ lean_ctor_set(x_32, 1, x_21); lean_ctor_set(x_32, 2, x_13); lean_ctor_set(x_32, 3, x_31); x_33 = lean_array_push(x_16, x_32); -x_34 = l_Lean_IR_IRType_isScalar(x_21); +x_34 = l_Lean_IR_IRType_isScalarOrStruct(x_21); if (x_34 == 0) { lean_object* x_35; lean_object* x_36; lean_object* x_37; @@ -935,7 +1021,7 @@ lean_ctor_set(x_75, 1, x_63); lean_ctor_set(x_75, 2, x_73); lean_ctor_set(x_75, 3, x_74); x_76 = lean_array_push(x_57, x_75); -x_77 = l_Lean_IR_IRType_isScalar(x_63); +x_77 = l_Lean_IR_IRType_isScalarOrStruct(x_63); if (x_77 == 0) { lean_object* x_78; lean_object* x_79; lean_object* x_80; @@ -1079,246 +1165,646 @@ lean_dec_ref(x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ExplicitBoxing_addBoxedVersions_spec__0(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +static lean_object* _init_l_panic___at___00Lean_IR_ExplicitBoxing_transformBoxedPaps_spec__0___closed__0() { _start: { -lean_object* x_6; uint8_t x_11; -x_11 = lean_usize_dec_eq(x_3, x_4); -if (x_11 == 0) -{ -lean_object* x_12; uint8_t x_13; -x_12 = lean_array_uget(x_2, x_3); -lean_inc_ref(x_1); -x_13 = l_Lean_IR_ExplicitBoxing_requiresBoxedVersion(x_1, x_12); -if (x_13 == 0) -{ -lean_dec(x_12); -x_6 = x_5; -goto block_10; +lean_object* x_1; +x_1 = l_Lean_IR_instInhabitedDecl_default; +return x_1; } -else +} +LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_ExplicitBoxing_transformBoxedPaps_spec__0(lean_object* x_1) { +_start: { -lean_object* x_14; lean_object* x_15; -x_14 = l_Lean_IR_ExplicitBoxing_mkBoxedVersion(x_12); -lean_dec(x_12); -x_15 = lean_array_push(x_5, x_14); -x_6 = x_15; -goto block_10; +lean_object* x_2; lean_object* x_3; +x_2 = l_panic___at___00Lean_IR_ExplicitBoxing_transformBoxedPaps_spec__0___closed__0; +x_3 = lean_panic_fn(x_2, x_1); +return x_3; } } -else +static lean_object* _init_l_Lean_IR_ExplicitBoxing_transformBoxedPaps___closed__2() { +_start: { -lean_dec_ref(x_1); -return x_5; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("value is none", 13, 13); +return x_1; } -block_10: -{ -size_t x_7; size_t x_8; -x_7 = 1; -x_8 = lean_usize_add(x_3, x_7); -x_3 = x_8; -x_5 = x_6; -goto _start; } +static lean_object* _init_l_Lean_IR_ExplicitBoxing_transformBoxedPaps___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Option.get!", 11, 11); +return x_1; } } -static lean_object* _init_l_Lean_IR_ExplicitBoxing_addBoxedVersions___closed__0() { +static lean_object* _init_l_Lean_IR_ExplicitBoxing_transformBoxedPaps___closed__0() { _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_object* x_1; +x_1 = lean_mk_string_unchecked("Init.Data.Option.BasicAux", 25, 25); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_addBoxedVersions(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_IR_ExplicitBoxing_transformBoxedPaps___closed__3() { _start: { -lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; -x_3 = lean_unsigned_to_nat(0u); -x_4 = l_Lean_IR_ExplicitBoxing_addBoxedVersions___closed__0; -x_5 = lean_array_get_size(x_2); -x_6 = lean_nat_dec_lt(x_3, x_5); -if (x_6 == 0) -{ -lean_object* x_7; -lean_dec_ref(x_1); -x_7 = l_Array_append___redArg(x_2, x_4); -return x_7; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_IR_ExplicitBoxing_transformBoxedPaps___closed__2; +x_2 = lean_unsigned_to_nat(14u); +x_3 = lean_unsigned_to_nat(22u); +x_4 = l_Lean_IR_ExplicitBoxing_transformBoxedPaps___closed__1; +x_5 = l_Lean_IR_ExplicitBoxing_transformBoxedPaps___closed__0; +x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); +return x_6; } -else -{ -uint8_t x_8; -x_8 = lean_nat_dec_le(x_5, x_5); -if (x_8 == 0) +} +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_transformBoxedPaps(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -if (x_6 == 0) +switch (lean_obj_tag(x_3)) { +case 0: { lean_object* x_9; -lean_dec_ref(x_1); -x_9 = l_Array_append___redArg(x_2, x_4); -return x_9; -} -else +x_9 = lean_ctor_get(x_3, 2); +lean_inc_ref(x_9); +if (lean_obj_tag(x_9) == 7) { -size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; -x_10 = 0; -x_11 = lean_usize_of_nat(x_5); -x_12 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ExplicitBoxing_addBoxedVersions_spec__0(x_1, x_2, x_10, x_11, x_4); -x_13 = l_Array_append___redArg(x_2, x_12); -lean_dec_ref(x_12); -return x_13; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_22; lean_object* x_26; +x_10 = lean_ctor_get(x_3, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_3, 1); +lean_inc(x_11); +x_12 = lean_ctor_get(x_3, 3); +lean_inc(x_12); +if (lean_is_exclusive(x_3)) { + lean_ctor_release(x_3, 0); + lean_ctor_release(x_3, 1); + lean_ctor_release(x_3, 2); + lean_ctor_release(x_3, 3); + x_13 = x_3; +} else { + lean_dec_ref(x_3); + x_13 = lean_box(0); } +x_14 = lean_ctor_get(x_9, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_9, 1); +lean_inc_ref(x_15); +if (lean_is_exclusive(x_9)) { + lean_ctor_release(x_9, 0); + lean_ctor_release(x_9, 1); + x_16 = x_9; +} else { + lean_dec_ref(x_9); + x_16 = lean_box(0); } -else +lean_inc(x_14); +lean_inc_ref(x_2); +x_26 = l_Lean_IR_findEnvDecl_x27(x_2, x_14, x_1); +if (lean_obj_tag(x_26) == 0) { -size_t x_14; size_t x_15; lean_object* x_16; lean_object* x_17; -x_14 = 0; -x_15 = lean_usize_of_nat(x_5); -x_16 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ExplicitBoxing_addBoxedVersions_spec__0(x_1, x_2, x_14, x_15, x_4); -x_17 = l_Array_append___redArg(x_2, x_16); -lean_dec_ref(x_16); -return x_17; -} -} +lean_object* x_27; lean_object* x_28; +x_27 = l_Lean_IR_ExplicitBoxing_transformBoxedPaps___closed__3; +x_28 = l_panic___at___00Lean_IR_ExplicitBoxing_transformBoxedPaps_spec__0(x_27); +x_22 = x_28; +goto block_25; } +else +{ +lean_object* x_29; +x_29 = lean_ctor_get(x_26, 0); +lean_inc(x_29); +lean_dec_ref(x_26); +x_22 = x_29; +goto block_25; } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ExplicitBoxing_addBoxedVersions_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +block_21: { -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ExplicitBoxing_addBoxedVersions_spec__0(x_1, x_2, x_6, x_7, x_5); -lean_dec_ref(x_2); -return x_8; +lean_object* x_18; lean_object* x_19; lean_object* x_20; +if (lean_is_scalar(x_16)) { + x_18 = lean_alloc_ctor(7, 2, 0); +} else { + x_18 = x_16; } +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_15); +x_19 = l_Lean_IR_ExplicitBoxing_transformBoxedPaps(x_1, x_2, x_12); +if (lean_is_scalar(x_13)) { + x_20 = lean_alloc_ctor(0, 4, 0); +} else { + x_20 = x_13; } -LEAN_EXPORT uint8_t l_Lean_IR_ExplicitBoxing_eqvTypes(lean_object* x_1, lean_object* x_2) { -_start: +lean_ctor_set(x_20, 0, x_10); +lean_ctor_set(x_20, 1, x_11); +lean_ctor_set(x_20, 2, x_18); +lean_ctor_set(x_20, 3, x_19); +return x_20; +} +block_25: { -uint8_t x_3; uint8_t x_7; uint8_t x_9; uint8_t x_10; -x_9 = l_Lean_IR_IRType_isScalar(x_1); -x_10 = l_Lean_IR_IRType_isScalar(x_2); -if (x_9 == 0) +uint8_t x_23; +lean_inc_ref(x_2); +x_23 = l_Lean_IR_ExplicitBoxing_requiresBoxedVersion(x_2, x_22); +lean_dec_ref(x_22); +if (x_23 == 0) { -if (x_10 == 0) +x_17 = x_14; +goto block_21; +} +else { -uint8_t x_11; -x_11 = 1; -x_3 = x_11; -goto block_6; +lean_object* x_24; +x_24 = l_Lean_IR_ExplicitBoxing_mkBoxedName(x_14); +x_17 = x_24; +goto block_21; +} +} } else { -x_7 = x_9; +lean_dec_ref(x_9); goto block_8; } } +case 1: +{ +uint8_t x_30; +x_30 = !lean_is_exclusive(x_3); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_31 = lean_ctor_get(x_3, 2); +x_32 = lean_ctor_get(x_3, 3); +lean_inc_ref(x_2); +x_33 = l_Lean_IR_ExplicitBoxing_transformBoxedPaps(x_1, x_2, x_31); +x_34 = l_Lean_IR_ExplicitBoxing_transformBoxedPaps(x_1, x_2, x_32); +lean_ctor_set(x_3, 3, x_34); +lean_ctor_set(x_3, 2, x_33); +return x_3; +} else { -x_7 = x_10; -goto block_8; +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_35 = lean_ctor_get(x_3, 0); +x_36 = lean_ctor_get(x_3, 1); +x_37 = lean_ctor_get(x_3, 2); +x_38 = lean_ctor_get(x_3, 3); +lean_inc(x_38); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_3); +lean_inc_ref(x_2); +x_39 = l_Lean_IR_ExplicitBoxing_transformBoxedPaps(x_1, x_2, x_37); +x_40 = l_Lean_IR_ExplicitBoxing_transformBoxedPaps(x_1, x_2, x_38); +x_41 = lean_alloc_ctor(1, 4, 0); +lean_ctor_set(x_41, 0, x_35); +lean_ctor_set(x_41, 1, x_36); +lean_ctor_set(x_41, 2, x_39); +lean_ctor_set(x_41, 3, x_40); +return x_41; +} } -block_6: +case 9: { -uint8_t x_4; -x_4 = l_Lean_IR_IRType_isScalar(x_1); -if (x_4 == 0) +uint8_t x_42; +x_42 = !lean_is_exclusive(x_3); +if (x_42 == 0) { -lean_dec(x_2); -lean_dec(x_1); +lean_object* x_43; size_t x_44; size_t x_45; lean_object* x_46; +x_43 = lean_ctor_get(x_3, 3); +x_44 = lean_array_size(x_43); +x_45 = 0; +x_46 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_transformBoxedPaps_spec__1(x_1, x_2, x_44, x_45, x_43); +lean_ctor_set(x_3, 3, x_46); return x_3; } else { -uint8_t x_5; -x_5 = l_Lean_IR_instBEqIRType_beq(x_1, x_2); -return x_5; +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; size_t x_51; size_t x_52; lean_object* x_53; lean_object* x_54; +x_47 = lean_ctor_get(x_3, 0); +x_48 = lean_ctor_get(x_3, 1); +x_49 = lean_ctor_get(x_3, 2); +x_50 = lean_ctor_get(x_3, 3); +lean_inc(x_50); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_dec(x_3); +x_51 = lean_array_size(x_50); +x_52 = 0; +x_53 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_transformBoxedPaps_spec__1(x_1, x_2, x_51, x_52, x_50); +x_54 = lean_alloc_ctor(9, 4, 0); +lean_ctor_set(x_54, 0, x_47); +lean_ctor_set(x_54, 1, x_48); +lean_ctor_set(x_54, 2, x_49); +lean_ctor_set(x_54, 3, x_53); +return x_54; +} +} +default: +{ +goto block_8; } } block_8: { -if (x_7 == 0) +uint8_t x_4; +x_4 = l_Lean_IR_FnBody_isTerminal(x_3); +if (x_4 == 0) { -lean_dec(x_2); -lean_dec(x_1); +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = l_Lean_IR_FnBody_body(x_3); +x_6 = l_Lean_IR_ExplicitBoxing_transformBoxedPaps(x_1, x_2, x_5); +x_7 = l_Lean_IR_FnBody_setBody(x_3, x_6); return x_7; } else { -x_3 = x_7; -goto block_6; +lean_dec_ref(x_2); +return x_3; } } } } -LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_eqvTypes___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_transformBoxedPaps_spec__1(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { _start: { -uint8_t x_3; lean_object* x_4; -x_3 = l_Lean_IR_ExplicitBoxing_eqvTypes(x_1, x_2); -x_4 = lean_box(x_3); -return x_4; -} +uint8_t x_6; +x_6 = lean_usize_dec_lt(x_4, x_3); +if (x_6 == 0) +{ +lean_dec_ref(x_2); +return x_5; } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(lean_object* x_1) { -_start: +else { -uint8_t x_2; -x_2 = !lean_is_exclusive(x_1); -if (x_2 == 0) +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_7 = lean_array_uget(x_5, x_4); +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_array_uset(x_5, x_4, x_8); +if (lean_obj_tag(x_7) == 0) { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_3 = lean_ctor_get(x_1, 0); -x_4 = lean_unsigned_to_nat(1u); -x_5 = lean_nat_add(x_3, x_4); -lean_ctor_set(x_1, 0, x_5); -x_6 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_6, 0, x_3); -lean_ctor_set(x_6, 1, x_1); -return x_6; +uint8_t x_16; +x_16 = !lean_is_exclusive(x_7); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_7, 1); +lean_inc_ref(x_2); +x_18 = l_Lean_IR_ExplicitBoxing_transformBoxedPaps(x_1, x_2, x_17); +lean_ctor_set(x_7, 1, x_18); +x_10 = x_7; +goto block_15; } else { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_7 = lean_ctor_get(x_1, 0); -x_8 = lean_ctor_get(x_1, 1); -x_9 = lean_ctor_get(x_1, 2); -x_10 = lean_ctor_get(x_1, 3); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_dec(x_1); -x_11 = lean_unsigned_to_nat(1u); -x_12 = lean_nat_add(x_7, x_11); -x_13 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_8); -lean_ctor_set(x_13, 2, x_9); -lean_ctor_set(x_13, 3, x_10); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_7); -lean_ctor_set(x_14, 1, x_13); -return x_14; +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_19 = lean_ctor_get(x_7, 0); +x_20 = lean_ctor_get(x_7, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_7); +lean_inc_ref(x_2); +x_21 = l_Lean_IR_ExplicitBoxing_transformBoxedPaps(x_1, x_2, x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_19); +lean_ctor_set(x_22, 1, x_21); +x_10 = x_22; +goto block_15; } } +else +{ +uint8_t x_23; +x_23 = !lean_is_exclusive(x_7); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_7, 0); +lean_inc_ref(x_2); +x_25 = l_Lean_IR_ExplicitBoxing_transformBoxedPaps(x_1, x_2, x_24); +lean_ctor_set(x_7, 0, x_25); +x_10 = x_7; +goto block_15; } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh(lean_object* x_1, lean_object* x_2) { -_start: +else { -lean_object* x_3; -x_3 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(x_2); -return x_3; +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_7, 0); +lean_inc(x_26); +lean_dec(x_7); +lean_inc_ref(x_2); +x_27 = l_Lean_IR_ExplicitBoxing_transformBoxedPaps(x_1, x_2, x_26); +x_28 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_28, 0, x_27); +x_10 = x_28; +goto block_15; } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___boxed(lean_object* x_1, lean_object* x_2) { -_start: +block_15: { -lean_object* x_3; +size_t x_11; size_t x_12; lean_object* x_13; +x_11 = 1; +x_12 = lean_usize_add(x_4, x_11); +x_13 = lean_array_uset(x_9, x_4, x_10); +x_4 = x_12; +x_5 = x_13; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_transformBoxedPaps_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_transformBoxedPaps_spec__1(x_1, x_2, x_6, x_7, x_5); +lean_dec_ref(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_transformBoxedPaps___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_IR_ExplicitBoxing_transformBoxedPaps(x_1, x_2, x_3); +lean_dec_ref(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ExplicitBoxing_addBoxedVersions_spec__1(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; uint8_t x_11; +x_11 = lean_usize_dec_eq(x_3, x_4); +if (x_11 == 0) +{ +lean_object* x_12; uint8_t x_13; +x_12 = lean_array_uget(x_2, x_3); +lean_inc_ref(x_1); +x_13 = l_Lean_IR_ExplicitBoxing_requiresBoxedVersion(x_1, x_12); +if (x_13 == 0) +{ +lean_dec(x_12); +x_6 = x_5; +goto block_10; +} +else +{ +lean_object* x_14; lean_object* x_15; +x_14 = l_Lean_IR_ExplicitBoxing_mkBoxedVersion(x_12); +lean_dec(x_12); +x_15 = lean_array_push(x_5, x_14); +x_6 = x_15; +goto block_10; +} +} +else +{ +lean_dec_ref(x_1); +return x_5; +} +block_10: +{ +size_t x_7; size_t x_8; +x_7 = 1; +x_8 = lean_usize_add(x_3, x_7); +x_3 = x_8; +x_5 = x_6; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_addBoxedVersions_spec__0(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_usize_dec_lt(x_4, x_3); +if (x_6 == 0) +{ +lean_dec_ref(x_2); +return x_5; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_7 = lean_array_uget(x_5, x_4); +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_array_uset(x_5, x_4, x_8); +if (lean_obj_tag(x_7) == 0) +{ +uint8_t x_16; +x_16 = !lean_is_exclusive(x_7); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_7, 3); +lean_inc_ref(x_2); +x_18 = l_Lean_IR_ExplicitBoxing_transformBoxedPaps(x_1, x_2, x_17); +lean_ctor_set(x_7, 3, x_18); +x_10 = x_7; +goto block_15; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_19 = lean_ctor_get(x_7, 0); +x_20 = lean_ctor_get(x_7, 1); +x_21 = lean_ctor_get(x_7, 2); +x_22 = lean_ctor_get(x_7, 3); +x_23 = lean_ctor_get(x_7, 4); +lean_inc(x_23); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_7); +lean_inc_ref(x_2); +x_24 = l_Lean_IR_ExplicitBoxing_transformBoxedPaps(x_1, x_2, x_22); +x_25 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_25, 0, x_19); +lean_ctor_set(x_25, 1, x_20); +lean_ctor_set(x_25, 2, x_21); +lean_ctor_set(x_25, 3, x_24); +lean_ctor_set(x_25, 4, x_23); +x_10 = x_25; +goto block_15; +} +} +else +{ +x_10 = x_7; +goto block_15; +} +block_15: +{ +size_t x_11; size_t x_12; lean_object* x_13; +x_11 = 1; +x_12 = lean_usize_add(x_4, x_11); +x_13 = lean_array_uset(x_9, x_4, x_10); +x_4 = x_12; +x_5 = x_13; +goto _start; +} +} +} +} +static lean_object* _init_l_Lean_IR_ExplicitBoxing_addBoxedVersions___closed__0() { +_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_IR_ExplicitBoxing_addBoxedVersions(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_9 = lean_unsigned_to_nat(0u); +x_10 = l_Lean_IR_ExplicitBoxing_addBoxedVersions___closed__0; +x_11 = lean_array_get_size(x_2); +x_12 = lean_nat_dec_lt(x_9, x_11); +if (x_12 == 0) +{ +x_3 = x_10; +goto block_8; +} +else +{ +uint8_t x_13; +x_13 = lean_nat_dec_le(x_11, x_11); +if (x_13 == 0) +{ +if (x_12 == 0) +{ +x_3 = x_10; +goto block_8; +} +else +{ +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = 0; +x_15 = lean_usize_of_nat(x_11); +lean_inc_ref(x_1); +x_16 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ExplicitBoxing_addBoxedVersions_spec__1(x_1, x_2, x_14, x_15, x_10); +x_3 = x_16; +goto block_8; +} +} +else +{ +size_t x_17; size_t x_18; lean_object* x_19; +x_17 = 0; +x_18 = lean_usize_of_nat(x_11); +lean_inc_ref(x_1); +x_19 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ExplicitBoxing_addBoxedVersions_spec__1(x_1, x_2, x_17, x_18, x_10); +x_3 = x_19; +goto block_8; +} +} +block_8: +{ +size_t x_4; size_t x_5; lean_object* x_6; lean_object* x_7; +x_4 = lean_array_size(x_2); +x_5 = 0; +lean_inc_ref(x_2); +x_6 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_addBoxedVersions_spec__0(x_2, x_1, x_4, x_5, x_2); +lean_dec_ref(x_2); +x_7 = l_Array_append___redArg(x_6, x_3); +lean_dec_ref(x_3); +return x_7; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ExplicitBoxing_addBoxedVersions_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ExplicitBoxing_addBoxedVersions_spec__1(x_1, x_2, x_6, x_7, x_5); +lean_dec_ref(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_addBoxedVersions_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_addBoxedVersions_spec__0(x_1, x_2, x_6, x_7, x_5); +lean_dec_ref(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(lean_object* x_1) { +_start: +{ +uint8_t x_2; +x_2 = !lean_is_exclusive(x_1); +if (x_2 == 0) +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_unsigned_to_nat(1u); +x_5 = lean_nat_add(x_3, x_4); +lean_ctor_set(x_1, 0, x_5); +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_3); +lean_ctor_set(x_6, 1, x_1); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_7 = lean_ctor_get(x_1, 0); +x_8 = lean_ctor_get(x_1, 1); +x_9 = lean_ctor_get(x_1, 2); +x_10 = lean_ctor_get(x_1, 3); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_dec(x_1); +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_add(x_7, x_11); +x_13 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_8); +lean_ctor_set(x_13, 2, x_9); +lean_ctor_set(x_13, 3, x_10); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_7); +lean_ctor_set(x_14, 1, x_13); +return x_14; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; x_3 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh(x_1, x_2); lean_dec_ref(x_1); return x_3; @@ -1540,14 +2026,6 @@ lean_dec(x_1); return x_4; } } -static lean_object* _init_l_Lean_IR_ExplicitBoxing_getDecl___closed__0() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_IR_instInhabitedDecl_default; -return x_1; -} -} LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_getDecl(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { @@ -1562,7 +2040,7 @@ lean_dec_ref(x_4); if (lean_obj_tag(x_6) == 0) { lean_object* x_7; lean_object* x_8; -x_7 = l_Lean_IR_ExplicitBoxing_getDecl___closed__0; +x_7 = l_panic___at___00Lean_IR_ExplicitBoxing_transformBoxedPaps_spec__0___closed__0; x_8 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_8, 0, x_7); lean_ctor_set(x_8, 1, x_3); @@ -2095,7 +2573,7 @@ static lean_object* _init_l_Lean_IR_ExplicitBoxing_mkCast___closed__2() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("_boxed_const", 12, 12); +x_1 = lean_mk_string_unchecked("_boxed_const_", 13, 13); return x_1; } } @@ -2103,15 +2581,6 @@ static lean_object* _init_l_Lean_IR_ExplicitBoxing_mkCast___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_IR_ExplicitBoxing_mkCast___closed__2; -x_2 = l_Lean_Name_mkStr1(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_IR_ExplicitBoxing_mkCast___closed__4() { -_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; @@ -2121,7 +2590,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_mkCast(lean_object* x_1, lean_ _start: { uint8_t x_6; -x_6 = l_Lean_IR_IRType_isScalar(x_3); +x_6 = l_Lean_IR_IRType_isScalarOrStruct(x_3); if (x_6 == 0) { lean_object* x_7; lean_object* x_8; @@ -2162,457 +2631,383 @@ return x_14; } else { -uint8_t x_15; +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_dec(x_1); -x_15 = !lean_is_exclusive(x_7); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_16 = lean_ctor_get(x_7, 1); -x_17 = lean_ctor_get(x_7, 0); -lean_dec(x_17); -x_18 = lean_ctor_get(x_8, 0); -lean_inc(x_18); +x_15 = lean_ctor_get(x_7, 1); +lean_inc(x_15); +if (lean_is_exclusive(x_7)) { + lean_ctor_release(x_7, 0); + lean_ctor_release(x_7, 1); + x_16 = x_7; +} else { + lean_dec_ref(x_7); + x_16 = lean_box(0); +} +x_17 = lean_ctor_get(x_8, 0); +lean_inc(x_17); lean_dec_ref(x_8); -x_19 = lean_ctor_get(x_16, 0); -x_20 = lean_ctor_get(x_16, 1); -x_21 = lean_ctor_get(x_16, 2); -x_22 = lean_ctor_get(x_16, 3); -x_23 = lean_unsigned_to_nat(1u); -x_24 = lean_unsigned_to_nat(2u); +if (lean_obj_tag(x_17) == 6) +{ +lean_object* x_84; lean_object* x_85; uint8_t x_86; lean_object* x_98; lean_object* x_99; uint8_t x_100; +x_84 = lean_ctor_get(x_17, 0); +x_85 = lean_ctor_get(x_17, 1); +x_98 = lean_array_get_size(x_85); +x_99 = lean_unsigned_to_nat(0u); +x_100 = lean_nat_dec_eq(x_98, x_99); +if (x_100 == 0) +{ +x_18 = x_4; +x_19 = x_15; +goto block_83; +} +else +{ +uint8_t x_101; +x_101 = l_Lean_IR_IRType_isObj(x_3); +if (x_101 == 0) +{ +x_86 = x_101; +goto block_97; +} +else +{ +uint8_t x_102; +x_102 = l_Lean_IR_IRType_isStruct(x_2); +x_86 = x_102; +goto block_97; +} +} +block_97: +{ +if (x_86 == 0) +{ +x_18 = x_4; +x_19 = x_15; +goto block_83; +} +else +{ +uint8_t x_87; +lean_inc(x_84); +lean_dec(x_16); +lean_dec_ref(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_87 = !lean_is_exclusive(x_17); +if (x_87 == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_88 = lean_ctor_get(x_17, 1); +lean_dec(x_88); +x_89 = lean_ctor_get(x_17, 0); +lean_dec(x_89); +x_90 = l_Lean_IR_ExplicitBoxing_mkBoxedName(x_84); +x_91 = l_Lean_IR_ExplicitBoxing_mkCast___closed__3; +lean_ctor_set(x_17, 1, x_91); +lean_ctor_set(x_17, 0, x_90); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_17); +lean_ctor_set(x_92, 1, x_15); +return x_92; +} +else +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; +lean_dec(x_17); +x_93 = l_Lean_IR_ExplicitBoxing_mkBoxedName(x_84); +x_94 = l_Lean_IR_ExplicitBoxing_mkCast___closed__3; +x_95 = lean_alloc_ctor(6, 2, 0); +lean_ctor_set(x_95, 0, x_93); +lean_ctor_set(x_95, 1, x_94); +x_96 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_96, 0, x_95); +lean_ctor_set(x_96, 1, x_15); +return x_96; +} +} +} +} +else +{ +x_18 = x_4; +x_19 = x_15; +goto block_83; +} +block_83: +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_20 = lean_ctor_get(x_19, 0); +x_21 = lean_ctor_get(x_19, 1); +x_22 = lean_ctor_get(x_19, 2); +x_23 = lean_ctor_get(x_19, 3); +x_24 = lean_unsigned_to_nat(1u); +x_25 = lean_unsigned_to_nat(2u); lean_inc(x_2); -x_25 = lean_alloc_ctor(9, 2, 0); -lean_ctor_set(x_25, 0, x_2); -lean_ctor_set(x_25, 1, x_23); -x_26 = l_Lean_IR_ExplicitBoxing_mkCast___closed__1; +x_26 = lean_alloc_ctor(9, 2, 0); +lean_ctor_set(x_26, 0, x_2); +lean_ctor_set(x_26, 1, x_24); +x_27 = l_Lean_IR_ExplicitBoxing_mkCast___closed__1; lean_inc(x_3); -x_27 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_27, 0, x_24); -lean_ctor_set(x_27, 1, x_3); -lean_ctor_set(x_27, 2, x_25); -lean_ctor_set(x_27, 3, x_26); x_28 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_28, 0, x_23); -lean_ctor_set(x_28, 1, x_2); -lean_ctor_set(x_28, 2, x_18); +lean_ctor_set(x_28, 0, x_25); +lean_ctor_set(x_28, 1, x_3); +lean_ctor_set(x_28, 2, x_26); lean_ctor_set(x_28, 3, x_27); -lean_inc(x_21); -lean_inc_ref(x_28); -x_29 = l_Lean_AssocList_find_x3f___at___00Lean_IR_ExplicitBoxing_mkCast_spec__0___redArg(x_28, x_21); -if (lean_obj_tag(x_29) == 0) +x_29 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_29, 0, x_24); +lean_ctor_set(x_29, 1, x_2); +lean_ctor_set(x_29, 2, x_17); +lean_ctor_set(x_29, 3, x_28); +lean_inc(x_22); +lean_inc_ref(x_29); +x_30 = l_Lean_AssocList_find_x3f___at___00Lean_IR_ExplicitBoxing_mkCast_spec__0___redArg(x_29, x_22); +if (lean_obj_tag(x_30) == 0) { -uint8_t x_30; +uint8_t x_31; +lean_inc(x_23); lean_inc(x_22); -lean_inc(x_21); -lean_inc_ref(x_20); -lean_inc(x_19); -x_30 = !lean_is_exclusive(x_16); -if (x_30 == 0) +lean_inc_ref(x_21); +lean_inc(x_20); +x_31 = !lean_is_exclusive(x_19); +if (x_31 == 0) { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; -x_31 = lean_ctor_get(x_16, 3); -lean_dec(x_31); -x_32 = lean_ctor_get(x_16, 2); +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_32 = lean_ctor_get(x_19, 3); lean_dec(x_32); -x_33 = lean_ctor_get(x_16, 1); +x_33 = lean_ctor_get(x_19, 2); lean_dec(x_33); -x_34 = lean_ctor_get(x_16, 0); +x_34 = lean_ctor_get(x_19, 1); lean_dec(x_34); -x_35 = !lean_is_exclusive(x_4); -if (x_35 == 0) +x_35 = lean_ctor_get(x_19, 0); +lean_dec(x_35); +x_36 = !lean_is_exclusive(x_18); +if (x_36 == 0) { -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_36 = lean_ctor_get(x_4, 0); -x_37 = lean_ctor_get(x_4, 4); -lean_dec(x_37); -x_38 = lean_ctor_get(x_4, 3); +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_37 = lean_ctor_get(x_18, 0); +x_38 = lean_ctor_get(x_18, 4); lean_dec(x_38); -x_39 = lean_ctor_get(x_4, 2); +x_39 = lean_ctor_get(x_18, 3); lean_dec(x_39); -x_40 = lean_ctor_get(x_4, 1); +x_40 = lean_ctor_get(x_18, 2); lean_dec(x_40); -x_41 = l_Lean_IR_ExplicitBoxing_mkCast___closed__3; -lean_inc(x_22); -x_42 = lean_name_append_index_after(x_41, x_22); -x_43 = l_Lean_Name_append(x_36, x_42); -x_44 = l_Lean_IR_ExplicitBoxing_mkCast___closed__4; -lean_inc(x_43); -x_45 = lean_alloc_ctor(6, 2, 0); -lean_ctor_set(x_45, 0, x_43); -lean_ctor_set(x_45, 1, x_44); -x_46 = lean_box(0); -lean_inc_ref(x_28); -lean_ctor_set(x_4, 4, x_46); -lean_ctor_set(x_4, 3, x_28); -lean_ctor_set(x_4, 2, x_3); -lean_ctor_set(x_4, 1, x_44); -lean_ctor_set(x_4, 0, x_43); -x_47 = lean_array_push(x_20, x_4); -lean_inc_ref(x_45); -x_48 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_48, 0, x_28); -lean_ctor_set(x_48, 1, x_45); -lean_ctor_set(x_48, 2, x_21); -x_49 = lean_nat_add(x_22, x_23); -lean_dec(x_22); -lean_ctor_set(x_16, 3, x_49); -lean_ctor_set(x_16, 2, x_48); -lean_ctor_set(x_16, 1, x_47); -lean_ctor_set(x_7, 0, x_45); -return x_7; +x_41 = lean_ctor_get(x_18, 1); +lean_dec(x_41); +x_42 = l_Lean_IR_ExplicitBoxing_mkCast___closed__2; +lean_inc(x_23); +x_43 = l_Nat_reprFast(x_23); +x_44 = lean_string_append(x_42, x_43); +lean_dec_ref(x_43); +x_45 = l_Lean_Name_str___override(x_37, x_44); +x_46 = l_Lean_IR_ExplicitBoxing_mkCast___closed__3; +lean_inc(x_45); +x_47 = lean_alloc_ctor(6, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +x_48 = lean_box(0); +lean_inc_ref(x_29); +lean_ctor_set(x_18, 4, x_48); +lean_ctor_set(x_18, 3, x_29); +lean_ctor_set(x_18, 2, x_3); +lean_ctor_set(x_18, 1, x_46); +lean_ctor_set(x_18, 0, x_45); +x_49 = lean_array_push(x_21, x_18); +lean_inc_ref(x_47); +x_50 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_50, 0, x_29); +lean_ctor_set(x_50, 1, x_47); +lean_ctor_set(x_50, 2, x_22); +x_51 = lean_nat_add(x_23, x_24); +lean_dec(x_23); +lean_ctor_set(x_19, 3, x_51); +lean_ctor_set(x_19, 2, x_50); +lean_ctor_set(x_19, 1, x_49); +if (lean_is_scalar(x_16)) { + x_52 = lean_alloc_ctor(0, 2, 0); +} else { + x_52 = x_16; +} +lean_ctor_set(x_52, 0, x_47); +lean_ctor_set(x_52, 1, x_19); +return x_52; } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_50 = lean_ctor_get(x_4, 0); -lean_inc(x_50); -lean_dec(x_4); -x_51 = l_Lean_IR_ExplicitBoxing_mkCast___closed__3; -lean_inc(x_22); -x_52 = lean_name_append_index_after(x_51, x_22); -x_53 = l_Lean_Name_append(x_50, x_52); -x_54 = l_Lean_IR_ExplicitBoxing_mkCast___closed__4; +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_53 = lean_ctor_get(x_18, 0); lean_inc(x_53); -x_55 = lean_alloc_ctor(6, 2, 0); -lean_ctor_set(x_55, 0, x_53); -lean_ctor_set(x_55, 1, x_54); -x_56 = lean_box(0); -lean_inc_ref(x_28); -x_57 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_57, 0, x_53); -lean_ctor_set(x_57, 1, x_54); -lean_ctor_set(x_57, 2, x_3); -lean_ctor_set(x_57, 3, x_28); -lean_ctor_set(x_57, 4, x_56); -x_58 = lean_array_push(x_20, x_57); -lean_inc_ref(x_55); -x_59 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_59, 0, x_28); -lean_ctor_set(x_59, 1, x_55); -lean_ctor_set(x_59, 2, x_21); -x_60 = lean_nat_add(x_22, x_23); -lean_dec(x_22); -lean_ctor_set(x_16, 3, x_60); -lean_ctor_set(x_16, 2, x_59); -lean_ctor_set(x_16, 1, x_58); -lean_ctor_set(x_7, 0, x_55); -return x_7; +lean_dec(x_18); +x_54 = l_Lean_IR_ExplicitBoxing_mkCast___closed__2; +lean_inc(x_23); +x_55 = l_Nat_reprFast(x_23); +x_56 = lean_string_append(x_54, x_55); +lean_dec_ref(x_55); +x_57 = l_Lean_Name_str___override(x_53, x_56); +x_58 = l_Lean_IR_ExplicitBoxing_mkCast___closed__3; +lean_inc(x_57); +x_59 = lean_alloc_ctor(6, 2, 0); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_58); +x_60 = lean_box(0); +lean_inc_ref(x_29); +x_61 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_61, 0, x_57); +lean_ctor_set(x_61, 1, x_58); +lean_ctor_set(x_61, 2, x_3); +lean_ctor_set(x_61, 3, x_29); +lean_ctor_set(x_61, 4, x_60); +x_62 = lean_array_push(x_21, x_61); +lean_inc_ref(x_59); +x_63 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_63, 0, x_29); +lean_ctor_set(x_63, 1, x_59); +lean_ctor_set(x_63, 2, x_22); +x_64 = lean_nat_add(x_23, x_24); +lean_dec(x_23); +lean_ctor_set(x_19, 3, x_64); +lean_ctor_set(x_19, 2, x_63); +lean_ctor_set(x_19, 1, x_62); +if (lean_is_scalar(x_16)) { + x_65 = lean_alloc_ctor(0, 2, 0); +} else { + x_65 = x_16; +} +lean_ctor_set(x_65, 0, x_59); +lean_ctor_set(x_65, 1, x_19); +return x_65; } } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; -lean_dec(x_16); -x_61 = lean_ctor_get(x_4, 0); -lean_inc(x_61); -if (lean_is_exclusive(x_4)) { - lean_ctor_release(x_4, 0); - lean_ctor_release(x_4, 1); - lean_ctor_release(x_4, 2); - lean_ctor_release(x_4, 3); - lean_ctor_release(x_4, 4); - x_62 = x_4; +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; +lean_dec(x_19); +x_66 = lean_ctor_get(x_18, 0); +lean_inc(x_66); +if (lean_is_exclusive(x_18)) { + lean_ctor_release(x_18, 0); + lean_ctor_release(x_18, 1); + lean_ctor_release(x_18, 2); + lean_ctor_release(x_18, 3); + lean_ctor_release(x_18, 4); + x_67 = x_18; } else { - lean_dec_ref(x_4); - x_62 = lean_box(0); + lean_dec_ref(x_18); + x_67 = lean_box(0); } -x_63 = l_Lean_IR_ExplicitBoxing_mkCast___closed__3; -lean_inc(x_22); -x_64 = lean_name_append_index_after(x_63, x_22); -x_65 = l_Lean_Name_append(x_61, x_64); -x_66 = l_Lean_IR_ExplicitBoxing_mkCast___closed__4; -lean_inc(x_65); -x_67 = lean_alloc_ctor(6, 2, 0); -lean_ctor_set(x_67, 0, x_65); -lean_ctor_set(x_67, 1, x_66); -x_68 = lean_box(0); -lean_inc_ref(x_28); -if (lean_is_scalar(x_62)) { - x_69 = lean_alloc_ctor(0, 5, 0); +x_68 = l_Lean_IR_ExplicitBoxing_mkCast___closed__2; +lean_inc(x_23); +x_69 = l_Nat_reprFast(x_23); +x_70 = lean_string_append(x_68, x_69); +lean_dec_ref(x_69); +x_71 = l_Lean_Name_str___override(x_66, x_70); +x_72 = l_Lean_IR_ExplicitBoxing_mkCast___closed__3; +lean_inc(x_71); +x_73 = lean_alloc_ctor(6, 2, 0); +lean_ctor_set(x_73, 0, x_71); +lean_ctor_set(x_73, 1, x_72); +x_74 = lean_box(0); +lean_inc_ref(x_29); +if (lean_is_scalar(x_67)) { + x_75 = lean_alloc_ctor(0, 5, 0); } else { - x_69 = x_62; + x_75 = x_67; } -lean_ctor_set(x_69, 0, x_65); -lean_ctor_set(x_69, 1, x_66); -lean_ctor_set(x_69, 2, x_3); -lean_ctor_set(x_69, 3, x_28); -lean_ctor_set(x_69, 4, x_68); -x_70 = lean_array_push(x_20, x_69); -lean_inc_ref(x_67); -x_71 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_71, 0, x_28); -lean_ctor_set(x_71, 1, x_67); -lean_ctor_set(x_71, 2, x_21); -x_72 = lean_nat_add(x_22, x_23); -lean_dec(x_22); -x_73 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_73, 0, x_19); -lean_ctor_set(x_73, 1, x_70); -lean_ctor_set(x_73, 2, x_71); -lean_ctor_set(x_73, 3, x_72); -lean_ctor_set(x_7, 1, x_73); -lean_ctor_set(x_7, 0, x_67); -return x_7; +lean_ctor_set(x_75, 0, x_71); +lean_ctor_set(x_75, 1, x_72); +lean_ctor_set(x_75, 2, x_3); +lean_ctor_set(x_75, 3, x_29); +lean_ctor_set(x_75, 4, x_74); +x_76 = lean_array_push(x_21, x_75); +lean_inc_ref(x_73); +x_77 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_77, 0, x_29); +lean_ctor_set(x_77, 1, x_73); +lean_ctor_set(x_77, 2, x_22); +x_78 = lean_nat_add(x_23, x_24); +lean_dec(x_23); +x_79 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_79, 0, x_20); +lean_ctor_set(x_79, 1, x_76); +lean_ctor_set(x_79, 2, x_77); +lean_ctor_set(x_79, 3, x_78); +if (lean_is_scalar(x_16)) { + x_80 = lean_alloc_ctor(0, 2, 0); +} else { + x_80 = x_16; +} +lean_ctor_set(x_80, 0, x_73); +lean_ctor_set(x_80, 1, x_79); +return x_80; } } else { -lean_object* x_74; -lean_dec_ref(x_28); -lean_dec_ref(x_4); -lean_dec(x_3); -x_74 = lean_ctor_get(x_29, 0); -lean_inc(x_74); +lean_object* x_81; lean_object* x_82; lean_dec_ref(x_29); -lean_ctor_set(x_7, 0, x_74); -return x_7; +lean_dec_ref(x_18); +lean_dec(x_3); +x_81 = lean_ctor_get(x_30, 0); +lean_inc(x_81); +lean_dec_ref(x_30); +if (lean_is_scalar(x_16)) { + x_82 = lean_alloc_ctor(0, 2, 0); +} else { + x_82 = x_16; } +lean_ctor_set(x_82, 0, x_81); +lean_ctor_set(x_82, 1, x_19); +return x_82; } -else -{ -lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_75 = lean_ctor_get(x_7, 1); -lean_inc(x_75); -lean_dec(x_7); -x_76 = lean_ctor_get(x_8, 0); -lean_inc(x_76); -lean_dec_ref(x_8); -x_77 = lean_ctor_get(x_75, 0); -x_78 = lean_ctor_get(x_75, 1); -x_79 = lean_ctor_get(x_75, 2); -x_80 = lean_ctor_get(x_75, 3); -x_81 = lean_unsigned_to_nat(1u); -x_82 = lean_unsigned_to_nat(2u); -lean_inc(x_2); -x_83 = lean_alloc_ctor(9, 2, 0); -lean_ctor_set(x_83, 0, x_2); -lean_ctor_set(x_83, 1, x_81); -x_84 = l_Lean_IR_ExplicitBoxing_mkCast___closed__1; -lean_inc(x_3); -x_85 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_85, 0, x_82); -lean_ctor_set(x_85, 1, x_3); -lean_ctor_set(x_85, 2, x_83); -lean_ctor_set(x_85, 3, x_84); -x_86 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_86, 0, x_81); -lean_ctor_set(x_86, 1, x_2); -lean_ctor_set(x_86, 2, x_76); -lean_ctor_set(x_86, 3, x_85); -lean_inc(x_79); -lean_inc_ref(x_86); -x_87 = l_Lean_AssocList_find_x3f___at___00Lean_IR_ExplicitBoxing_mkCast_spec__0___redArg(x_86, x_79); -if (lean_obj_tag(x_87) == 0) -{ -lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; -lean_inc(x_80); -lean_inc(x_79); -lean_inc_ref(x_78); -lean_inc(x_77); -if (lean_is_exclusive(x_75)) { - lean_ctor_release(x_75, 0); - lean_ctor_release(x_75, 1); - lean_ctor_release(x_75, 2); - lean_ctor_release(x_75, 3); - x_88 = x_75; -} else { - lean_dec_ref(x_75); - x_88 = lean_box(0); } -x_89 = lean_ctor_get(x_4, 0); -lean_inc(x_89); -if (lean_is_exclusive(x_4)) { - lean_ctor_release(x_4, 0); - lean_ctor_release(x_4, 1); - lean_ctor_release(x_4, 2); - lean_ctor_release(x_4, 3); - lean_ctor_release(x_4, 4); - x_90 = x_4; -} else { - lean_dec_ref(x_4); - x_90 = lean_box(0); } -x_91 = l_Lean_IR_ExplicitBoxing_mkCast___closed__3; -lean_inc(x_80); -x_92 = lean_name_append_index_after(x_91, x_80); -x_93 = l_Lean_Name_append(x_89, x_92); -x_94 = l_Lean_IR_ExplicitBoxing_mkCast___closed__4; -lean_inc(x_93); -x_95 = lean_alloc_ctor(6, 2, 0); -lean_ctor_set(x_95, 0, x_93); -lean_ctor_set(x_95, 1, x_94); -x_96 = lean_box(0); -lean_inc_ref(x_86); -if (lean_is_scalar(x_90)) { - x_97 = lean_alloc_ctor(0, 5, 0); -} else { - x_97 = x_90; -} -lean_ctor_set(x_97, 0, x_93); -lean_ctor_set(x_97, 1, x_94); -lean_ctor_set(x_97, 2, x_3); -lean_ctor_set(x_97, 3, x_86); -lean_ctor_set(x_97, 4, x_96); -x_98 = lean_array_push(x_78, x_97); -lean_inc_ref(x_95); -x_99 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_99, 0, x_86); -lean_ctor_set(x_99, 1, x_95); -lean_ctor_set(x_99, 2, x_79); -x_100 = lean_nat_add(x_80, x_81); -lean_dec(x_80); -if (lean_is_scalar(x_88)) { - x_101 = lean_alloc_ctor(0, 4, 0); -} else { - x_101 = x_88; -} -lean_ctor_set(x_101, 0, x_77); -lean_ctor_set(x_101, 1, x_98); -lean_ctor_set(x_101, 2, x_99); -lean_ctor_set(x_101, 3, x_100); -x_102 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_102, 0, x_95); -lean_ctor_set(x_102, 1, x_101); -return x_102; } else { -lean_object* x_103; lean_object* x_104; -lean_dec_ref(x_86); +uint8_t x_103; lean_dec_ref(x_4); lean_dec(x_3); -x_103 = lean_ctor_get(x_87, 0); -lean_inc(x_103); -lean_dec_ref(x_87); -x_104 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_104, 0, x_103); -lean_ctor_set(x_104, 1, x_75); -return x_104; -} -} -} +x_103 = l_Lean_IR_IRType_isStruct(x_2); +if (x_103 == 0) +{ +lean_object* x_104; lean_object* x_105; +lean_dec(x_2); +x_104 = lean_alloc_ctor(10, 1, 0); +lean_ctor_set(x_104, 0, x_1); +x_105 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_105, 0, x_104); +lean_ctor_set(x_105, 1, x_5); +return x_105; } else { -lean_object* x_105; lean_object* x_106; -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_105 = lean_alloc_ctor(10, 1, 0); -lean_ctor_set(x_105, 0, x_1); -x_106 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_106, 0, x_105); -lean_ctor_set(x_106, 1, x_5); -return x_106; +lean_object* x_106; lean_object* x_107; +x_106 = lean_alloc_ctor(9, 2, 0); +lean_ctor_set(x_106, 0, x_2); +lean_ctor_set(x_106, 1, x_1); +x_107 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_107, 0, x_106); +lean_ctor_set(x_107, 1, x_5); +return x_107; +} } } } LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_castVarIfNeeded(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; uint8_t x_10; x_6 = l_Lean_IR_ExplicitBoxing_getVarType(x_1, x_4, x_5); x_7 = lean_ctor_get(x_6, 0); lean_inc(x_7); x_8 = lean_ctor_get(x_6, 1); lean_inc(x_8); lean_dec_ref(x_6); -lean_inc(x_2); -lean_inc(x_7); -x_9 = l_Lean_IR_ExplicitBoxing_eqvTypes(x_7, x_2); -if (x_9 == 0) -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_10 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(x_8); -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); -lean_dec_ref(x_10); -lean_inc_ref(x_4); -lean_inc(x_2); -x_13 = l_Lean_IR_ExplicitBoxing_mkCast(x_1, x_7, x_2, x_4, x_12); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec_ref(x_13); -lean_inc(x_11); -x_16 = lean_apply_3(x_3, x_11, x_4, x_15); -x_17 = !lean_is_exclusive(x_16); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; -x_18 = lean_ctor_get(x_16, 0); -x_19 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_19, 0, x_11); -lean_ctor_set(x_19, 1, x_2); -lean_ctor_set(x_19, 2, x_14); -lean_ctor_set(x_19, 3, x_18); -lean_ctor_set(x_16, 0, x_19); -return x_16; -} -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_20 = lean_ctor_get(x_16, 0); -x_21 = lean_ctor_get(x_16, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_16); -x_22 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_22, 0, x_11); -lean_ctor_set(x_22, 1, x_2); -lean_ctor_set(x_22, 2, x_14); -lean_ctor_set(x_22, 3, x_20); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_21); -return x_23; -} -} -else -{ -lean_object* x_24; -lean_dec(x_7); -lean_dec(x_2); -x_24 = lean_apply_3(x_3, x_1, x_4, x_8); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_castArgIfNeeded___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; lean_object* x_6; -x_5 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_5, 0, x_2); -x_6 = lean_apply_3(x_1, x_5, x_3, x_4); -return x_6; -} -} -LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_castArgIfNeeded(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; -x_6 = lean_ctor_get(x_1, 0); -lean_inc(x_6); -lean_dec_ref(x_1); -x_7 = l_Lean_IR_ExplicitBoxing_getVarType(x_6, x_4, x_5); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_7, 1); -lean_inc(x_9); -lean_dec_ref(x_7); -lean_inc(x_2); -lean_inc(x_8); -x_10 = l_Lean_IR_ExplicitBoxing_eqvTypes(x_8, x_2); +x_9 = 0; +x_10 = l_Lean_IR_IRType_compatibleWith(x_7, x_2, x_9); if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_11 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(x_9); +x_11 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(x_8); x_12 = lean_ctor_get(x_11, 0); lean_inc(x_12); x_13 = lean_ctor_get(x_11, 1); @@ -2620,14 +3015,14 @@ lean_inc(x_13); lean_dec_ref(x_11); lean_inc_ref(x_4); lean_inc(x_2); -x_14 = l_Lean_IR_ExplicitBoxing_mkCast(x_6, x_8, x_2, x_4, x_13); +x_14 = l_Lean_IR_ExplicitBoxing_mkCast(x_1, x_7, x_2, x_4, x_13); x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); x_16 = lean_ctor_get(x_14, 1); lean_inc(x_16); lean_dec_ref(x_14); lean_inc(x_12); -x_17 = l_Lean_IR_ExplicitBoxing_castArgIfNeeded___lam__0(x_3, x_12, x_4, x_16); +x_17 = lean_apply_3(x_3, x_12, x_4, x_16); x_18 = !lean_is_exclusive(x_17); if (x_18 == 0) { @@ -2663,46 +3058,134 @@ return x_24; else { lean_object* x_25; -lean_dec(x_8); +lean_dec(x_7); lean_dec(x_2); -x_25 = l_Lean_IR_ExplicitBoxing_castArgIfNeeded___lam__0(x_3, x_6, x_4, x_9); +x_25 = lean_apply_3(x_3, x_1, x_4, x_8); return x_25; } } -else -{ -lean_object* x_26; -lean_dec(x_2); -x_26 = lean_apply_3(x_3, x_1, x_4, x_5); -return x_26; -} -} } -static lean_object* _init_l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux___closed__0() { +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_castArgIfNeeded___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_IR_ExplicitBoxing_mkCast___closed__4; -x_2 = lean_unsigned_to_nat(0u); -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; +lean_object* x_5; lean_object* x_6; +x_5 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_5, 0, x_2); +x_6 = lean_apply_3(x_1, x_5, x_3, x_4); +return x_6; } } -static lean_object* _init_l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux___closed__1() { +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_castArgIfNeeded(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux___closed__0; -x_2 = l_Lean_IR_ExplicitBoxing_mkCast___closed__4; -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_ExplicitBoxing_castArgsIfNeededAux_spec__0(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; uint8_t x_11; +x_6 = lean_ctor_get(x_1, 0); +lean_inc(x_6); +lean_dec_ref(x_1); +x_7 = l_Lean_IR_ExplicitBoxing_getVarType(x_6, x_4, x_5); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +lean_dec_ref(x_7); +x_10 = 0; +x_11 = l_Lean_IR_IRType_compatibleWith(x_8, x_2, x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_12 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(x_9); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec_ref(x_12); +lean_inc_ref(x_4); +lean_inc(x_2); +x_15 = l_Lean_IR_ExplicitBoxing_mkCast(x_6, x_8, x_2, x_4, x_14); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec_ref(x_15); +lean_inc(x_13); +x_18 = l_Lean_IR_ExplicitBoxing_castArgIfNeeded___lam__0(x_3, x_13, x_4, x_17); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_18, 0); +x_21 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_21, 0, x_13); +lean_ctor_set(x_21, 1, x_2); +lean_ctor_set(x_21, 2, x_16); +lean_ctor_set(x_21, 3, x_20); +lean_ctor_set(x_18, 0, x_21); +return x_18; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_22 = lean_ctor_get(x_18, 0); +x_23 = lean_ctor_get(x_18, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_18); +x_24 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_24, 0, x_13); +lean_ctor_set(x_24, 1, x_2); +lean_ctor_set(x_24, 2, x_16); +lean_ctor_set(x_24, 3, x_22); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_23); +return x_25; +} +} +else +{ +lean_object* x_26; +lean_dec(x_8); +lean_dec(x_2); +x_26 = l_Lean_IR_ExplicitBoxing_castArgIfNeeded___lam__0(x_3, x_6, x_4, x_9); +return x_26; +} +} +else +{ +lean_object* x_27; +lean_dec(x_2); +x_27 = lean_apply_3(x_3, x_1, x_4, x_5); +return x_27; +} +} +} +static lean_object* _init_l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux___closed__0() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_ExplicitBoxing_mkCast___closed__3; +x_2 = lean_unsigned_to_nat(0u); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux___closed__0; +x_2 = l_Lean_IR_ExplicitBoxing_mkCast___closed__3; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_ExplicitBoxing_castArgsIfNeededAux_spec__0(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { uint8_t x_8; @@ -2747,7 +3230,7 @@ if (lean_is_exclusive(x_10)) { x_27 = lean_array_uget(x_2, x_4); if (lean_obj_tag(x_27) == 0) { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; uint8_t x_34; x_28 = lean_ctor_get(x_27, 0); lean_inc(x_28); x_29 = l_Lean_IR_ExplicitBoxing_getVarType(x_28, x_6, x_7); @@ -2759,101 +3242,100 @@ lean_dec_ref(x_29); lean_inc_ref(x_1); lean_inc(x_13); x_32 = lean_apply_1(x_1, x_13); -lean_inc(x_32); -lean_inc(x_30); -x_33 = l_Lean_IR_ExplicitBoxing_eqvTypes(x_30, x_32); -if (x_33 == 0) -{ -uint8_t x_34; -x_34 = !lean_is_exclusive(x_27); +x_33 = 0; +x_34 = l_Lean_IR_IRType_compatibleWith(x_30, x_32, x_33); if (x_34 == 0) { -lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_35 = lean_ctor_get(x_27, 0); -lean_dec(x_35); -x_36 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(x_31); -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_36, 1); +uint8_t x_35; +x_35 = !lean_is_exclusive(x_27); +if (x_35 == 0) +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_36 = lean_ctor_get(x_27, 0); +lean_dec(x_36); +x_37 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(x_31); +x_38 = lean_ctor_get(x_37, 0); lean_inc(x_38); -lean_dec_ref(x_36); +x_39 = lean_ctor_get(x_37, 1); +lean_inc(x_39); +lean_dec_ref(x_37); lean_inc_ref(x_6); lean_inc(x_32); -x_39 = l_Lean_IR_ExplicitBoxing_mkCast(x_28, x_30, x_32, x_6, x_38); -x_40 = lean_ctor_get(x_39, 0); -lean_inc(x_40); -x_41 = lean_ctor_get(x_39, 1); +x_40 = l_Lean_IR_ExplicitBoxing_mkCast(x_28, x_30, x_32, x_6, x_39); +x_41 = lean_ctor_get(x_40, 0); lean_inc(x_41); -lean_dec_ref(x_39); -x_42 = lean_box(12); -lean_inc(x_37); -x_43 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_43, 0, x_37); -lean_ctor_set(x_43, 1, x_32); -lean_ctor_set(x_43, 2, x_40); -lean_ctor_set(x_43, 3, x_42); -lean_ctor_set(x_27, 0, x_37); -x_44 = lean_array_push(x_14, x_27); -x_45 = lean_array_push(x_11, x_43); -x_16 = x_45; -x_17 = x_44; -x_18 = x_41; +x_42 = lean_ctor_get(x_40, 1); +lean_inc(x_42); +lean_dec_ref(x_40); +x_43 = lean_box(12); +lean_inc(x_38); +x_44 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_44, 0, x_38); +lean_ctor_set(x_44, 1, x_32); +lean_ctor_set(x_44, 2, x_41); +lean_ctor_set(x_44, 3, x_43); +lean_ctor_set(x_27, 0, x_38); +x_45 = lean_array_push(x_14, x_27); +x_46 = lean_array_push(x_11, x_44); +x_16 = x_46; +x_17 = x_45; +x_18 = x_42; goto block_26; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_dec(x_27); -x_46 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(x_31); -x_47 = lean_ctor_get(x_46, 0); -lean_inc(x_47); -x_48 = lean_ctor_get(x_46, 1); +x_47 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(x_31); +x_48 = lean_ctor_get(x_47, 0); lean_inc(x_48); -lean_dec_ref(x_46); +x_49 = lean_ctor_get(x_47, 1); +lean_inc(x_49); +lean_dec_ref(x_47); lean_inc_ref(x_6); lean_inc(x_32); -x_49 = l_Lean_IR_ExplicitBoxing_mkCast(x_28, x_30, x_32, x_6, x_48); -x_50 = lean_ctor_get(x_49, 0); -lean_inc(x_50); -x_51 = lean_ctor_get(x_49, 1); +x_50 = l_Lean_IR_ExplicitBoxing_mkCast(x_28, x_30, x_32, x_6, x_49); +x_51 = lean_ctor_get(x_50, 0); lean_inc(x_51); -lean_dec_ref(x_49); -x_52 = lean_box(12); -lean_inc(x_47); -x_53 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_53, 0, x_47); -lean_ctor_set(x_53, 1, x_32); -lean_ctor_set(x_53, 2, x_50); -lean_ctor_set(x_53, 3, x_52); -x_54 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_54, 0, x_47); -x_55 = lean_array_push(x_14, x_54); -x_56 = lean_array_push(x_11, x_53); -x_16 = x_56; -x_17 = x_55; -x_18 = x_51; +x_52 = lean_ctor_get(x_50, 1); +lean_inc(x_52); +lean_dec_ref(x_50); +x_53 = lean_box(12); +lean_inc(x_48); +x_54 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_54, 0, x_48); +lean_ctor_set(x_54, 1, x_32); +lean_ctor_set(x_54, 2, x_51); +lean_ctor_set(x_54, 3, x_53); +x_55 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_55, 0, x_48); +x_56 = lean_array_push(x_14, x_55); +x_57 = lean_array_push(x_11, x_54); +x_16 = x_57; +x_17 = x_56; +x_18 = x_52; goto block_26; } } else { -lean_object* x_57; +lean_object* x_58; lean_dec(x_32); lean_dec(x_30); lean_dec(x_28); -x_57 = lean_array_push(x_14, x_27); +x_58 = lean_array_push(x_14, x_27); x_16 = x_11; -x_17 = x_57; +x_17 = x_58; x_18 = x_31; goto block_26; } } else { -lean_object* x_58; -x_58 = lean_array_push(x_14, x_27); +lean_object* x_59; +x_59 = lean_array_push(x_14, x_27); x_16 = x_11; -x_17 = x_58; +x_17 = x_59; x_18 = x_7; goto block_26; } @@ -3139,7 +3621,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_unboxResultIfNeeded___redArg(l _start: { uint8_t x_6; -x_6 = l_Lean_IR_IRType_isScalar(x_2); +x_6 = l_Lean_IR_IRType_isScalarOrStruct(x_2); if (x_6 == 0) { lean_object* x_7; lean_object* x_8; @@ -3229,745 +3711,167 @@ return x_7; LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_castResultIfNeeded(lean_object* x_1, lean_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_inc(x_2); -x_8 = l_Lean_IR_ExplicitBoxing_eqvTypes(x_2, x_4); -if (x_8 == 0) +uint8_t x_8; uint8_t x_9; +x_8 = 0; +x_9 = l_Lean_IR_IRType_compatibleWith(x_2, x_4, x_8); +if (x_9 == 0) { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_9 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(x_7); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_9, 1); +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_10 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(x_7); +x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); -lean_dec_ref(x_9); -x_12 = l_Lean_IR_IRType_boxed(x_2); -lean_inc(x_2); +x_12 = lean_ctor_get(x_10, 1); lean_inc(x_12); -lean_inc(x_10); -x_13 = l_Lean_IR_ExplicitBoxing_mkCast(x_10, x_12, x_2, x_6, x_11); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) +lean_dec_ref(x_10); +x_13 = l_Lean_IR_IRType_boxed(x_2); +lean_inc(x_2); +lean_inc(x_13); +lean_inc(x_11); +x_14 = l_Lean_IR_ExplicitBoxing_mkCast(x_11, x_13, x_2, x_6, x_12); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_16, 0, x_1); -lean_ctor_set(x_16, 1, x_2); -lean_ctor_set(x_16, 2, x_15); -lean_ctor_set(x_16, 3, x_5); +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_14, 0); x_17 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_17, 0, x_10); -lean_ctor_set(x_17, 1, x_12); -lean_ctor_set(x_17, 2, x_3); -lean_ctor_set(x_17, 3, x_16); -lean_ctor_set(x_13, 0, x_17); -return x_13; +lean_ctor_set(x_17, 0, x_1); +lean_ctor_set(x_17, 1, x_2); +lean_ctor_set(x_17, 2, x_16); +lean_ctor_set(x_17, 3, x_5); +x_18 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_18, 0, x_11); +lean_ctor_set(x_18, 1, x_13); +lean_ctor_set(x_18, 2, x_3); +lean_ctor_set(x_18, 3, x_17); +lean_ctor_set(x_14, 0, x_18); +return x_14; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_18 = lean_ctor_get(x_13, 0); -x_19 = lean_ctor_get(x_13, 1); +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_19 = lean_ctor_get(x_14, 0); +x_20 = lean_ctor_get(x_14, 1); +lean_inc(x_20); lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_13); -x_20 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_20, 0, x_1); -lean_ctor_set(x_20, 1, x_2); -lean_ctor_set(x_20, 2, x_18); -lean_ctor_set(x_20, 3, x_5); +lean_dec(x_14); x_21 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_21, 0, x_10); -lean_ctor_set(x_21, 1, x_12); -lean_ctor_set(x_21, 2, x_3); -lean_ctor_set(x_21, 3, x_20); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_19); -return x_22; +lean_ctor_set(x_21, 0, x_1); +lean_ctor_set(x_21, 1, x_2); +lean_ctor_set(x_21, 2, x_19); +lean_ctor_set(x_21, 3, x_5); +x_22 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_22, 0, x_11); +lean_ctor_set(x_22, 1, x_13); +lean_ctor_set(x_22, 2, x_3); +lean_ctor_set(x_22, 3, x_21); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_20); +return x_23; } } else { -lean_object* x_23; lean_object* x_24; +lean_object* x_24; lean_object* x_25; lean_dec_ref(x_6); -x_23 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_23, 0, x_1); -lean_ctor_set(x_23, 1, x_2); -lean_ctor_set(x_23, 2, x_3); -lean_ctor_set(x_23, 3, x_5); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_7); -return x_24; +x_24 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_24, 0, x_1); +lean_ctor_set(x_24, 1, x_2); +lean_ctor_set(x_24, 2, x_3); +lean_ctor_set(x_24, 3, x_5); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_7); +return x_25; } } } -LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitVDeclExpr___lam__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_castResultIfNeeded___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_4; lean_object* x_5; -x_4 = lean_array_get_borrowed(x_1, x_2, x_3); -x_5 = lean_ctor_get(x_4, 1); -lean_inc(x_5); -return x_5; +lean_object* x_8; +x_8 = l_Lean_IR_ExplicitBoxing_castResultIfNeeded(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_4); +return x_8; } } -static lean_object* _init_l_Lean_IR_ExplicitBoxing_visitVDeclExpr___closed__0() { +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitCtorExpr___lam__0(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = lean_box(0); +x_4 = lean_array_get_borrowed(x_3, x_1, x_2); +lean_inc(x_4); +return x_4; +} +} +static lean_object* _init_l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__0() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_IR_ExplicitBoxing_boxArgsIfNeeded___lam__0___boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Id_instMonad___lam__0), 4, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitVDeclExpr___lam__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__1() { _start: { -lean_object* x_4; -x_4 = l_Lean_IR_ExplicitBoxing_visitVDeclExpr___lam__2(x_1, x_2, x_3); -lean_dec(x_3); -lean_dec_ref(x_2); -return x_4; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Id_instMonad___lam__1___boxed), 4, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitVDeclExpr(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +static lean_object* _init_l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__2() { _start: { -switch (lean_obj_tag(x_3)) { -case 0: -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; uint8_t x_33; -x_7 = lean_ctor_get(x_3, 0); -lean_inc_ref(x_7); -x_8 = lean_ctor_get(x_3, 1); -lean_inc_ref(x_8); -if (lean_is_exclusive(x_3)) { - lean_ctor_release(x_3, 0); - lean_ctor_release(x_3, 1); - x_9 = x_3; -} else { - lean_dec_ref(x_3); - x_9 = lean_box(0); +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Id_instMonad___lam__2___boxed), 2, 0); +return x_1; } -x_10 = l_Lean_IR_ExplicitBoxing_visitVDeclExpr___closed__0; -x_33 = l_Lean_IR_CtorInfo_isScalar(x_7); -if (x_33 == 0) -{ -x_11 = x_33; -goto block_32; } -else +static lean_object* _init_l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__3() { +_start: { -uint8_t x_34; -x_34 = l_Lean_IR_IRType_isScalar(x_2); -x_11 = x_34; -goto block_32; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Id_instMonad___lam__3), 4, 0); +return x_1; +} } -block_32: +static lean_object* _init_l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__4() { +_start: { -if (x_11 == 0) +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Id_instMonad___lam__4___boxed), 4, 0); +return x_1; +} +} +static lean_object* _init_l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__5() { +_start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_12 = l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux(x_8, x_10, x_5, x_6); -lean_dec_ref(x_8); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec_ref(x_12); -x_15 = !lean_is_exclusive(x_13); -if (x_15 == 0) +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Id_instMonad___lam__5___boxed), 4, 0); +return x_1; +} +} +static lean_object* _init_l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__6() { +_start: { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_16 = lean_ctor_get(x_13, 0); -x_17 = lean_ctor_get(x_13, 1); -if (lean_is_scalar(x_9)) { - x_18 = lean_alloc_ctor(0, 2, 0); -} else { - x_18 = x_9; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Id_instMonad___lam__6), 4, 0); +return x_1; } -lean_ctor_set(x_18, 0, x_7); -lean_ctor_set(x_18, 1, x_16); -x_19 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_19, 0, x_1); -lean_ctor_set(x_19, 1, x_2); -lean_ctor_set(x_19, 2, x_18); -lean_ctor_set(x_19, 3, x_4); -x_20 = l_Lean_IR_reshape(x_17, x_19); -lean_ctor_set(x_13, 1, x_14); -lean_ctor_set(x_13, 0, x_20); -return x_13; } -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_21 = lean_ctor_get(x_13, 0); -x_22 = lean_ctor_get(x_13, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_13); -if (lean_is_scalar(x_9)) { - x_23 = lean_alloc_ctor(0, 2, 0); -} else { - x_23 = x_9; -} -lean_ctor_set(x_23, 0, x_7); -lean_ctor_set(x_23, 1, x_21); -x_24 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_24, 0, x_1); -lean_ctor_set(x_24, 1, x_2); -lean_ctor_set(x_24, 2, x_23); -lean_ctor_set(x_24, 3, x_4); -x_25 = l_Lean_IR_reshape(x_22, x_24); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_14); -return x_26; -} -} -else -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -lean_dec(x_9); -lean_dec_ref(x_8); -lean_dec_ref(x_5); -x_27 = lean_ctor_get(x_7, 1); -lean_inc(x_27); -lean_dec_ref(x_7); -x_28 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_28, 0, x_27); -x_29 = lean_alloc_ctor(11, 1, 0); -lean_ctor_set(x_29, 0, x_28); -x_30 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_30, 0, x_1); -lean_ctor_set(x_30, 1, x_2); -lean_ctor_set(x_30, 2, x_29); -lean_ctor_set(x_30, 3, x_4); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_6); -return x_31; -} -} -} -case 2: -{ -uint8_t x_35; -x_35 = !lean_is_exclusive(x_3); -if (x_35 == 0) -{ -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; -x_36 = lean_ctor_get(x_3, 2); -x_37 = l_Lean_IR_ExplicitBoxing_visitVDeclExpr___closed__0; -x_38 = l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux(x_36, x_37, x_5, x_6); -lean_dec_ref(x_36); -x_39 = lean_ctor_get(x_38, 0); -lean_inc(x_39); -x_40 = lean_ctor_get(x_38, 1); -lean_inc(x_40); -lean_dec_ref(x_38); -x_41 = !lean_is_exclusive(x_39); -if (x_41 == 0) -{ -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_42 = lean_ctor_get(x_39, 0); -x_43 = lean_ctor_get(x_39, 1); -lean_ctor_set(x_3, 2, x_42); -x_44 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_44, 0, x_1); -lean_ctor_set(x_44, 1, x_2); -lean_ctor_set(x_44, 2, x_3); -lean_ctor_set(x_44, 3, x_4); -x_45 = l_Lean_IR_reshape(x_43, x_44); -lean_ctor_set(x_39, 1, x_40); -lean_ctor_set(x_39, 0, x_45); -return x_39; -} -else -{ -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_46 = lean_ctor_get(x_39, 0); -x_47 = lean_ctor_get(x_39, 1); -lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_39); -lean_ctor_set(x_3, 2, x_46); -x_48 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_48, 0, x_1); -lean_ctor_set(x_48, 1, x_2); -lean_ctor_set(x_48, 2, x_3); -lean_ctor_set(x_48, 3, x_4); -x_49 = l_Lean_IR_reshape(x_47, x_48); -x_50 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_50, 0, x_49); -lean_ctor_set(x_50, 1, x_40); -return x_50; -} -} -else -{ -lean_object* x_51; lean_object* x_52; uint8_t x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_51 = lean_ctor_get(x_3, 0); -x_52 = lean_ctor_get(x_3, 1); -x_53 = lean_ctor_get_uint8(x_3, sizeof(void*)*3); -x_54 = lean_ctor_get(x_3, 2); -lean_inc(x_54); -lean_inc(x_52); -lean_inc(x_51); -lean_dec(x_3); -x_55 = l_Lean_IR_ExplicitBoxing_visitVDeclExpr___closed__0; -x_56 = l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux(x_54, x_55, x_5, x_6); -lean_dec_ref(x_54); -x_57 = lean_ctor_get(x_56, 0); -lean_inc(x_57); -x_58 = lean_ctor_get(x_56, 1); -lean_inc(x_58); -lean_dec_ref(x_56); -x_59 = lean_ctor_get(x_57, 0); -lean_inc(x_59); -x_60 = lean_ctor_get(x_57, 1); -lean_inc(x_60); -if (lean_is_exclusive(x_57)) { - lean_ctor_release(x_57, 0); - lean_ctor_release(x_57, 1); - x_61 = x_57; -} else { - lean_dec_ref(x_57); - x_61 = lean_box(0); -} -x_62 = lean_alloc_ctor(2, 3, 1); -lean_ctor_set(x_62, 0, x_51); -lean_ctor_set(x_62, 1, x_52); -lean_ctor_set(x_62, 2, x_59); -lean_ctor_set_uint8(x_62, sizeof(void*)*3, x_53); -x_63 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_63, 0, x_1); -lean_ctor_set(x_63, 1, x_2); -lean_ctor_set(x_63, 2, x_62); -lean_ctor_set(x_63, 3, x_4); -x_64 = l_Lean_IR_reshape(x_60, x_63); -if (lean_is_scalar(x_61)) { - x_65 = lean_alloc_ctor(0, 2, 0); -} else { - x_65 = x_61; -} -lean_ctor_set(x_65, 0, x_64); -lean_ctor_set(x_65, 1, x_58); -return x_65; -} -} -case 6: -{ -uint8_t x_66; -x_66 = !lean_is_exclusive(x_3); -if (x_66 == 0) -{ -lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; uint8_t x_82; -x_67 = lean_ctor_get(x_3, 0); -x_68 = lean_ctor_get(x_3, 1); -lean_inc_ref(x_5); -lean_inc(x_67); -x_69 = l_Lean_IR_ExplicitBoxing_getDecl(x_67, x_5, x_6); -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_ref(x_69); -x_72 = l_Lean_IR_Decl_params(x_70); -x_73 = l_Lean_IR_ExplicitBoxing_mkBoxedVersionAux___closed__0; -x_74 = lean_alloc_closure((void*)(l_Lean_IR_ExplicitBoxing_visitVDeclExpr___lam__2___boxed), 3, 2); -lean_closure_set(x_74, 0, x_73); -lean_closure_set(x_74, 1, x_72); -lean_inc_ref(x_5); -x_75 = l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux(x_68, x_74, x_5, x_71); -lean_dec_ref(x_68); -x_76 = lean_ctor_get(x_75, 0); -lean_inc(x_76); -x_77 = lean_ctor_get(x_75, 1); -lean_inc(x_77); -lean_dec_ref(x_75); -x_78 = lean_ctor_get(x_76, 0); -lean_inc(x_78); -x_79 = lean_ctor_get(x_76, 1); -lean_inc(x_79); -lean_dec(x_76); -lean_ctor_set(x_3, 1, x_78); -x_80 = l_Lean_IR_Decl_resultType(x_70); -lean_dec(x_70); -x_81 = l_Lean_IR_ExplicitBoxing_castResultIfNeeded(x_1, x_2, x_3, x_80, x_4, x_5, x_77); -x_82 = !lean_is_exclusive(x_81); -if (x_82 == 0) -{ -lean_object* x_83; lean_object* x_84; -x_83 = lean_ctor_get(x_81, 0); -x_84 = l_Lean_IR_reshape(x_79, x_83); -lean_ctor_set(x_81, 0, x_84); -return x_81; -} -else -{ -lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_85 = lean_ctor_get(x_81, 0); -x_86 = lean_ctor_get(x_81, 1); -lean_inc(x_86); -lean_inc(x_85); -lean_dec(x_81); -x_87 = l_Lean_IR_reshape(x_79, x_85); -x_88 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_88, 0, x_87); -lean_ctor_set(x_88, 1, x_86); -return x_88; -} -} -else -{ -lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; -x_89 = lean_ctor_get(x_3, 0); -x_90 = lean_ctor_get(x_3, 1); -lean_inc(x_90); -lean_inc(x_89); -lean_dec(x_3); -lean_inc_ref(x_5); -lean_inc(x_89); -x_91 = l_Lean_IR_ExplicitBoxing_getDecl(x_89, x_5, x_6); -x_92 = lean_ctor_get(x_91, 0); -lean_inc(x_92); -x_93 = lean_ctor_get(x_91, 1); -lean_inc(x_93); -lean_dec_ref(x_91); -x_94 = l_Lean_IR_Decl_params(x_92); -x_95 = l_Lean_IR_ExplicitBoxing_mkBoxedVersionAux___closed__0; -x_96 = lean_alloc_closure((void*)(l_Lean_IR_ExplicitBoxing_visitVDeclExpr___lam__2___boxed), 3, 2); -lean_closure_set(x_96, 0, x_95); -lean_closure_set(x_96, 1, x_94); -lean_inc_ref(x_5); -x_97 = l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux(x_90, x_96, x_5, x_93); -lean_dec_ref(x_90); -x_98 = lean_ctor_get(x_97, 0); -lean_inc(x_98); -x_99 = lean_ctor_get(x_97, 1); -lean_inc(x_99); -lean_dec_ref(x_97); -x_100 = lean_ctor_get(x_98, 0); -lean_inc(x_100); -x_101 = lean_ctor_get(x_98, 1); -lean_inc(x_101); -lean_dec(x_98); -x_102 = lean_alloc_ctor(6, 2, 0); -lean_ctor_set(x_102, 0, x_89); -lean_ctor_set(x_102, 1, x_100); -x_103 = l_Lean_IR_Decl_resultType(x_92); -lean_dec(x_92); -x_104 = l_Lean_IR_ExplicitBoxing_castResultIfNeeded(x_1, x_2, x_102, x_103, x_4, x_5, x_99); -x_105 = lean_ctor_get(x_104, 0); -lean_inc(x_105); -x_106 = lean_ctor_get(x_104, 1); -lean_inc(x_106); -if (lean_is_exclusive(x_104)) { - lean_ctor_release(x_104, 0); - lean_ctor_release(x_104, 1); - x_107 = x_104; -} else { - lean_dec_ref(x_104); - x_107 = lean_box(0); -} -x_108 = l_Lean_IR_reshape(x_101, x_105); -if (lean_is_scalar(x_107)) { - x_109 = lean_alloc_ctor(0, 2, 0); -} else { - x_109 = x_107; -} -lean_ctor_set(x_109, 0, x_108); -lean_ctor_set(x_109, 1, x_106); -return x_109; -} -} -case 7: -{ -lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; uint8_t x_137; -x_110 = lean_ctor_get(x_3, 0); -lean_inc(x_110); -x_111 = lean_ctor_get(x_3, 1); -lean_inc_ref(x_111); -if (lean_is_exclusive(x_3)) { - lean_ctor_release(x_3, 0); - lean_ctor_release(x_3, 1); - x_112 = x_3; -} else { - lean_dec_ref(x_3); - x_112 = lean_box(0); -} -x_113 = l_Lean_IR_ExplicitBoxing_getEnv(x_5, x_6); -x_114 = lean_ctor_get(x_113, 0); -lean_inc(x_114); -x_115 = lean_ctor_get(x_113, 1); -lean_inc(x_115); -lean_dec_ref(x_113); -lean_inc_ref(x_5); -lean_inc(x_110); -x_116 = l_Lean_IR_ExplicitBoxing_getDecl(x_110, x_5, x_115); -x_117 = lean_ctor_get(x_116, 0); -lean_inc(x_117); -x_118 = lean_ctor_get(x_116, 1); -lean_inc(x_118); -lean_dec_ref(x_116); -x_119 = l_Lean_IR_ExplicitBoxing_visitVDeclExpr___closed__0; -x_137 = l_Lean_IR_ExplicitBoxing_requiresBoxedVersion(x_114, x_117); -lean_dec(x_117); -if (x_137 == 0) -{ -x_120 = x_110; -goto block_136; -} -else -{ -lean_object* x_138; -x_138 = l_Lean_IR_ExplicitBoxing_mkBoxedName(x_110); -x_120 = x_138; -goto block_136; -} -block_136: -{ -lean_object* x_121; lean_object* x_122; lean_object* x_123; uint8_t x_124; -x_121 = l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux(x_111, x_119, x_5, x_118); -lean_dec_ref(x_111); -x_122 = lean_ctor_get(x_121, 0); -lean_inc(x_122); -x_123 = lean_ctor_get(x_121, 1); -lean_inc(x_123); -lean_dec_ref(x_121); -x_124 = !lean_is_exclusive(x_122); -if (x_124 == 0) -{ -lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; -x_125 = lean_ctor_get(x_122, 0); -x_126 = lean_ctor_get(x_122, 1); -if (lean_is_scalar(x_112)) { - x_127 = lean_alloc_ctor(7, 2, 0); -} else { - x_127 = x_112; -} -lean_ctor_set(x_127, 0, x_120); -lean_ctor_set(x_127, 1, x_125); -x_128 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_128, 0, x_1); -lean_ctor_set(x_128, 1, x_2); -lean_ctor_set(x_128, 2, x_127); -lean_ctor_set(x_128, 3, x_4); -x_129 = l_Lean_IR_reshape(x_126, x_128); -lean_ctor_set(x_122, 1, x_123); -lean_ctor_set(x_122, 0, x_129); -return x_122; -} -else -{ -lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; -x_130 = lean_ctor_get(x_122, 0); -x_131 = lean_ctor_get(x_122, 1); -lean_inc(x_131); -lean_inc(x_130); -lean_dec(x_122); -if (lean_is_scalar(x_112)) { - x_132 = lean_alloc_ctor(7, 2, 0); -} else { - x_132 = x_112; -} -lean_ctor_set(x_132, 0, x_120); -lean_ctor_set(x_132, 1, x_130); -x_133 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_133, 0, x_1); -lean_ctor_set(x_133, 1, x_2); -lean_ctor_set(x_133, 2, x_132); -lean_ctor_set(x_133, 3, x_4); -x_134 = l_Lean_IR_reshape(x_131, x_133); -x_135 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_135, 0, x_134); -lean_ctor_set(x_135, 1, x_123); -return x_135; -} -} -} -case 8: -{ -uint8_t x_139; -x_139 = !lean_is_exclusive(x_3); -if (x_139 == 0) -{ -lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; uint8_t x_148; -x_140 = lean_ctor_get(x_3, 1); -x_141 = l_Lean_IR_ExplicitBoxing_visitVDeclExpr___closed__0; -x_142 = l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux(x_140, x_141, x_5, x_6); -lean_dec_ref(x_140); -x_143 = lean_ctor_get(x_142, 0); -lean_inc(x_143); -x_144 = lean_ctor_get(x_142, 1); -lean_inc(x_144); -lean_dec_ref(x_142); -x_145 = lean_ctor_get(x_143, 0); -lean_inc(x_145); -x_146 = lean_ctor_get(x_143, 1); -lean_inc(x_146); -lean_dec(x_143); -lean_ctor_set(x_3, 1, x_145); -x_147 = l_Lean_IR_ExplicitBoxing_unboxResultIfNeeded___redArg(x_1, x_2, x_3, x_4, x_144); -x_148 = !lean_is_exclusive(x_147); -if (x_148 == 0) -{ -lean_object* x_149; lean_object* x_150; -x_149 = lean_ctor_get(x_147, 0); -x_150 = l_Lean_IR_reshape(x_146, x_149); -lean_ctor_set(x_147, 0, x_150); -return x_147; -} -else -{ -lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; -x_151 = lean_ctor_get(x_147, 0); -x_152 = lean_ctor_get(x_147, 1); -lean_inc(x_152); -lean_inc(x_151); -lean_dec(x_147); -x_153 = l_Lean_IR_reshape(x_146, x_151); -x_154 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_154, 0, x_153); -lean_ctor_set(x_154, 1, x_152); -return x_154; -} -} -else -{ -lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; -x_155 = lean_ctor_get(x_3, 0); -x_156 = lean_ctor_get(x_3, 1); -lean_inc(x_156); -lean_inc(x_155); -lean_dec(x_3); -x_157 = l_Lean_IR_ExplicitBoxing_visitVDeclExpr___closed__0; -x_158 = l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux(x_156, x_157, x_5, x_6); -lean_dec_ref(x_156); -x_159 = lean_ctor_get(x_158, 0); -lean_inc(x_159); -x_160 = lean_ctor_get(x_158, 1); -lean_inc(x_160); -lean_dec_ref(x_158); -x_161 = lean_ctor_get(x_159, 0); -lean_inc(x_161); -x_162 = lean_ctor_get(x_159, 1); -lean_inc(x_162); -lean_dec(x_159); -x_163 = lean_alloc_ctor(8, 2, 0); -lean_ctor_set(x_163, 0, x_155); -lean_ctor_set(x_163, 1, x_161); -x_164 = l_Lean_IR_ExplicitBoxing_unboxResultIfNeeded___redArg(x_1, x_2, x_163, x_4, x_160); -x_165 = lean_ctor_get(x_164, 0); -lean_inc(x_165); -x_166 = lean_ctor_get(x_164, 1); -lean_inc(x_166); -if (lean_is_exclusive(x_164)) { - lean_ctor_release(x_164, 0); - lean_ctor_release(x_164, 1); - x_167 = x_164; -} else { - lean_dec_ref(x_164); - x_167 = lean_box(0); -} -x_168 = l_Lean_IR_reshape(x_162, x_165); -if (lean_is_scalar(x_167)) { - x_169 = lean_alloc_ctor(0, 2, 0); -} else { - x_169 = x_167; -} -lean_ctor_set(x_169, 0, x_168); -lean_ctor_set(x_169, 1, x_166); -return x_169; -} -} -default: -{ -lean_object* x_170; lean_object* x_171; -lean_dec_ref(x_5); -x_170 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_170, 0, x_1); -lean_ctor_set(x_170, 1, x_2); -lean_ctor_set(x_170, 2, x_3); -lean_ctor_set(x_170, 3, x_4); -x_171 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_171, 0, x_170); -lean_ctor_set(x_171, 1, x_6); -return x_171; -} -} -} -} -static lean_object* _init_l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__0() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Id_instMonad___lam__0), 4, 0); -return x_1; -} -} -static lean_object* _init_l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Id_instMonad___lam__1___boxed), 4, 0); -return x_1; -} -} -static lean_object* _init_l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Id_instMonad___lam__2___boxed), 2, 0); -return x_1; -} -} -static lean_object* _init_l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Id_instMonad___lam__3), 4, 0); -return x_1; -} -} -static lean_object* _init_l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Id_instMonad___lam__4___boxed), 4, 0); -return x_1; -} -} -static lean_object* _init_l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__5() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Id_instMonad___lam__5___boxed), 4, 0); -return x_1; -} -} -static lean_object* _init_l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__6() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Id_instMonad___lam__6), 4, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; 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_4 = l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__0; -x_5 = l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__1; -x_6 = l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__2; -x_7 = l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__3; -x_8 = l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__4; -x_9 = l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__5; -x_10 = l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__6; +x_4 = l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__0; +x_5 = l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__1; +x_6 = l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__2; +x_7 = l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__3; +x_8 = l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__4; +x_9 = l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__5; +x_10 = l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__6; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_4); lean_ctor_set(x_11, 1, x_5); @@ -4018,7 +3922,7 @@ lean_closure_set(x_22, 2, x_13); x_23 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_23, 0, x_21); lean_ctor_set(x_23, 1, x_22); -x_24 = lean_box(0); +x_24 = l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___at___00Lean_IR_ExplicitBoxing_mkBoxedVersionAux_spec__1___redArg___closed__0; x_25 = l_instInhabitedOfMonad___redArg(x_23, x_24); x_26 = lean_alloc_closure((void*)(l_instInhabitedForall___redArg___lam__0___boxed), 2, 1); lean_closure_set(x_26, 0, x_25); @@ -4027,7 +3931,15 @@ x_28 = lean_apply_2(x_27, x_2, x_3); return x_28; } } -static lean_object* _init_l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__0() { +static lean_object* _init_l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_IR_ExplicitBoxing_boxArgsIfNeeded___lam__0___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__1() { _start: { lean_object* x_1; @@ -4035,15 +3947,15 @@ x_1 = lean_mk_string_unchecked("Lean.Compiler.IR.Boxing", 23, 23); return x_1; } } -static lean_object* _init_l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__1() { +static lean_object* _init_l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__2() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.IR.ExplicitBoxing.tryCorrectVDeclType", 42, 42); +x_1 = lean_mk_string_unchecked("Lean.IR.ExplicitBoxing.visitCtorExpr", 36, 36); return x_1; } } -static lean_object* _init_l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__2() { +static lean_object* _init_l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__3() { _start: { lean_object* x_1; @@ -4051,316 +3963,310 @@ x_1 = lean_mk_string_unchecked("unreachable code has been reached", 33, 33); return x_1; } } -static lean_object* _init_l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__3() { +static lean_object* _init_l_Lean_IR_ExplicitBoxing_visitCtorExpr___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_IR_ExplicitBoxing_tryCorrectVDeclType___closed__2; -x_2 = lean_unsigned_to_nat(42u); -x_3 = lean_unsigned_to_nat(288u); -x_4 = l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__1; -x_5 = l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__0; +x_1 = l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__3; +x_2 = lean_unsigned_to_nat(44u); +x_3 = lean_unsigned_to_nat(278u); +x_4 = l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__2; +x_5 = l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__1; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitCtorExpr___lam__0___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_5; lean_object* x_6; -switch (lean_obj_tag(x_2)) { -case 0: +lean_object* x_3; +x_3 = l_Lean_IR_ExplicitBoxing_visitCtorExpr___lam__0(x_1, x_2); +lean_dec(x_2); +lean_dec_ref(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitCtorExpr(lean_object* x_1, lean_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_10; -lean_dec_ref(x_3); -x_10 = !lean_is_exclusive(x_2); -if (x_10 == 0) +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_37; uint8_t x_38; uint8_t x_77; +x_37 = l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__0; +x_77 = l_Lean_IR_CtorInfo_isScalar(x_3); +if (x_77 == 0) { -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_2, 1); -lean_dec(x_11); -x_12 = lean_ctor_get(x_2, 0); -lean_dec(x_12); -lean_ctor_set(x_2, 1, x_4); -lean_ctor_set(x_2, 0, x_1); -return x_2; +x_38 = x_77; +goto block_76; } else { -lean_object* x_13; -lean_dec(x_2); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_1); -lean_ctor_set(x_13, 1, x_4); -return x_13; -} +uint8_t x_78; +x_78 = l_Lean_IR_IRType_isScalar(x_2); +x_38 = x_78; +goto block_76; } -case 2: +block_36: { -lean_object* x_14; -lean_dec_ref(x_3); -lean_dec_ref(x_2); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_1); -lean_ctor_set(x_14, 1, x_4); -return x_14; -} -case 4: +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = lean_alloc_closure((void*)(l_Lean_IR_ExplicitBoxing_visitCtorExpr___lam__0___boxed), 2, 1); +lean_closure_set(x_11, 0, x_8); +x_12 = l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux(x_4, x_11, x_9, x_10); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) { -uint8_t x_15; -lean_dec_ref(x_3); -lean_dec(x_1); -x_15 = !lean_is_exclusive(x_2); +lean_object* x_14; uint8_t x_15; +x_14 = lean_ctor_get(x_12, 0); +x_15 = !lean_is_exclusive(x_14); if (x_15 == 0) { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_2, 1); -lean_dec(x_16); -x_17 = lean_ctor_get(x_2, 0); -lean_dec(x_17); -x_18 = lean_box(5); -lean_ctor_set_tag(x_2, 0); -lean_ctor_set(x_2, 1, x_4); -lean_ctor_set(x_2, 0, x_18); -return x_2; +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_16 = lean_ctor_get(x_12, 1); +x_17 = lean_ctor_get(x_14, 0); +x_18 = lean_ctor_get(x_14, 1); +lean_ctor_set(x_12, 1, x_17); +lean_ctor_set(x_12, 0, x_3); +x_19 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_19, 0, x_1); +lean_ctor_set(x_19, 1, x_2); +lean_ctor_set(x_19, 2, x_12); +lean_ctor_set(x_19, 3, x_5); +x_20 = l_Lean_IR_reshape(x_18, x_19); +lean_ctor_set(x_14, 1, x_16); +lean_ctor_set(x_14, 0, x_20); +return x_14; } else { -lean_object* x_19; lean_object* x_20; -lean_dec(x_2); -x_19 = lean_box(5); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_4); -return x_20; -} -} -case 5: -{ -lean_object* x_21; -lean_dec_ref(x_3); -lean_dec_ref(x_2); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_1); -lean_ctor_set(x_21, 1, x_4); -return x_21; -} -case 6: -{ -lean_object* x_22; lean_object* x_23; uint8_t x_24; -lean_dec(x_1); -x_22 = lean_ctor_get(x_2, 0); +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_21 = lean_ctor_get(x_12, 1); +x_22 = lean_ctor_get(x_14, 0); +x_23 = lean_ctor_get(x_14, 1); +lean_inc(x_23); lean_inc(x_22); -lean_dec_ref(x_2); -x_23 = l_Lean_IR_ExplicitBoxing_getDecl(x_22, x_3, x_4); -x_24 = !lean_is_exclusive(x_23); -if (x_24 == 0) -{ -lean_object* x_25; lean_object* x_26; -x_25 = lean_ctor_get(x_23, 0); -x_26 = l_Lean_IR_Decl_resultType(x_25); -lean_dec(x_25); -lean_ctor_set(x_23, 0, x_26); -return x_23; +lean_dec(x_14); +lean_ctor_set(x_12, 1, x_22); +lean_ctor_set(x_12, 0, x_3); +x_24 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_24, 0, x_1); +lean_ctor_set(x_24, 1, x_2); +lean_ctor_set(x_24, 2, x_12); +lean_ctor_set(x_24, 3, x_5); +x_25 = l_Lean_IR_reshape(x_23, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_21); +return x_26; +} } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_27 = lean_ctor_get(x_23, 0); -x_28 = lean_ctor_get(x_23, 1); +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_27 = lean_ctor_get(x_12, 0); +x_28 = lean_ctor_get(x_12, 1); lean_inc(x_28); lean_inc(x_27); -lean_dec(x_23); -x_29 = l_Lean_IR_Decl_resultType(x_27); -lean_dec(x_27); -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_28); -return x_30; -} +lean_dec(x_12); +x_29 = lean_ctor_get(x_27, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_27, 1); +lean_inc(x_30); +if (lean_is_exclusive(x_27)) { + lean_ctor_release(x_27, 0); + lean_ctor_release(x_27, 1); + x_31 = x_27; +} else { + lean_dec_ref(x_27); + x_31 = lean_box(0); } -case 7: -{ -uint8_t x_31; -lean_dec_ref(x_3); -lean_dec(x_1); -x_31 = !lean_is_exclusive(x_2); -if (x_31 == 0) -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_2, 1); -lean_dec(x_32); -x_33 = lean_ctor_get(x_2, 0); -lean_dec(x_33); -x_34 = lean_box(7); -lean_ctor_set_tag(x_2, 0); -lean_ctor_set(x_2, 1, x_4); -lean_ctor_set(x_2, 0, x_34); -return x_2; +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_3); +lean_ctor_set(x_32, 1, x_29); +x_33 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_33, 0, x_1); +lean_ctor_set(x_33, 1, x_2); +lean_ctor_set(x_33, 2, x_32); +lean_ctor_set(x_33, 3, x_5); +x_34 = l_Lean_IR_reshape(x_30, x_33); +if (lean_is_scalar(x_31)) { + x_35 = lean_alloc_ctor(0, 2, 0); +} else { + x_35 = x_31; } -else -{ -lean_object* x_35; lean_object* x_36; -lean_dec(x_2); -x_35 = lean_box(7); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_4); -return x_36; +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_28); +return x_35; } } -case 8: +block_76: { -uint8_t x_37; -lean_dec_ref(x_3); -x_37 = !lean_is_exclusive(x_2); -if (x_37 == 0) +if (x_38 == 0) { -lean_object* x_38; lean_object* x_39; -x_38 = lean_ctor_get(x_2, 1); -lean_dec(x_38); -x_39 = lean_ctor_get(x_2, 0); -lean_dec(x_39); -lean_ctor_set_tag(x_2, 0); -lean_ctor_set(x_2, 1, x_4); -lean_ctor_set(x_2, 0, x_1); -return x_2; +if (lean_obj_tag(x_2) == 10) +{ +lean_object* x_39; +x_39 = lean_ctor_get(x_2, 1); +lean_inc_ref(x_39); +x_8 = x_39; +x_9 = x_6; +x_10 = x_7; +goto block_36; } else { -lean_object* x_40; -lean_dec(x_2); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_1); -lean_ctor_set(x_40, 1, x_4); -return x_40; -} -} -case 9: +if (lean_obj_tag(x_2) == 11) { -lean_object* x_41; lean_object* x_42; -lean_dec_ref(x_2); -lean_dec(x_1); -x_41 = l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__3; -x_42 = l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0(x_41, x_3, x_4); -return x_42; -} -case 10: +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_40 = lean_ctor_get(x_2, 1); +x_41 = lean_ctor_get(x_3, 1); +x_42 = lean_box(0); +x_43 = lean_array_get_borrowed(x_42, x_40, x_41); +if (lean_obj_tag(x_43) == 10) { -lean_dec_ref(x_2); -lean_dec(x_1); -x_5 = x_3; -x_6 = x_4; -goto block_9; +lean_object* x_44; +x_44 = lean_ctor_get(x_43, 1); +lean_inc_ref(x_44); +x_8 = x_44; +x_9 = x_6; +x_10 = x_7; +goto block_36; } -case 11: +else { -lean_object* x_43; +lean_object* x_45; lean_object* x_46; +lean_dec(x_5); lean_dec_ref(x_3); lean_dec_ref(x_2); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_1); -lean_ctor_set(x_43, 1, x_4); -return x_43; -} -case 12: -{ -lean_dec_ref(x_2); lean_dec(x_1); -x_5 = x_3; -x_6 = x_4; -goto block_9; -} -default: -{ -lean_object* x_44; -lean_dec_ref(x_3); -lean_dec_ref(x_2); -x_44 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_44, 0, x_1); -lean_ctor_set(x_44, 1, x_4); -return x_44; +x_45 = l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__4; +x_46 = l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0(x_45, x_6, x_7); +return x_46; } } -block_9: +else { -lean_object* x_7; lean_object* x_8; -x_7 = l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__3; -x_8 = l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0(x_7, x_5, x_6); -return x_8; +lean_object* x_47; uint8_t x_48; +x_47 = l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux(x_4, x_37, x_6, x_7); +x_48 = !lean_is_exclusive(x_47); +if (x_48 == 0) +{ +lean_object* x_49; uint8_t x_50; +x_49 = lean_ctor_get(x_47, 0); +x_50 = !lean_is_exclusive(x_49); +if (x_50 == 0) +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_51 = lean_ctor_get(x_47, 1); +x_52 = lean_ctor_get(x_49, 0); +x_53 = lean_ctor_get(x_49, 1); +lean_ctor_set(x_47, 1, x_52); +lean_ctor_set(x_47, 0, x_3); +x_54 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_54, 0, x_1); +lean_ctor_set(x_54, 1, x_2); +lean_ctor_set(x_54, 2, x_47); +lean_ctor_set(x_54, 3, x_5); +x_55 = l_Lean_IR_reshape(x_53, x_54); +lean_ctor_set(x_49, 1, x_51); +lean_ctor_set(x_49, 0, x_55); +return x_49; } +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_56 = lean_ctor_get(x_47, 1); +x_57 = lean_ctor_get(x_49, 0); +x_58 = lean_ctor_get(x_49, 1); +lean_inc(x_58); +lean_inc(x_57); +lean_dec(x_49); +lean_ctor_set(x_47, 1, x_57); +lean_ctor_set(x_47, 0, x_3); +x_59 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_59, 0, x_1); +lean_ctor_set(x_59, 1, x_2); +lean_ctor_set(x_59, 2, x_47); +lean_ctor_set(x_59, 3, x_5); +x_60 = l_Lean_IR_reshape(x_58, x_59); +x_61 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_56); +return x_61; } } -LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +else { -lean_object* x_7; lean_object* x_8; -x_7 = lean_alloc_ctor(4, 4, 0); -lean_ctor_set(x_7, 0, x_1); -lean_ctor_set(x_7, 1, x_2); -lean_ctor_set(x_7, 2, x_4); -lean_ctor_set(x_7, 3, x_3); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_7); -lean_ctor_set(x_8, 1, x_6); -return x_8; +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_62 = lean_ctor_get(x_47, 0); +x_63 = lean_ctor_get(x_47, 1); +lean_inc(x_63); +lean_inc(x_62); +lean_dec(x_47); +x_64 = lean_ctor_get(x_62, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_62, 1); +lean_inc(x_65); +if (lean_is_exclusive(x_62)) { + lean_ctor_release(x_62, 0); + lean_ctor_release(x_62, 1); + x_66 = x_62; +} else { + lean_dec_ref(x_62); + x_66 = lean_box(0); +} +x_67 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_67, 0, x_3); +lean_ctor_set(x_67, 1, x_64); +x_68 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_68, 0, x_1); +lean_ctor_set(x_68, 1, x_2); +lean_ctor_set(x_68, 2, x_67); +lean_ctor_set(x_68, 3, x_5); +x_69 = l_Lean_IR_reshape(x_65, x_68); +if (lean_is_scalar(x_66)) { + x_70 = lean_alloc_ctor(0, 2, 0); +} else { + x_70 = x_66; } +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_63); +return x_70; } -LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; lean_object* x_10; -x_9 = lean_alloc_ctor(5, 6, 0); -lean_ctor_set(x_9, 0, x_1); -lean_ctor_set(x_9, 1, x_2); -lean_ctor_set(x_9, 2, x_3); -lean_ctor_set(x_9, 3, x_6); -lean_ctor_set(x_9, 4, x_4); -lean_ctor_set(x_9, 5, x_5); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_8); -return x_10; } } -LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +} +else { -lean_object* x_7; lean_object* x_8; -x_7 = lean_alloc_ctor(9, 4, 0); -lean_ctor_set(x_7, 0, x_1); -lean_ctor_set(x_7, 1, x_4); -lean_ctor_set(x_7, 2, x_2); -lean_ctor_set(x_7, 3, x_3); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_7); -lean_ctor_set(x_8, 1, x_6); -return x_8; +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; +lean_dec_ref(x_6); +x_71 = lean_ctor_get(x_3, 1); +lean_inc(x_71); +lean_dec_ref(x_3); +x_72 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_72, 0, x_71); +x_73 = lean_alloc_ctor(11, 1, 0); +lean_ctor_set(x_73, 0, x_72); +x_74 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_74, 0, x_1); +lean_ctor_set(x_74, 1, x_2); +lean_ctor_set(x_74, 2, x_73); +lean_ctor_set(x_74, 3, x_5); +x_75 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_75, 0, x_74); +lean_ctor_set(x_75, 1, x_7); +return x_75; } } -LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; lean_object* x_5; -x_4 = lean_alloc_ctor(10, 1, 0); -lean_ctor_set(x_4, 0, x_1); -x_5 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_5, 0, x_4); -lean_ctor_set(x_5, 1, x_3); -return x_5; } } -LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitCtorExpr___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_5; lean_object* x_6; -x_5 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_5, 0, x_2); -x_6 = lean_apply_3(x_1, x_5, x_3, x_4); -return x_6; +lean_object* x_8; +x_8 = l_Lean_IR_ExplicitBoxing_visitCtorExpr(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec_ref(x_4); +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitVDeclExpr___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; @@ -4370,972 +4276,4366 @@ lean_inc(x_5); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitVDeclExpr___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__3(x_1, x_2, x_3); +x_4 = l_Lean_IR_ExplicitBoxing_visitVDeclExpr___lam__1(x_1, x_2, x_3); +lean_dec(x_3); lean_dec_ref(x_2); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitVDeclExpr(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_4; -x_4 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__5(x_1, x_2, x_3); -lean_dec(x_3); -lean_dec_ref(x_2); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -switch (lean_obj_tag(x_1)) { +switch (lean_obj_tag(x_3)) { case 0: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); -x_5 = lean_ctor_get(x_1, 1); -lean_inc(x_5); -x_6 = lean_ctor_get(x_1, 2); -lean_inc_ref(x_6); -x_7 = lean_ctor_get(x_1, 3); -lean_inc(x_7); -lean_dec_ref(x_1); -lean_inc_ref(x_2); -lean_inc_ref(x_6); -x_8 = l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType(x_5, x_6, x_2, x_3); -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_8, 1); -lean_inc(x_10); +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_ctor_get(x_3, 0); +lean_inc_ref(x_7); +x_8 = lean_ctor_get(x_3, 1); +lean_inc_ref(x_8); +lean_dec_ref(x_3); +x_9 = l_Lean_IR_ExplicitBoxing_visitCtorExpr(x_1, x_2, x_7, x_8, x_4, x_5, x_6); lean_dec_ref(x_8); -x_11 = lean_ctor_get(x_2, 0); -x_12 = lean_ctor_get(x_2, 1); -x_13 = lean_ctor_get(x_2, 2); -x_14 = lean_ctor_get(x_2, 3); -x_15 = lean_ctor_get(x_2, 4); -lean_inc_ref(x_6); -lean_inc(x_9); -lean_inc(x_4); -lean_inc(x_12); -x_16 = l_Lean_IR_LocalContext_addLocal(x_12, x_4, x_9, x_6); -lean_inc_ref(x_15); -lean_inc_ref(x_14); -lean_inc(x_13); -lean_inc(x_11); -x_17 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_17, 0, x_11); -lean_ctor_set(x_17, 1, x_16); -lean_ctor_set(x_17, 2, x_13); -lean_ctor_set(x_17, 3, x_14); -lean_ctor_set(x_17, 4, x_15); -x_18 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_7, x_17, x_10); -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); -lean_inc(x_20); -lean_dec_ref(x_18); -x_21 = l_Lean_IR_ExplicitBoxing_visitVDeclExpr(x_4, x_9, x_6, x_19, x_2, x_20); -return x_21; +return x_9; } -case 1: +case 2: { -uint8_t x_22; -x_22 = !lean_is_exclusive(x_1); -if (x_22 == 0) +uint8_t x_10; +x_10 = !lean_is_exclusive(x_3); +if (x_10 == 0) { -uint8_t x_23; -x_23 = !lean_is_exclusive(x_2); -if (x_23 == 0) +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_11 = lean_ctor_get(x_3, 2); +x_12 = l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__0; +x_13 = l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux(x_11, x_12, x_5, x_6); +lean_dec_ref(x_11); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec_ref(x_13); +x_16 = !lean_is_exclusive(x_14); +if (x_16 == 0) { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; -x_24 = lean_ctor_get(x_1, 0); -x_25 = lean_ctor_get(x_1, 1); -x_26 = lean_ctor_get(x_1, 2); -x_27 = lean_ctor_get(x_1, 3); -x_28 = lean_ctor_get(x_2, 0); -x_29 = lean_ctor_get(x_2, 1); -x_30 = lean_ctor_get(x_2, 2); -x_31 = lean_ctor_get(x_2, 3); -x_32 = lean_ctor_get(x_2, 4); +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = lean_ctor_get(x_14, 0); +x_18 = lean_ctor_get(x_14, 1); +lean_ctor_set(x_3, 2, x_17); +x_19 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_19, 0, x_1); +lean_ctor_set(x_19, 1, x_2); +lean_ctor_set(x_19, 2, x_3); +lean_ctor_set(x_19, 3, x_4); +x_20 = l_Lean_IR_reshape(x_18, x_19); +lean_ctor_set(x_14, 1, x_15); +lean_ctor_set(x_14, 0, x_20); +return x_14; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +lean_ctor_set(x_3, 2, x_21); +x_23 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_23, 0, x_1); +lean_ctor_set(x_23, 1, x_2); +lean_ctor_set(x_23, 2, x_3); +lean_ctor_set(x_23, 3, x_4); +x_24 = l_Lean_IR_reshape(x_22, x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_15); +return x_25; +} +} +else +{ +lean_object* x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_26 = lean_ctor_get(x_3, 0); +x_27 = lean_ctor_get(x_3, 1); +x_28 = lean_ctor_get_uint8(x_3, sizeof(void*)*3); +x_29 = lean_ctor_get(x_3, 2); lean_inc(x_29); -x_33 = l_Lean_IR_LocalContext_addParams(x_29, x_25); -lean_inc_ref(x_32); -lean_inc_ref(x_31); -lean_inc(x_30); -lean_inc(x_28); -lean_ctor_set(x_2, 1, x_33); -x_34 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_26, x_2, x_3); -x_35 = lean_ctor_get(x_34, 0); -lean_inc(x_35); -x_36 = lean_ctor_get(x_34, 1); -lean_inc(x_36); -lean_dec_ref(x_34); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_3); +x_30 = l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__0; +x_31 = l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux(x_29, x_30, x_5, x_6); +lean_dec_ref(x_29); +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec_ref(x_31); +x_34 = lean_ctor_get(x_32, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_32, 1); lean_inc(x_35); -lean_inc_ref(x_25); -lean_inc(x_24); -x_37 = l_Lean_IR_LocalContext_addJP(x_29, x_24, x_25, x_35); -x_38 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_38, 0, x_28); -lean_ctor_set(x_38, 1, x_37); -lean_ctor_set(x_38, 2, x_30); -lean_ctor_set(x_38, 3, x_31); -lean_ctor_set(x_38, 4, x_32); -x_39 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_27, x_38, x_36); -x_40 = !lean_is_exclusive(x_39); -if (x_40 == 0) +if (lean_is_exclusive(x_32)) { + lean_ctor_release(x_32, 0); + lean_ctor_release(x_32, 1); + x_36 = x_32; +} else { + lean_dec_ref(x_32); + x_36 = lean_box(0); +} +x_37 = lean_alloc_ctor(2, 3, 1); +lean_ctor_set(x_37, 0, x_26); +lean_ctor_set(x_37, 1, x_27); +lean_ctor_set(x_37, 2, x_34); +lean_ctor_set_uint8(x_37, sizeof(void*)*3, x_28); +x_38 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_38, 0, x_1); +lean_ctor_set(x_38, 1, x_2); +lean_ctor_set(x_38, 2, x_37); +lean_ctor_set(x_38, 3, x_4); +x_39 = l_Lean_IR_reshape(x_35, x_38); +if (lean_is_scalar(x_36)) { + x_40 = lean_alloc_ctor(0, 2, 0); +} else { + x_40 = x_36; +} +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_33); +return x_40; +} +} +case 6: { -lean_object* x_41; -x_41 = lean_ctor_get(x_39, 0); -lean_ctor_set(x_1, 3, x_41); -lean_ctor_set(x_1, 2, x_35); -lean_ctor_set(x_39, 0, x_1); -return x_39; +uint8_t x_41; +x_41 = !lean_is_exclusive(x_3); +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; +x_42 = lean_ctor_get(x_3, 0); +x_43 = lean_ctor_get(x_3, 1); +lean_inc_ref(x_5); +lean_inc(x_42); +x_44 = l_Lean_IR_ExplicitBoxing_getDecl(x_42, x_5, x_6); +x_45 = lean_ctor_get(x_44, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_44, 1); +lean_inc(x_46); +lean_dec_ref(x_44); +x_47 = l_Lean_IR_Decl_params(x_45); +x_48 = l_Lean_IR_ExplicitBoxing_mkBoxedVersionAux___closed__0; +x_49 = lean_alloc_closure((void*)(l_Lean_IR_ExplicitBoxing_visitVDeclExpr___lam__1___boxed), 3, 2); +lean_closure_set(x_49, 0, x_48); +lean_closure_set(x_49, 1, x_47); +lean_inc_ref(x_5); +x_50 = l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux(x_43, x_49, x_5, x_46); +lean_dec_ref(x_43); +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +x_52 = lean_ctor_get(x_50, 1); +lean_inc(x_52); +lean_dec_ref(x_50); +x_53 = lean_ctor_get(x_51, 0); +lean_inc(x_53); +x_54 = lean_ctor_get(x_51, 1); +lean_inc(x_54); +lean_dec(x_51); +lean_ctor_set(x_3, 1, x_53); +x_55 = l_Lean_IR_Decl_resultType(x_45); +lean_dec(x_45); +x_56 = l_Lean_IR_ExplicitBoxing_castResultIfNeeded(x_1, x_2, x_3, x_55, x_4, x_5, x_52); +lean_dec(x_55); +x_57 = !lean_is_exclusive(x_56); +if (x_57 == 0) +{ +lean_object* x_58; lean_object* x_59; +x_58 = lean_ctor_get(x_56, 0); +x_59 = l_Lean_IR_reshape(x_54, x_58); +lean_ctor_set(x_56, 0, x_59); +return x_56; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_39, 0); -x_43 = lean_ctor_get(x_39, 1); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_39); -lean_ctor_set(x_1, 3, x_42); -lean_ctor_set(x_1, 2, x_35); -x_44 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_44, 0, x_1); -lean_ctor_set(x_44, 1, x_43); -return x_44; +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_60 = lean_ctor_get(x_56, 0); +x_61 = lean_ctor_get(x_56, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_56); +x_62 = l_Lean_IR_reshape(x_54, x_60); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_62); +lean_ctor_set(x_63, 1, x_61); +return x_63; } } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_45 = lean_ctor_get(x_1, 0); -x_46 = lean_ctor_get(x_1, 1); -x_47 = lean_ctor_get(x_1, 2); -x_48 = lean_ctor_get(x_1, 3); -x_49 = lean_ctor_get(x_2, 0); -x_50 = lean_ctor_get(x_2, 1); -x_51 = lean_ctor_get(x_2, 2); -x_52 = lean_ctor_get(x_2, 3); -x_53 = lean_ctor_get(x_2, 4); -lean_inc(x_53); -lean_inc(x_52); -lean_inc(x_51); -lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_2); -lean_inc(x_50); -x_54 = l_Lean_IR_LocalContext_addParams(x_50, x_46); -lean_inc_ref(x_53); -lean_inc_ref(x_52); -lean_inc(x_51); -lean_inc(x_49); -x_55 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_55, 0, x_49); -lean_ctor_set(x_55, 1, x_54); -lean_ctor_set(x_55, 2, x_51); -lean_ctor_set(x_55, 3, x_52); -lean_ctor_set(x_55, 4, x_53); -x_56 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_47, x_55, x_3); -x_57 = lean_ctor_get(x_56, 0); -lean_inc(x_57); -x_58 = lean_ctor_get(x_56, 1); -lean_inc(x_58); -lean_dec_ref(x_56); -lean_inc(x_57); -lean_inc_ref(x_46); -lean_inc(x_45); -x_59 = l_Lean_IR_LocalContext_addJP(x_50, x_45, x_46, x_57); -x_60 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_60, 0, x_49); -lean_ctor_set(x_60, 1, x_59); -lean_ctor_set(x_60, 2, x_51); -lean_ctor_set(x_60, 3, x_52); -lean_ctor_set(x_60, 4, x_53); -x_61 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_48, x_60, x_58); -x_62 = lean_ctor_get(x_61, 0); -lean_inc(x_62); -x_63 = lean_ctor_get(x_61, 1); -lean_inc(x_63); -if (lean_is_exclusive(x_61)) { - lean_ctor_release(x_61, 0); - lean_ctor_release(x_61, 1); - x_64 = x_61; +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_64 = lean_ctor_get(x_3, 0); +x_65 = lean_ctor_get(x_3, 1); +lean_inc(x_65); +lean_inc(x_64); +lean_dec(x_3); +lean_inc_ref(x_5); +lean_inc(x_64); +x_66 = l_Lean_IR_ExplicitBoxing_getDecl(x_64, x_5, x_6); +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec_ref(x_66); +x_69 = l_Lean_IR_Decl_params(x_67); +x_70 = l_Lean_IR_ExplicitBoxing_mkBoxedVersionAux___closed__0; +x_71 = lean_alloc_closure((void*)(l_Lean_IR_ExplicitBoxing_visitVDeclExpr___lam__1___boxed), 3, 2); +lean_closure_set(x_71, 0, x_70); +lean_closure_set(x_71, 1, x_69); +lean_inc_ref(x_5); +x_72 = l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux(x_65, x_71, x_5, x_68); +lean_dec_ref(x_65); +x_73 = lean_ctor_get(x_72, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_72, 1); +lean_inc(x_74); +lean_dec_ref(x_72); +x_75 = lean_ctor_get(x_73, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_73, 1); +lean_inc(x_76); +lean_dec(x_73); +x_77 = lean_alloc_ctor(6, 2, 0); +lean_ctor_set(x_77, 0, x_64); +lean_ctor_set(x_77, 1, x_75); +x_78 = l_Lean_IR_Decl_resultType(x_67); +lean_dec(x_67); +x_79 = l_Lean_IR_ExplicitBoxing_castResultIfNeeded(x_1, x_2, x_77, x_78, x_4, x_5, x_74); +lean_dec(x_78); +x_80 = lean_ctor_get(x_79, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_79, 1); +lean_inc(x_81); +if (lean_is_exclusive(x_79)) { + lean_ctor_release(x_79, 0); + lean_ctor_release(x_79, 1); + x_82 = x_79; } else { - lean_dec_ref(x_61); - x_64 = lean_box(0); + lean_dec_ref(x_79); + x_82 = lean_box(0); } -lean_ctor_set(x_1, 3, x_62); -lean_ctor_set(x_1, 2, x_57); -if (lean_is_scalar(x_64)) { - x_65 = lean_alloc_ctor(0, 2, 0); +x_83 = l_Lean_IR_reshape(x_76, x_80); +if (lean_is_scalar(x_82)) { + x_84 = lean_alloc_ctor(0, 2, 0); } else { - x_65 = x_64; + x_84 = x_82; } -lean_ctor_set(x_65, 0, x_1); -lean_ctor_set(x_65, 1, x_63); -return x_65; +lean_ctor_set(x_84, 0, x_83); +lean_ctor_set(x_84, 1, x_81); +return x_84; } } -else +case 7: { -lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_66 = lean_ctor_get(x_1, 0); -x_67 = lean_ctor_get(x_1, 1); -x_68 = lean_ctor_get(x_1, 2); -x_69 = lean_ctor_get(x_1, 3); -lean_inc(x_69); -lean_inc(x_68); -lean_inc(x_67); -lean_inc(x_66); -lean_dec(x_1); -x_70 = lean_ctor_get(x_2, 0); -lean_inc(x_70); -x_71 = lean_ctor_get(x_2, 1); -lean_inc(x_71); -x_72 = lean_ctor_get(x_2, 2); -lean_inc(x_72); -x_73 = lean_ctor_get(x_2, 3); -lean_inc_ref(x_73); -x_74 = lean_ctor_get(x_2, 4); -lean_inc_ref(x_74); -if (lean_is_exclusive(x_2)) { - lean_ctor_release(x_2, 0); - lean_ctor_release(x_2, 1); - lean_ctor_release(x_2, 2); - lean_ctor_release(x_2, 3); - lean_ctor_release(x_2, 4); - x_75 = x_2; -} else { - lean_dec_ref(x_2); - x_75 = lean_box(0); +uint8_t x_85; +x_85 = !lean_is_exclusive(x_3); +if (x_85 == 0) +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; uint8_t x_91; +x_86 = lean_ctor_get(x_3, 1); +x_87 = l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__0; +x_88 = l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux(x_86, x_87, x_5, x_6); +lean_dec_ref(x_86); +x_89 = lean_ctor_get(x_88, 0); +lean_inc(x_89); +x_90 = lean_ctor_get(x_88, 1); +lean_inc(x_90); +lean_dec_ref(x_88); +x_91 = !lean_is_exclusive(x_89); +if (x_91 == 0) +{ +lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_92 = lean_ctor_get(x_89, 0); +x_93 = lean_ctor_get(x_89, 1); +lean_ctor_set(x_3, 1, x_92); +x_94 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_94, 0, x_1); +lean_ctor_set(x_94, 1, x_2); +lean_ctor_set(x_94, 2, x_3); +lean_ctor_set(x_94, 3, x_4); +x_95 = l_Lean_IR_reshape(x_93, x_94); +lean_ctor_set(x_89, 1, x_90); +lean_ctor_set(x_89, 0, x_95); +return x_89; } -lean_inc(x_71); -x_76 = l_Lean_IR_LocalContext_addParams(x_71, x_67); -lean_inc_ref(x_74); -lean_inc_ref(x_73); -lean_inc(x_72); -lean_inc(x_70); -if (lean_is_scalar(x_75)) { - x_77 = lean_alloc_ctor(0, 5, 0); +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; +x_96 = lean_ctor_get(x_89, 0); +x_97 = lean_ctor_get(x_89, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_89); +lean_ctor_set(x_3, 1, x_96); +x_98 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_98, 0, x_1); +lean_ctor_set(x_98, 1, x_2); +lean_ctor_set(x_98, 2, x_3); +lean_ctor_set(x_98, 3, x_4); +x_99 = l_Lean_IR_reshape(x_97, x_98); +x_100 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_100, 0, x_99); +lean_ctor_set(x_100, 1, x_90); +return x_100; +} +} +else +{ +lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; +x_101 = lean_ctor_get(x_3, 0); +x_102 = lean_ctor_get(x_3, 1); +lean_inc(x_102); +lean_inc(x_101); +lean_dec(x_3); +x_103 = l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__0; +x_104 = l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux(x_102, x_103, x_5, x_6); +lean_dec_ref(x_102); +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec_ref(x_104); +x_107 = lean_ctor_get(x_105, 0); +lean_inc(x_107); +x_108 = lean_ctor_get(x_105, 1); +lean_inc(x_108); +if (lean_is_exclusive(x_105)) { + lean_ctor_release(x_105, 0); + lean_ctor_release(x_105, 1); + x_109 = x_105; } else { - x_77 = x_75; + lean_dec_ref(x_105); + x_109 = lean_box(0); +} +x_110 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_110, 0, x_101); +lean_ctor_set(x_110, 1, x_107); +x_111 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_111, 0, x_1); +lean_ctor_set(x_111, 1, x_2); +lean_ctor_set(x_111, 2, x_110); +lean_ctor_set(x_111, 3, x_4); +x_112 = l_Lean_IR_reshape(x_108, x_111); +if (lean_is_scalar(x_109)) { + x_113 = lean_alloc_ctor(0, 2, 0); +} else { + x_113 = x_109; } -lean_ctor_set(x_77, 0, x_70); -lean_ctor_set(x_77, 1, x_76); -lean_ctor_set(x_77, 2, x_72); -lean_ctor_set(x_77, 3, x_73); -lean_ctor_set(x_77, 4, x_74); -x_78 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_68, x_77, x_3); -x_79 = lean_ctor_get(x_78, 0); -lean_inc(x_79); -x_80 = lean_ctor_get(x_78, 1); -lean_inc(x_80); -lean_dec_ref(x_78); -lean_inc(x_79); -lean_inc_ref(x_67); -lean_inc(x_66); -x_81 = l_Lean_IR_LocalContext_addJP(x_71, x_66, x_67, x_79); -x_82 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_82, 0, x_70); -lean_ctor_set(x_82, 1, x_81); -lean_ctor_set(x_82, 2, x_72); -lean_ctor_set(x_82, 3, x_73); -lean_ctor_set(x_82, 4, x_74); -x_83 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_69, x_82, x_80); -x_84 = lean_ctor_get(x_83, 0); -lean_inc(x_84); -x_85 = lean_ctor_get(x_83, 1); -lean_inc(x_85); -if (lean_is_exclusive(x_83)) { - lean_ctor_release(x_83, 0); - lean_ctor_release(x_83, 1); - x_86 = x_83; +lean_ctor_set(x_113, 0, x_112); +lean_ctor_set(x_113, 1, x_106); +return x_113; +} +} +case 8: +{ +uint8_t x_114; +x_114 = !lean_is_exclusive(x_3); +if (x_114 == 0) +{ +lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; uint8_t x_123; +x_115 = lean_ctor_get(x_3, 1); +x_116 = l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__0; +x_117 = l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux(x_115, x_116, x_5, x_6); +lean_dec_ref(x_115); +x_118 = lean_ctor_get(x_117, 0); +lean_inc(x_118); +x_119 = lean_ctor_get(x_117, 1); +lean_inc(x_119); +lean_dec_ref(x_117); +x_120 = lean_ctor_get(x_118, 0); +lean_inc(x_120); +x_121 = lean_ctor_get(x_118, 1); +lean_inc(x_121); +lean_dec(x_118); +lean_ctor_set(x_3, 1, x_120); +x_122 = l_Lean_IR_ExplicitBoxing_unboxResultIfNeeded___redArg(x_1, x_2, x_3, x_4, x_119); +x_123 = !lean_is_exclusive(x_122); +if (x_123 == 0) +{ +lean_object* x_124; lean_object* x_125; +x_124 = lean_ctor_get(x_122, 0); +x_125 = l_Lean_IR_reshape(x_121, x_124); +lean_ctor_set(x_122, 0, x_125); +return x_122; +} +else +{ +lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; +x_126 = lean_ctor_get(x_122, 0); +x_127 = lean_ctor_get(x_122, 1); +lean_inc(x_127); +lean_inc(x_126); +lean_dec(x_122); +x_128 = l_Lean_IR_reshape(x_121, x_126); +x_129 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_129, 0, x_128); +lean_ctor_set(x_129, 1, x_127); +return x_129; +} +} +else +{ +lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; +x_130 = lean_ctor_get(x_3, 0); +x_131 = lean_ctor_get(x_3, 1); +lean_inc(x_131); +lean_inc(x_130); +lean_dec(x_3); +x_132 = l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__0; +x_133 = l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux(x_131, x_132, x_5, x_6); +lean_dec_ref(x_131); +x_134 = lean_ctor_get(x_133, 0); +lean_inc(x_134); +x_135 = lean_ctor_get(x_133, 1); +lean_inc(x_135); +lean_dec_ref(x_133); +x_136 = lean_ctor_get(x_134, 0); +lean_inc(x_136); +x_137 = lean_ctor_get(x_134, 1); +lean_inc(x_137); +lean_dec(x_134); +x_138 = lean_alloc_ctor(8, 2, 0); +lean_ctor_set(x_138, 0, x_130); +lean_ctor_set(x_138, 1, x_136); +x_139 = l_Lean_IR_ExplicitBoxing_unboxResultIfNeeded___redArg(x_1, x_2, x_138, x_4, x_135); +x_140 = lean_ctor_get(x_139, 0); +lean_inc(x_140); +x_141 = lean_ctor_get(x_139, 1); +lean_inc(x_141); +if (lean_is_exclusive(x_139)) { + lean_ctor_release(x_139, 0); + lean_ctor_release(x_139, 1); + x_142 = x_139; } else { - lean_dec_ref(x_83); - x_86 = lean_box(0); + lean_dec_ref(x_139); + x_142 = lean_box(0); } -x_87 = lean_alloc_ctor(1, 4, 0); -lean_ctor_set(x_87, 0, x_66); -lean_ctor_set(x_87, 1, x_67); -lean_ctor_set(x_87, 2, x_79); -lean_ctor_set(x_87, 3, x_84); -if (lean_is_scalar(x_86)) { - x_88 = lean_alloc_ctor(0, 2, 0); +x_143 = l_Lean_IR_reshape(x_137, x_140); +if (lean_is_scalar(x_142)) { + x_144 = lean_alloc_ctor(0, 2, 0); } else { - x_88 = x_86; + x_144 = x_142; } -lean_ctor_set(x_88, 0, x_87); -lean_ctor_set(x_88, 1, x_85); -return x_88; +lean_ctor_set(x_144, 0, x_143); +lean_ctor_set(x_144, 1, x_141); +return x_144; } } -case 4: +default: { -uint8_t x_89; -x_89 = !lean_is_exclusive(x_1); -if (x_89 == 0) +lean_object* x_145; lean_object* x_146; +lean_dec_ref(x_5); +x_145 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_145, 0, x_1); +lean_ctor_set(x_145, 1, x_2); +lean_ctor_set(x_145, 2, x_3); +lean_ctor_set(x_145, 3, x_4); +x_146 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_146, 0, x_145); +lean_ctor_set(x_146, 1, x_6); +return x_146; +} +} +} +} +LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; uint8_t x_101; -x_90 = lean_ctor_get(x_1, 0); -x_91 = lean_ctor_get(x_1, 1); -x_92 = lean_ctor_get(x_1, 2); -x_93 = lean_ctor_get(x_1, 3); -lean_inc_ref(x_2); -x_94 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_93, x_2, x_3); -x_95 = lean_ctor_get(x_94, 0); -lean_inc(x_95); -x_96 = lean_ctor_get(x_94, 1); -lean_inc(x_96); -lean_dec_ref(x_94); -x_97 = l_Lean_IR_ExplicitBoxing_getVarType(x_92, x_2, x_96); -x_98 = lean_ctor_get(x_97, 0); -lean_inc(x_98); -x_99 = lean_ctor_get(x_97, 1); -lean_inc(x_99); -lean_dec_ref(x_97); -x_100 = lean_box(5); -lean_inc(x_98); -x_101 = l_Lean_IR_ExplicitBoxing_eqvTypes(x_98, x_100); -if (x_101 == 0) +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; 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_4 = l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__0; +x_5 = l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__1; +x_6 = l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__2; +x_7 = l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__3; +x_8 = l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__4; +x_9 = l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__5; +x_10 = l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__6; +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_4); +lean_ctor_set(x_11, 1, x_5); +x_12 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_6); +lean_ctor_set(x_12, 2, x_7); +lean_ctor_set(x_12, 3, x_8); +lean_ctor_set(x_12, 4, x_9); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_10); +lean_inc_ref(x_13); +x_14 = lean_alloc_closure((void*)(l_StateT_instMonad___redArg___lam__1), 6, 1); +lean_closure_set(x_14, 0, x_13); +lean_inc_ref(x_13); +x_15 = lean_alloc_closure((void*)(l_StateT_instMonad___redArg___lam__4), 6, 1); +lean_closure_set(x_15, 0, x_13); +lean_inc_ref(x_13); +x_16 = lean_alloc_closure((void*)(l_StateT_instMonad___redArg___lam__7), 6, 1); +lean_closure_set(x_16, 0, x_13); +lean_inc_ref(x_13); +x_17 = lean_alloc_closure((void*)(l_StateT_instMonad___redArg___lam__9), 6, 1); +lean_closure_set(x_17, 0, x_13); +lean_inc_ref(x_13); +x_18 = lean_alloc_closure((void*)(l_StateT_map), 8, 3); +lean_closure_set(x_18, 0, lean_box(0)); +lean_closure_set(x_18, 1, lean_box(0)); +lean_closure_set(x_18, 2, x_13); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_14); +lean_inc_ref(x_13); +x_20 = lean_alloc_closure((void*)(l_StateT_pure), 6, 3); +lean_closure_set(x_20, 0, lean_box(0)); +lean_closure_set(x_20, 1, lean_box(0)); +lean_closure_set(x_20, 2, x_13); +x_21 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +lean_ctor_set(x_21, 2, x_15); +lean_ctor_set(x_21, 3, x_16); +lean_ctor_set(x_21, 4, x_17); +x_22 = lean_alloc_closure((void*)(l_StateT_bind), 8, 3); +lean_closure_set(x_22, 0, lean_box(0)); +lean_closure_set(x_22, 1, lean_box(0)); +lean_closure_set(x_22, 2, x_13); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +x_24 = lean_box(0); +x_25 = l_instInhabitedOfMonad___redArg(x_23, x_24); +x_26 = lean_alloc_closure((void*)(l_instInhabitedForall___redArg___lam__0___boxed), 2, 1); +lean_closure_set(x_26, 0, x_25); +x_27 = lean_panic_fn(x_26, x_1); +x_28 = lean_apply_2(x_27, x_2, x_3); +return x_28; +} +} +static lean_object* _init_l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__0() { +_start: { -lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; uint8_t x_109; -x_102 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(x_99); -x_103 = lean_ctor_get(x_102, 0); -lean_inc(x_103); -x_104 = lean_ctor_get(x_102, 1); -lean_inc(x_104); -lean_dec_ref(x_102); -lean_inc_ref(x_2); -x_105 = l_Lean_IR_ExplicitBoxing_mkCast(x_92, x_98, x_100, x_2, x_104); -x_106 = lean_ctor_get(x_105, 0); -lean_inc(x_106); -x_107 = lean_ctor_get(x_105, 1); -lean_inc(x_107); -lean_dec_ref(x_105); -lean_inc(x_103); -x_108 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__0(x_90, x_91, x_95, x_103, x_2, x_107); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.IR.ExplicitBoxing.tryCorrectVDeclType", 42, 42); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__3; +x_2 = lean_unsigned_to_nat(42u); +x_3 = lean_unsigned_to_nat(328u); +x_4 = l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__0; +x_5 = l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__1; +x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); +return x_6; +} +} +static lean_object* _init_l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__3; +x_2 = lean_unsigned_to_nat(41u); +x_3 = lean_unsigned_to_nat(323u); +x_4 = l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__0; +x_5 = l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__1; +x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType(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; +switch (lean_obj_tag(x_2)) { +case 3: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_10 = lean_ctor_get(x_2, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_2, 1); +lean_inc(x_11); +x_12 = lean_ctor_get(x_2, 2); +lean_inc(x_12); lean_dec_ref(x_2); -x_109 = !lean_is_exclusive(x_108); -if (x_109 == 0) +x_13 = l_Lean_IR_ExplicitBoxing_getVarType(x_12, x_3, x_4); +lean_dec(x_12); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +if (lean_is_exclusive(x_13)) { + lean_ctor_release(x_13, 0); + lean_ctor_release(x_13, 1); + x_16 = x_13; +} else { + lean_dec_ref(x_13); + x_16 = lean_box(0); +} +x_17 = lean_box(0); +switch (lean_obj_tag(x_14)) { +case 10: +{ +lean_object* x_22; +lean_dec(x_10); +lean_dec_ref(x_3); +lean_dec(x_1); +x_22 = lean_ctor_get(x_14, 1); +lean_inc_ref(x_22); +lean_dec_ref(x_14); +x_18 = x_22; +goto block_21; +} +case 11: +{ +lean_object* x_23; lean_object* x_24; +lean_dec(x_1); +x_23 = lean_ctor_get(x_14, 1); +lean_inc_ref(x_23); +lean_dec_ref(x_14); +x_24 = lean_array_get(x_17, x_23, x_10); +lean_dec(x_10); +lean_dec_ref(x_23); +if (lean_obj_tag(x_24) == 10) +{ +lean_object* x_25; +lean_dec_ref(x_3); +x_25 = lean_ctor_get(x_24, 1); +lean_inc_ref(x_25); +lean_dec_ref(x_24); +x_18 = x_25; +goto block_21; +} +else +{ +lean_object* x_26; lean_object* x_27; +lean_dec(x_24); +lean_dec(x_16); +lean_dec(x_11); +x_26 = l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__2; +x_27 = l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0(x_26, x_3, x_15); +return x_27; +} +} +default: +{ +lean_object* x_28; lean_object* x_29; +lean_dec(x_16); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_10); +lean_dec_ref(x_3); +x_28 = l_Lean_IR_IRType_boxed(x_1); +lean_dec(x_1); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_15); +return x_29; +} +} +block_21: +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_array_get(x_17, x_18, x_11); +lean_dec(x_11); +lean_dec_ref(x_18); +if (lean_is_scalar(x_16)) { + x_20 = lean_alloc_ctor(0, 2, 0); +} else { + x_20 = x_16; +} +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_15); +return x_20; +} +} +case 4: +{ +lean_object* x_30; lean_object* x_31; +lean_dec_ref(x_3); +lean_dec_ref(x_2); +lean_dec(x_1); +x_30 = lean_box(5); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_4); +return x_31; +} +case 6: +{ +lean_object* x_32; lean_object* x_33; uint8_t x_34; +lean_dec(x_1); +x_32 = lean_ctor_get(x_2, 0); +lean_inc(x_32); +lean_dec_ref(x_2); +x_33 = l_Lean_IR_ExplicitBoxing_getDecl(x_32, x_3, x_4); +x_34 = !lean_is_exclusive(x_33); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; +x_35 = lean_ctor_get(x_33, 0); +x_36 = l_Lean_IR_Decl_resultType(x_35); +lean_dec(x_35); +lean_ctor_set(x_33, 0, x_36); +return x_33; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_37 = lean_ctor_get(x_33, 0); +x_38 = lean_ctor_get(x_33, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_33); +x_39 = l_Lean_IR_Decl_resultType(x_37); +lean_dec(x_37); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_38); +return x_40; +} +} +case 7: +{ +uint8_t x_41; +lean_dec_ref(x_3); +lean_dec(x_1); +x_41 = !lean_is_exclusive(x_2); +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_2, 1); +lean_dec(x_42); +x_43 = lean_ctor_get(x_2, 0); +lean_dec(x_43); +x_44 = lean_box(7); +lean_ctor_set_tag(x_2, 0); +lean_ctor_set(x_2, 1, x_4); +lean_ctor_set(x_2, 0, x_44); +return x_2; +} +else +{ +lean_object* x_45; lean_object* x_46; +lean_dec(x_2); +x_45 = lean_box(7); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_4); +return x_46; +} +} +case 9: +{ +lean_object* x_47; lean_object* x_48; +lean_dec_ref(x_2); +lean_dec(x_1); +x_47 = l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__1; +x_48 = l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0(x_47, x_3, x_4); +return x_48; +} +case 10: +{ +lean_dec_ref(x_2); +lean_dec(x_1); +x_5 = x_3; +x_6 = x_4; +goto block_9; +} +case 12: +{ +lean_dec_ref(x_2); +lean_dec(x_1); +x_5 = x_3; +x_6 = x_4; +goto block_9; +} +default: +{ +lean_object* x_49; +lean_dec_ref(x_3); +lean_dec_ref(x_2); +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_1); +lean_ctor_set(x_49, 1, x_4); +return x_49; +} +} +block_9: +{ +lean_object* x_7; lean_object* x_8; +x_7 = l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__1; +x_8 = l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0(x_7, x_5, x_6); +return x_8; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_alloc_ctor(4, 5, 0); +lean_ctor_set(x_8, 0, x_1); +lean_ctor_set(x_8, 1, x_2); +lean_ctor_set(x_8, 2, x_3); +lean_ctor_set(x_8, 3, x_5); +lean_ctor_set(x_8, 4, x_4); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_7); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_alloc_ctor(5, 7, 0); +lean_ctor_set(x_10, 0, x_1); +lean_ctor_set(x_10, 1, x_2); +lean_ctor_set(x_10, 2, x_3); +lean_ctor_set(x_10, 3, x_4); +lean_ctor_set(x_10, 4, x_7); +lean_ctor_set(x_10, 5, x_5); +lean_ctor_set(x_10, 6, x_6); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_9); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__2(lean_object* x_1, 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_ctor(9, 4, 0); +lean_ctor_set(x_7, 0, x_1); +lean_ctor_set(x_7, 1, x_4); +lean_ctor_set(x_7, 2, x_2); +lean_ctor_set(x_7, 3, x_3); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_6); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_alloc_ctor(10, 1, 0); +lean_ctor_set(x_4, 0, x_1); +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_4); +lean_ctor_set(x_5, 1, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_5, 0, x_2); +x_6 = lean_apply_3(x_1, x_5, x_3, x_4); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_array_get_borrowed(x_1, x_2, x_3); +x_5 = lean_ctor_get(x_4, 1); +lean_inc(x_5); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__3(x_1, x_2, x_3); +lean_dec_ref(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__5(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec_ref(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 0: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_1, 1); +lean_inc(x_5); +x_6 = lean_ctor_get(x_1, 2); +lean_inc_ref(x_6); +x_7 = lean_ctor_get(x_1, 3); +lean_inc(x_7); +lean_dec_ref(x_1); +lean_inc_ref(x_2); +lean_inc_ref(x_6); +x_8 = l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType(x_5, x_6, x_2, x_3); +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_ref(x_8); +x_11 = lean_ctor_get(x_2, 0); +x_12 = lean_ctor_get(x_2, 1); +x_13 = lean_ctor_get(x_2, 2); +x_14 = lean_ctor_get(x_2, 3); +x_15 = lean_ctor_get(x_2, 4); +lean_inc_ref(x_6); +lean_inc(x_9); +lean_inc(x_4); +lean_inc(x_12); +x_16 = l_Lean_IR_LocalContext_addLocal(x_12, x_4, x_9, x_6); +lean_inc_ref(x_15); +lean_inc_ref(x_14); +lean_inc(x_13); +lean_inc(x_11); +x_17 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_17, 0, x_11); +lean_ctor_set(x_17, 1, x_16); +lean_ctor_set(x_17, 2, x_13); +lean_ctor_set(x_17, 3, x_14); +lean_ctor_set(x_17, 4, x_15); +x_18 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_7, x_17, x_10); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec_ref(x_18); +x_21 = l_Lean_IR_ExplicitBoxing_visitVDeclExpr(x_4, x_9, x_6, x_19, x_2, x_20); +return x_21; +} +case 1: +{ +uint8_t x_22; +x_22 = !lean_is_exclusive(x_1); +if (x_22 == 0) +{ +uint8_t x_23; +x_23 = !lean_is_exclusive(x_2); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; +x_24 = lean_ctor_get(x_1, 0); +x_25 = lean_ctor_get(x_1, 1); +x_26 = lean_ctor_get(x_1, 2); +x_27 = lean_ctor_get(x_1, 3); +x_28 = lean_ctor_get(x_2, 0); +x_29 = lean_ctor_get(x_2, 1); +x_30 = lean_ctor_get(x_2, 2); +x_31 = lean_ctor_get(x_2, 3); +x_32 = lean_ctor_get(x_2, 4); +lean_inc(x_29); +x_33 = l_Lean_IR_LocalContext_addParams(x_29, x_25); +lean_inc_ref(x_32); +lean_inc_ref(x_31); +lean_inc(x_30); +lean_inc(x_28); +lean_ctor_set(x_2, 1, x_33); +x_34 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_26, x_2, x_3); +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec_ref(x_34); +lean_inc(x_35); +lean_inc_ref(x_25); +lean_inc(x_24); +x_37 = l_Lean_IR_LocalContext_addJP(x_29, x_24, x_25, x_35); +x_38 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_38, 0, x_28); +lean_ctor_set(x_38, 1, x_37); +lean_ctor_set(x_38, 2, x_30); +lean_ctor_set(x_38, 3, x_31); +lean_ctor_set(x_38, 4, x_32); +x_39 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_27, x_38, x_36); +x_40 = !lean_is_exclusive(x_39); +if (x_40 == 0) +{ +lean_object* x_41; +x_41 = lean_ctor_get(x_39, 0); +lean_ctor_set(x_1, 3, x_41); +lean_ctor_set(x_1, 2, x_35); +lean_ctor_set(x_39, 0, x_1); +return x_39; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_39, 0); +x_43 = lean_ctor_get(x_39, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_39); +lean_ctor_set(x_1, 3, x_42); +lean_ctor_set(x_1, 2, x_35); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_1); +lean_ctor_set(x_44, 1, x_43); +return x_44; +} +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_45 = lean_ctor_get(x_1, 0); +x_46 = lean_ctor_get(x_1, 1); +x_47 = lean_ctor_get(x_1, 2); +x_48 = lean_ctor_get(x_1, 3); +x_49 = lean_ctor_get(x_2, 0); +x_50 = lean_ctor_get(x_2, 1); +x_51 = lean_ctor_get(x_2, 2); +x_52 = lean_ctor_get(x_2, 3); +x_53 = lean_ctor_get(x_2, 4); +lean_inc(x_53); +lean_inc(x_52); +lean_inc(x_51); +lean_inc(x_50); +lean_inc(x_49); +lean_dec(x_2); +lean_inc(x_50); +x_54 = l_Lean_IR_LocalContext_addParams(x_50, x_46); +lean_inc_ref(x_53); +lean_inc_ref(x_52); +lean_inc(x_51); +lean_inc(x_49); +x_55 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_55, 0, x_49); +lean_ctor_set(x_55, 1, x_54); +lean_ctor_set(x_55, 2, x_51); +lean_ctor_set(x_55, 3, x_52); +lean_ctor_set(x_55, 4, x_53); +x_56 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_47, x_55, x_3); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_56, 1); +lean_inc(x_58); +lean_dec_ref(x_56); +lean_inc(x_57); +lean_inc_ref(x_46); +lean_inc(x_45); +x_59 = l_Lean_IR_LocalContext_addJP(x_50, x_45, x_46, x_57); +x_60 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_60, 0, x_49); +lean_ctor_set(x_60, 1, x_59); +lean_ctor_set(x_60, 2, x_51); +lean_ctor_set(x_60, 3, x_52); +lean_ctor_set(x_60, 4, x_53); +x_61 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_48, x_60, x_58); +x_62 = lean_ctor_get(x_61, 0); +lean_inc(x_62); +x_63 = lean_ctor_get(x_61, 1); +lean_inc(x_63); +if (lean_is_exclusive(x_61)) { + lean_ctor_release(x_61, 0); + lean_ctor_release(x_61, 1); + x_64 = x_61; +} else { + lean_dec_ref(x_61); + x_64 = lean_box(0); +} +lean_ctor_set(x_1, 3, x_62); +lean_ctor_set(x_1, 2, x_57); +if (lean_is_scalar(x_64)) { + x_65 = lean_alloc_ctor(0, 2, 0); +} else { + x_65 = x_64; +} +lean_ctor_set(x_65, 0, x_1); +lean_ctor_set(x_65, 1, x_63); +return x_65; +} +} +else +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_66 = lean_ctor_get(x_1, 0); +x_67 = lean_ctor_get(x_1, 1); +x_68 = lean_ctor_get(x_1, 2); +x_69 = lean_ctor_get(x_1, 3); +lean_inc(x_69); +lean_inc(x_68); +lean_inc(x_67); +lean_inc(x_66); +lean_dec(x_1); +x_70 = lean_ctor_get(x_2, 0); +lean_inc(x_70); +x_71 = lean_ctor_get(x_2, 1); +lean_inc(x_71); +x_72 = lean_ctor_get(x_2, 2); +lean_inc(x_72); +x_73 = lean_ctor_get(x_2, 3); +lean_inc_ref(x_73); +x_74 = lean_ctor_get(x_2, 4); +lean_inc_ref(x_74); +if (lean_is_exclusive(x_2)) { + lean_ctor_release(x_2, 0); + lean_ctor_release(x_2, 1); + lean_ctor_release(x_2, 2); + lean_ctor_release(x_2, 3); + lean_ctor_release(x_2, 4); + x_75 = x_2; +} else { + lean_dec_ref(x_2); + x_75 = lean_box(0); +} +lean_inc(x_71); +x_76 = l_Lean_IR_LocalContext_addParams(x_71, x_67); +lean_inc_ref(x_74); +lean_inc_ref(x_73); +lean_inc(x_72); +lean_inc(x_70); +if (lean_is_scalar(x_75)) { + x_77 = lean_alloc_ctor(0, 5, 0); +} else { + x_77 = x_75; +} +lean_ctor_set(x_77, 0, x_70); +lean_ctor_set(x_77, 1, x_76); +lean_ctor_set(x_77, 2, x_72); +lean_ctor_set(x_77, 3, x_73); +lean_ctor_set(x_77, 4, x_74); +x_78 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_68, x_77, x_3); +x_79 = lean_ctor_get(x_78, 0); +lean_inc(x_79); +x_80 = lean_ctor_get(x_78, 1); +lean_inc(x_80); +lean_dec_ref(x_78); +lean_inc(x_79); +lean_inc_ref(x_67); +lean_inc(x_66); +x_81 = l_Lean_IR_LocalContext_addJP(x_71, x_66, x_67, x_79); +x_82 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_82, 0, x_70); +lean_ctor_set(x_82, 1, x_81); +lean_ctor_set(x_82, 2, x_72); +lean_ctor_set(x_82, 3, x_73); +lean_ctor_set(x_82, 4, x_74); +x_83 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_69, x_82, x_80); +x_84 = lean_ctor_get(x_83, 0); +lean_inc(x_84); +x_85 = lean_ctor_get(x_83, 1); +lean_inc(x_85); +if (lean_is_exclusive(x_83)) { + lean_ctor_release(x_83, 0); + lean_ctor_release(x_83, 1); + x_86 = x_83; +} else { + lean_dec_ref(x_83); + x_86 = lean_box(0); +} +x_87 = lean_alloc_ctor(1, 4, 0); +lean_ctor_set(x_87, 0, x_66); +lean_ctor_set(x_87, 1, x_67); +lean_ctor_set(x_87, 2, x_79); +lean_ctor_set(x_87, 3, x_84); +if (lean_is_scalar(x_86)) { + x_88 = lean_alloc_ctor(0, 2, 0); +} else { + x_88 = x_86; +} +lean_ctor_set(x_88, 0, x_87); +lean_ctor_set(x_88, 1, x_85); +return x_88; +} +} +case 4: +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; uint8_t x_101; uint8_t x_102; +x_89 = lean_ctor_get(x_1, 0); +lean_inc(x_89); +x_90 = lean_ctor_get(x_1, 1); +lean_inc(x_90); +x_91 = lean_ctor_get(x_1, 2); +lean_inc(x_91); +x_92 = lean_ctor_get(x_1, 3); +lean_inc(x_92); +x_93 = lean_ctor_get(x_1, 4); +lean_inc(x_93); +lean_dec_ref(x_1); +lean_inc_ref(x_2); +x_94 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_93, x_2, x_3); +x_95 = lean_ctor_get(x_94, 0); +lean_inc(x_95); +x_96 = lean_ctor_get(x_94, 1); +lean_inc(x_96); +lean_dec_ref(x_94); +x_97 = l_Lean_IR_ExplicitBoxing_getVarType(x_92, x_2, x_96); +x_98 = lean_ctor_get(x_97, 0); +lean_inc(x_98); +x_99 = lean_ctor_get(x_97, 1); +lean_inc(x_99); +lean_dec_ref(x_97); +x_100 = lean_box(5); +x_101 = 0; +x_102 = l_Lean_IR_IRType_compatibleWith(x_98, x_100, x_101); +if (x_102 == 0) +{ +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; uint8_t x_110; +x_103 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(x_99); +x_104 = lean_ctor_get(x_103, 0); +lean_inc(x_104); +x_105 = lean_ctor_get(x_103, 1); +lean_inc(x_105); +lean_dec_ref(x_103); +lean_inc_ref(x_2); +x_106 = l_Lean_IR_ExplicitBoxing_mkCast(x_92, x_98, x_100, x_2, x_105); +x_107 = lean_ctor_get(x_106, 0); +lean_inc(x_107); +x_108 = lean_ctor_get(x_106, 1); +lean_inc(x_108); +lean_dec_ref(x_106); +lean_inc(x_104); +x_109 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__0(x_89, x_90, x_91, x_95, x_104, x_2, x_108); +lean_dec_ref(x_2); +x_110 = !lean_is_exclusive(x_109); +if (x_110 == 0) +{ +lean_object* x_111; lean_object* x_112; +x_111 = lean_ctor_get(x_109, 0); +x_112 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_112, 0, x_104); +lean_ctor_set(x_112, 1, x_100); +lean_ctor_set(x_112, 2, x_107); +lean_ctor_set(x_112, 3, x_111); +lean_ctor_set(x_109, 0, x_112); +return x_109; +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_113 = lean_ctor_get(x_109, 0); +x_114 = lean_ctor_get(x_109, 1); +lean_inc(x_114); +lean_inc(x_113); +lean_dec(x_109); +x_115 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_115, 0, x_104); +lean_ctor_set(x_115, 1, x_100); +lean_ctor_set(x_115, 2, x_107); +lean_ctor_set(x_115, 3, x_113); +x_116 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_116, 0, x_115); +lean_ctor_set(x_116, 1, x_114); +return x_116; +} +} +else +{ +lean_object* x_117; +lean_dec(x_98); +x_117 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__0(x_89, x_90, x_91, x_95, x_92, x_2, x_99); +lean_dec_ref(x_2); +return x_117; +} +} +case 5: +{ +lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; uint8_t x_131; uint8_t x_132; +x_118 = lean_ctor_get(x_1, 0); +lean_inc(x_118); +x_119 = lean_ctor_get(x_1, 1); +lean_inc(x_119); +x_120 = lean_ctor_get(x_1, 2); +lean_inc(x_120); +x_121 = lean_ctor_get(x_1, 3); +lean_inc(x_121); +x_122 = lean_ctor_get(x_1, 4); +lean_inc(x_122); +x_123 = lean_ctor_get(x_1, 5); +lean_inc(x_123); +x_124 = lean_ctor_get(x_1, 6); +lean_inc(x_124); +lean_dec_ref(x_1); +lean_inc_ref(x_2); +x_125 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_124, x_2, x_3); +x_126 = lean_ctor_get(x_125, 0); +lean_inc(x_126); +x_127 = lean_ctor_get(x_125, 1); +lean_inc(x_127); +lean_dec_ref(x_125); +x_128 = l_Lean_IR_ExplicitBoxing_getVarType(x_122, x_2, x_127); +x_129 = lean_ctor_get(x_128, 0); +lean_inc(x_129); +x_130 = lean_ctor_get(x_128, 1); +lean_inc(x_130); +lean_dec_ref(x_128); +x_131 = 0; +x_132 = l_Lean_IR_IRType_compatibleWith(x_129, x_123, x_131); +if (x_132 == 0) +{ +lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; uint8_t x_140; +x_133 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(x_130); +x_134 = lean_ctor_get(x_133, 0); +lean_inc(x_134); +x_135 = lean_ctor_get(x_133, 1); +lean_inc(x_135); +lean_dec_ref(x_133); +lean_inc_ref(x_2); +lean_inc(x_123); +x_136 = l_Lean_IR_ExplicitBoxing_mkCast(x_122, x_129, x_123, x_2, x_135); +x_137 = lean_ctor_get(x_136, 0); +lean_inc(x_137); +x_138 = lean_ctor_get(x_136, 1); +lean_inc(x_138); +lean_dec_ref(x_136); +lean_inc(x_134); +lean_inc(x_123); +x_139 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__1(x_118, x_119, x_120, x_121, x_123, x_126, x_134, x_2, x_138); +lean_dec_ref(x_2); +x_140 = !lean_is_exclusive(x_139); +if (x_140 == 0) +{ +lean_object* x_141; lean_object* x_142; +x_141 = lean_ctor_get(x_139, 0); +x_142 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_142, 0, x_134); +lean_ctor_set(x_142, 1, x_123); +lean_ctor_set(x_142, 2, x_137); +lean_ctor_set(x_142, 3, x_141); +lean_ctor_set(x_139, 0, x_142); +return x_139; +} +else +{ +lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; +x_143 = lean_ctor_get(x_139, 0); +x_144 = lean_ctor_get(x_139, 1); +lean_inc(x_144); +lean_inc(x_143); +lean_dec(x_139); +x_145 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_145, 0, x_134); +lean_ctor_set(x_145, 1, x_123); +lean_ctor_set(x_145, 2, x_137); +lean_ctor_set(x_145, 3, x_143); +x_146 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_146, 0, x_145); +lean_ctor_set(x_146, 1, x_144); +return x_146; +} +} +else +{ +lean_object* x_147; +lean_dec(x_129); +x_147 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__1(x_118, x_119, x_120, x_121, x_123, x_126, x_122, x_2, x_130); +lean_dec_ref(x_2); +return x_147; +} +} +case 9: +{ +uint8_t x_148; +x_148 = !lean_is_exclusive(x_1); +if (x_148 == 0) +{ +lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; size_t x_153; size_t x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; uint8_t x_158; +x_149 = lean_ctor_get(x_1, 0); +x_150 = lean_ctor_get(x_1, 1); +x_151 = lean_ctor_get(x_1, 2); +x_152 = lean_ctor_get(x_1, 3); +x_153 = lean_array_size(x_152); +x_154 = 0; +lean_inc_ref(x_2); +x_155 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_visitFnBody_spec__0(x_153, x_154, x_152, x_2, x_3); +x_156 = lean_ctor_get(x_155, 0); +lean_inc(x_156); +x_157 = lean_ctor_get(x_155, 1); +lean_inc(x_157); +lean_dec_ref(x_155); +x_158 = l_Lean_IR_IRType_isObj(x_151); +if (x_158 == 0) +{ +lean_object* x_159; lean_object* x_160; lean_object* x_161; uint8_t x_162; +x_159 = l_Lean_IR_ExplicitBoxing_getVarType(x_150, x_2, x_157); +x_160 = lean_ctor_get(x_159, 0); +lean_inc(x_160); +x_161 = lean_ctor_get(x_159, 1); +lean_inc(x_161); +lean_dec_ref(x_159); +x_162 = l_Lean_IR_IRType_compatibleWith(x_160, x_151, x_158); +if (x_162 == 0) +{ +lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; uint8_t x_170; +x_163 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(x_161); +x_164 = lean_ctor_get(x_163, 0); +lean_inc(x_164); +x_165 = lean_ctor_get(x_163, 1); +lean_inc(x_165); +lean_dec_ref(x_163); +lean_inc_ref(x_2); +lean_inc(x_151); +x_166 = l_Lean_IR_ExplicitBoxing_mkCast(x_150, x_160, x_151, x_2, x_165); +x_167 = lean_ctor_get(x_166, 0); +lean_inc(x_167); +x_168 = lean_ctor_get(x_166, 1); +lean_inc(x_168); +lean_dec_ref(x_166); +lean_inc(x_164); +lean_inc(x_151); +x_169 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__2(x_149, x_151, x_156, x_164, x_2, x_168); +lean_dec_ref(x_2); +x_170 = !lean_is_exclusive(x_169); +if (x_170 == 0) +{ +lean_object* x_171; +x_171 = lean_ctor_get(x_169, 0); +lean_ctor_set_tag(x_1, 0); +lean_ctor_set(x_1, 3, x_171); +lean_ctor_set(x_1, 2, x_167); +lean_ctor_set(x_1, 1, x_151); +lean_ctor_set(x_1, 0, x_164); +lean_ctor_set(x_169, 0, x_1); +return x_169; +} +else +{ +lean_object* x_172; lean_object* x_173; lean_object* x_174; +x_172 = lean_ctor_get(x_169, 0); +x_173 = lean_ctor_get(x_169, 1); +lean_inc(x_173); +lean_inc(x_172); +lean_dec(x_169); +lean_ctor_set_tag(x_1, 0); +lean_ctor_set(x_1, 3, x_172); +lean_ctor_set(x_1, 2, x_167); +lean_ctor_set(x_1, 1, x_151); +lean_ctor_set(x_1, 0, x_164); +x_174 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_174, 0, x_1); +lean_ctor_set(x_174, 1, x_173); +return x_174; +} +} +else +{ +lean_object* x_175; +lean_dec(x_160); +lean_free_object(x_1); +x_175 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__2(x_149, x_151, x_156, x_150, x_2, x_161); +lean_dec_ref(x_2); +return x_175; +} +} +else +{ +lean_object* x_176; uint8_t x_177; +lean_dec(x_151); +x_176 = l_Lean_IR_ExplicitBoxing_getVarType(x_150, x_2, x_157); +lean_dec_ref(x_2); +x_177 = !lean_is_exclusive(x_176); +if (x_177 == 0) +{ +lean_object* x_178; +x_178 = lean_ctor_get(x_176, 0); +lean_ctor_set(x_1, 3, x_156); +lean_ctor_set(x_1, 2, x_178); +lean_ctor_set(x_176, 0, x_1); +return x_176; +} +else +{ +lean_object* x_179; lean_object* x_180; lean_object* x_181; +x_179 = lean_ctor_get(x_176, 0); +x_180 = lean_ctor_get(x_176, 1); +lean_inc(x_180); +lean_inc(x_179); +lean_dec(x_176); +lean_ctor_set(x_1, 3, x_156); +lean_ctor_set(x_1, 2, x_179); +x_181 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_181, 0, x_1); +lean_ctor_set(x_181, 1, x_180); +return x_181; +} +} +} +else +{ +lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; size_t x_186; size_t x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; uint8_t x_191; +x_182 = lean_ctor_get(x_1, 0); +x_183 = lean_ctor_get(x_1, 1); +x_184 = lean_ctor_get(x_1, 2); +x_185 = lean_ctor_get(x_1, 3); +lean_inc(x_185); +lean_inc(x_184); +lean_inc(x_183); +lean_inc(x_182); +lean_dec(x_1); +x_186 = lean_array_size(x_185); +x_187 = 0; +lean_inc_ref(x_2); +x_188 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_visitFnBody_spec__0(x_186, x_187, x_185, x_2, x_3); +x_189 = lean_ctor_get(x_188, 0); +lean_inc(x_189); +x_190 = lean_ctor_get(x_188, 1); +lean_inc(x_190); +lean_dec_ref(x_188); +x_191 = l_Lean_IR_IRType_isObj(x_184); +if (x_191 == 0) +{ +lean_object* x_192; lean_object* x_193; lean_object* x_194; uint8_t x_195; +x_192 = l_Lean_IR_ExplicitBoxing_getVarType(x_183, x_2, x_190); +x_193 = lean_ctor_get(x_192, 0); +lean_inc(x_193); +x_194 = lean_ctor_get(x_192, 1); +lean_inc(x_194); +lean_dec_ref(x_192); +x_195 = l_Lean_IR_IRType_compatibleWith(x_193, x_184, x_191); +if (x_195 == 0) +{ +lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; +x_196 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(x_194); +x_197 = lean_ctor_get(x_196, 0); +lean_inc(x_197); +x_198 = lean_ctor_get(x_196, 1); +lean_inc(x_198); +lean_dec_ref(x_196); +lean_inc_ref(x_2); +lean_inc(x_184); +x_199 = l_Lean_IR_ExplicitBoxing_mkCast(x_183, x_193, x_184, x_2, x_198); +x_200 = lean_ctor_get(x_199, 0); +lean_inc(x_200); +x_201 = lean_ctor_get(x_199, 1); +lean_inc(x_201); +lean_dec_ref(x_199); +lean_inc(x_197); +lean_inc(x_184); +x_202 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__2(x_182, x_184, x_189, x_197, x_2, x_201); +lean_dec_ref(x_2); +x_203 = lean_ctor_get(x_202, 0); +lean_inc(x_203); +x_204 = lean_ctor_get(x_202, 1); +lean_inc(x_204); +if (lean_is_exclusive(x_202)) { + lean_ctor_release(x_202, 0); + lean_ctor_release(x_202, 1); + x_205 = x_202; +} else { + lean_dec_ref(x_202); + x_205 = lean_box(0); +} +x_206 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_206, 0, x_197); +lean_ctor_set(x_206, 1, x_184); +lean_ctor_set(x_206, 2, x_200); +lean_ctor_set(x_206, 3, x_203); +if (lean_is_scalar(x_205)) { + x_207 = lean_alloc_ctor(0, 2, 0); +} else { + x_207 = x_205; +} +lean_ctor_set(x_207, 0, x_206); +lean_ctor_set(x_207, 1, x_204); +return x_207; +} +else +{ +lean_object* x_208; +lean_dec(x_193); +x_208 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__2(x_182, x_184, x_189, x_183, x_2, x_194); +lean_dec_ref(x_2); +return x_208; +} +} +else +{ +lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; +lean_dec(x_184); +x_209 = l_Lean_IR_ExplicitBoxing_getVarType(x_183, x_2, x_190); +lean_dec_ref(x_2); +x_210 = lean_ctor_get(x_209, 0); +lean_inc(x_210); +x_211 = lean_ctor_get(x_209, 1); +lean_inc(x_211); +if (lean_is_exclusive(x_209)) { + lean_ctor_release(x_209, 0); + lean_ctor_release(x_209, 1); + x_212 = x_209; +} else { + lean_dec_ref(x_209); + x_212 = lean_box(0); +} +x_213 = lean_alloc_ctor(9, 4, 0); +lean_ctor_set(x_213, 0, x_182); +lean_ctor_set(x_213, 1, x_183); +lean_ctor_set(x_213, 2, x_210); +lean_ctor_set(x_213, 3, x_189); +if (lean_is_scalar(x_212)) { + x_214 = lean_alloc_ctor(0, 2, 0); +} else { + x_214 = x_212; +} +lean_ctor_set(x_214, 0, x_213); +lean_ctor_set(x_214, 1, x_211); +return x_214; +} +} +} +case 10: +{ +lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; +x_215 = lean_ctor_get(x_1, 0); +lean_inc(x_215); +lean_dec_ref(x_1); +x_216 = l_Lean_IR_ExplicitBoxing_getResultType(x_2, x_3); +x_217 = lean_ctor_get(x_216, 0); +lean_inc(x_217); +x_218 = lean_ctor_get(x_216, 1); +lean_inc(x_218); +lean_dec_ref(x_216); +x_219 = lean_alloc_closure((void*)(l_Lean_IR_ExplicitBoxing_visitFnBody___lam__3___boxed), 3, 0); +if (lean_obj_tag(x_215) == 0) +{ +lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; uint8_t x_224; uint8_t x_225; +x_220 = lean_ctor_get(x_215, 0); +lean_inc(x_220); +lean_dec_ref(x_215); +x_221 = l_Lean_IR_ExplicitBoxing_getVarType(x_220, x_2, x_218); +x_222 = lean_ctor_get(x_221, 0); +lean_inc(x_222); +x_223 = lean_ctor_get(x_221, 1); +lean_inc(x_223); +lean_dec_ref(x_221); +x_224 = 0; +x_225 = l_Lean_IR_IRType_compatibleWith(x_222, x_217, x_224); +if (x_225 == 0) +{ +lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; uint8_t x_233; +x_226 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(x_223); +x_227 = lean_ctor_get(x_226, 0); +lean_inc(x_227); +x_228 = lean_ctor_get(x_226, 1); +lean_inc(x_228); +lean_dec_ref(x_226); +lean_inc_ref(x_2); +lean_inc(x_217); +x_229 = l_Lean_IR_ExplicitBoxing_mkCast(x_220, x_222, x_217, x_2, x_228); +x_230 = lean_ctor_get(x_229, 0); +lean_inc(x_230); +x_231 = lean_ctor_get(x_229, 1); +lean_inc(x_231); +lean_dec_ref(x_229); +lean_inc(x_227); +x_232 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__4(x_219, x_227, x_2, x_231); +x_233 = !lean_is_exclusive(x_232); +if (x_233 == 0) +{ +lean_object* x_234; lean_object* x_235; +x_234 = lean_ctor_get(x_232, 0); +x_235 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_235, 0, x_227); +lean_ctor_set(x_235, 1, x_217); +lean_ctor_set(x_235, 2, x_230); +lean_ctor_set(x_235, 3, x_234); +lean_ctor_set(x_232, 0, x_235); +return x_232; +} +else +{ +lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; +x_236 = lean_ctor_get(x_232, 0); +x_237 = lean_ctor_get(x_232, 1); +lean_inc(x_237); +lean_inc(x_236); +lean_dec(x_232); +x_238 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_238, 0, x_227); +lean_ctor_set(x_238, 1, x_217); +lean_ctor_set(x_238, 2, x_230); +lean_ctor_set(x_238, 3, x_236); +x_239 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_239, 0, x_238); +lean_ctor_set(x_239, 1, x_237); +return x_239; +} +} +else +{ +lean_object* x_240; +lean_dec(x_222); +lean_dec(x_217); +x_240 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__4(x_219, x_220, x_2, x_223); +return x_240; +} +} +else +{ +lean_object* x_241; +lean_dec_ref(x_219); +lean_dec(x_217); +x_241 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__3(x_215, x_2, x_218); +lean_dec_ref(x_2); +return x_241; +} +} +case 11: +{ +uint8_t x_242; +x_242 = !lean_is_exclusive(x_1); +if (x_242 == 0) +{ +lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; uint8_t x_253; +x_243 = lean_ctor_get(x_1, 0); +x_244 = lean_ctor_get(x_1, 1); +x_245 = l_Lean_IR_ExplicitBoxing_getJPParams(x_243, x_2, x_3); +x_246 = lean_ctor_get(x_245, 0); +lean_inc(x_246); +x_247 = lean_ctor_get(x_245, 1); +lean_inc(x_247); +lean_dec_ref(x_245); +x_248 = l_Lean_IR_ExplicitBoxing_mkBoxedVersionAux___closed__0; +x_249 = lean_alloc_closure((void*)(l_Lean_IR_ExplicitBoxing_visitFnBody___lam__5___boxed), 3, 2); +lean_closure_set(x_249, 0, x_248); +lean_closure_set(x_249, 1, x_246); +x_250 = l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux(x_244, x_249, x_2, x_247); +lean_dec_ref(x_244); +x_251 = lean_ctor_get(x_250, 0); +lean_inc(x_251); +x_252 = lean_ctor_get(x_250, 1); +lean_inc(x_252); +lean_dec_ref(x_250); +x_253 = !lean_is_exclusive(x_251); +if (x_253 == 0) +{ +lean_object* x_254; lean_object* x_255; lean_object* x_256; +x_254 = lean_ctor_get(x_251, 0); +x_255 = lean_ctor_get(x_251, 1); +lean_ctor_set(x_1, 1, x_254); +x_256 = l_Lean_IR_reshape(x_255, x_1); +lean_ctor_set(x_251, 1, x_252); +lean_ctor_set(x_251, 0, x_256); +return x_251; +} +else +{ +lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; +x_257 = lean_ctor_get(x_251, 0); +x_258 = lean_ctor_get(x_251, 1); +lean_inc(x_258); +lean_inc(x_257); +lean_dec(x_251); +lean_ctor_set(x_1, 1, x_257); +x_259 = l_Lean_IR_reshape(x_258, x_1); +x_260 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_260, 0, x_259); +lean_ctor_set(x_260, 1, x_252); +return x_260; +} +} +else +{ +lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; +x_261 = lean_ctor_get(x_1, 0); +x_262 = lean_ctor_get(x_1, 1); +lean_inc(x_262); +lean_inc(x_261); +lean_dec(x_1); +x_263 = l_Lean_IR_ExplicitBoxing_getJPParams(x_261, x_2, x_3); +x_264 = lean_ctor_get(x_263, 0); +lean_inc(x_264); +x_265 = lean_ctor_get(x_263, 1); +lean_inc(x_265); +lean_dec_ref(x_263); +x_266 = l_Lean_IR_ExplicitBoxing_mkBoxedVersionAux___closed__0; +x_267 = lean_alloc_closure((void*)(l_Lean_IR_ExplicitBoxing_visitFnBody___lam__5___boxed), 3, 2); +lean_closure_set(x_267, 0, x_266); +lean_closure_set(x_267, 1, x_264); +x_268 = l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux(x_262, x_267, x_2, x_265); +lean_dec_ref(x_262); +x_269 = lean_ctor_get(x_268, 0); +lean_inc(x_269); +x_270 = lean_ctor_get(x_268, 1); +lean_inc(x_270); +lean_dec_ref(x_268); +x_271 = lean_ctor_get(x_269, 0); +lean_inc(x_271); +x_272 = lean_ctor_get(x_269, 1); +lean_inc(x_272); +if (lean_is_exclusive(x_269)) { + lean_ctor_release(x_269, 0); + lean_ctor_release(x_269, 1); + x_273 = x_269; +} else { + lean_dec_ref(x_269); + x_273 = lean_box(0); +} +x_274 = lean_alloc_ctor(11, 2, 0); +lean_ctor_set(x_274, 0, x_261); +lean_ctor_set(x_274, 1, x_271); +x_275 = l_Lean_IR_reshape(x_272, x_274); +if (lean_is_scalar(x_273)) { + x_276 = lean_alloc_ctor(0, 2, 0); +} else { + x_276 = x_273; +} +lean_ctor_set(x_276, 0, x_275); +lean_ctor_set(x_276, 1, x_270); +return x_276; +} +} +default: +{ +lean_object* x_277; +lean_dec_ref(x_2); +x_277 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_277, 0, x_1); +lean_ctor_set(x_277, 1, x_3); +return x_277; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_visitFnBody_spec__0(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_usize_dec_lt(x_2, x_1); +if (x_6 == 0) +{ +lean_object* x_7; +lean_dec_ref(x_4); +x_7 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_7, 0, x_3); +lean_ctor_set(x_7, 1, x_5); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_8 = lean_array_uget(x_3, x_2); +x_9 = lean_unsigned_to_nat(0u); +x_10 = lean_array_uset(x_3, x_2, x_9); +if (lean_obj_tag(x_8) == 0) +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_8); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_19 = lean_ctor_get(x_8, 1); +lean_inc_ref(x_4); +x_20 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_19, x_4, x_5); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec_ref(x_20); +lean_ctor_set(x_8, 1, x_21); +x_11 = x_8; +x_12 = x_22; +goto block_17; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_23 = lean_ctor_get(x_8, 0); +x_24 = lean_ctor_get(x_8, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_8); +lean_inc_ref(x_4); +x_25 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_24, x_4, x_5); +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +lean_dec_ref(x_25); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_23); +lean_ctor_set(x_28, 1, x_26); +x_11 = x_28; +x_12 = x_27; +goto block_17; +} +} +else +{ +uint8_t x_29; +x_29 = !lean_is_exclusive(x_8); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_8, 0); +lean_inc_ref(x_4); +x_31 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_30, x_4, x_5); +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec_ref(x_31); +lean_ctor_set(x_8, 0, x_32); +x_11 = x_8; +x_12 = x_33; +goto block_17; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_34 = lean_ctor_get(x_8, 0); +lean_inc(x_34); +lean_dec(x_8); +lean_inc_ref(x_4); +x_35 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_34, x_4, x_5); +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec_ref(x_35); +x_38 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_38, 0, x_36); +x_11 = x_38; +x_12 = x_37; +goto block_17; +} +} +block_17: +{ +size_t x_13; size_t x_14; lean_object* x_15; +x_13 = 1; +x_14 = lean_usize_add(x_2, x_13); +x_15 = lean_array_uset(x_10, x_2, x_11); +x_2 = x_14; +x_3 = x_15; +x_5 = x_12; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__0___boxed(lean_object* x_1, lean_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_IR_ExplicitBoxing_visitFnBody___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec_ref(x_6); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec_ref(x_8); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__2___boxed(lean_object* x_1, 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_IR_ExplicitBoxing_visitFnBody___lam__2(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec_ref(x_5); +return x_7; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_visitFnBody_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_7 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_8 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_visitFnBody_spec__0(x_6, x_7, x_3, x_4, x_5); +return x_8; +} +} +LEAN_EXPORT uint8_t l_Lean_IR_ExplicitBoxing_BoxParams_isExpensiveScalar(lean_object* x_1) { +_start: +{ +uint8_t x_2; +x_2 = l_Lean_IR_IRType_isScalarOrStruct(x_1); +if (x_2 == 0) +{ +return x_2; +} +else +{ +switch (lean_obj_tag(x_1)) { +case 1: +{ +uint8_t x_3; +x_3 = 0; +return x_3; +} +case 2: +{ +uint8_t x_4; +x_4 = 0; +return x_4; +} +default: +{ +return x_2; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_BoxParams_isExpensiveScalar___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_IR_ExplicitBoxing_BoxParams_isExpensiveScalar(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT uint8_t l_Std_DTreeMap_Internal_Impl_contains___at___00__private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go_spec__0___redArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_3 = lean_ctor_get(x_2, 1); +x_4 = lean_ctor_get(x_2, 3); +x_5 = lean_ctor_get(x_2, 4); +x_6 = lean_nat_dec_lt(x_1, x_3); +if (x_6 == 0) +{ +uint8_t x_7; +x_7 = lean_nat_dec_eq(x_1, x_3); +if (x_7 == 0) +{ +x_2 = x_5; +goto _start; +} +else +{ +return x_7; +} +} +else +{ +x_2 = x_4; +goto _start; +} +} +else +{ +uint8_t x_10; +x_10 = 0; +return x_10; +} +} +} +LEAN_EXPORT uint8_t l_Std_DTreeMap_Internal_Impl_contains___at___00__private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = l_Std_DTreeMap_Internal_Impl_contains___at___00__private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go_spec__0___redArg(x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_insert___at___00__private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go_spec__1___redArg(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_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_3, 1); +lean_inc(x_5); +x_6 = lean_ctor_get(x_3, 2); +lean_inc(x_6); +x_7 = lean_ctor_get(x_3, 3); +lean_inc(x_7); +x_8 = lean_ctor_get(x_3, 4); +lean_inc(x_8); +if (lean_is_exclusive(x_3)) { + lean_ctor_release(x_3, 0); + lean_ctor_release(x_3, 1); + lean_ctor_release(x_3, 2); + lean_ctor_release(x_3, 3); + lean_ctor_release(x_3, 4); + x_9 = x_3; +} else { + lean_dec_ref(x_3); + x_9 = lean_box(0); +} +x_10 = lean_nat_dec_lt(x_1, x_5); +if (x_10 == 0) +{ +uint8_t x_11; +x_11 = lean_nat_dec_eq(x_1, x_5); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_4); +x_12 = l_Std_DTreeMap_Internal_Impl_insert___at___00__private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go_spec__1___redArg(x_1, x_2, x_8); +x_13 = lean_unsigned_to_nat(1u); +if (lean_obj_tag(x_7) == 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; uint8_t x_22; +x_14 = lean_ctor_get(x_7, 0); +x_15 = lean_ctor_get(x_12, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_12, 1); +lean_inc(x_16); +x_17 = lean_ctor_get(x_12, 2); +lean_inc(x_17); +x_18 = lean_ctor_get(x_12, 3); +lean_inc(x_18); +x_19 = lean_ctor_get(x_12, 4); +lean_inc(x_19); +x_20 = lean_unsigned_to_nat(3u); +x_21 = lean_nat_mul(x_20, x_14); +x_22 = lean_nat_dec_lt(x_21, x_15); +lean_dec(x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +x_23 = lean_nat_add(x_13, x_14); +x_24 = lean_nat_add(x_23, x_15); +lean_dec(x_15); +lean_dec(x_23); +if (lean_is_scalar(x_9)) { + x_25 = lean_alloc_ctor(0, 5, 0); +} else { + x_25 = x_9; +} +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_5); +lean_ctor_set(x_25, 2, x_6); +lean_ctor_set(x_25, 3, x_7); +lean_ctor_set(x_25, 4, x_12); +return x_25; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +if (lean_is_exclusive(x_12)) { + lean_ctor_release(x_12, 0); + lean_ctor_release(x_12, 1); + lean_ctor_release(x_12, 2); + lean_ctor_release(x_12, 3); + lean_ctor_release(x_12, 4); + x_26 = x_12; +} else { + lean_dec_ref(x_12); + x_26 = lean_box(0); +} +x_27 = lean_ctor_get(x_18, 0); +x_28 = lean_ctor_get(x_18, 1); +x_29 = lean_ctor_get(x_18, 2); +x_30 = lean_ctor_get(x_18, 3); +x_31 = lean_ctor_get(x_18, 4); +x_32 = lean_ctor_get(x_19, 0); +x_33 = lean_unsigned_to_nat(2u); +x_34 = lean_nat_mul(x_33, x_32); +x_35 = lean_nat_dec_lt(x_27, x_34); +lean_dec(x_34); +if (x_35 == 0) +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_46; +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +if (lean_is_exclusive(x_18)) { + lean_ctor_release(x_18, 0); + lean_ctor_release(x_18, 1); + lean_ctor_release(x_18, 2); + lean_ctor_release(x_18, 3); + lean_ctor_release(x_18, 4); + x_36 = x_18; +} else { + lean_dec_ref(x_18); + x_36 = lean_box(0); +} +x_37 = lean_nat_add(x_13, x_14); +x_38 = lean_nat_add(x_37, x_15); +lean_dec(x_15); +if (lean_obj_tag(x_30) == 0) +{ +lean_object* x_53; +x_53 = lean_ctor_get(x_30, 0); +lean_inc(x_53); +x_46 = x_53; +goto block_52; +} +else +{ +lean_object* x_54; +x_54 = lean_unsigned_to_nat(0u); +x_46 = x_54; +goto block_52; +} +block_45: +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_nat_add(x_39, x_41); +lean_dec(x_41); +lean_dec(x_39); +if (lean_is_scalar(x_36)) { + x_43 = lean_alloc_ctor(0, 5, 0); +} else { + x_43 = x_36; +} +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_16); +lean_ctor_set(x_43, 2, x_17); +lean_ctor_set(x_43, 3, x_31); +lean_ctor_set(x_43, 4, x_19); +if (lean_is_scalar(x_26)) { + x_44 = lean_alloc_ctor(0, 5, 0); +} else { + x_44 = x_26; +} +lean_ctor_set(x_44, 0, x_38); +lean_ctor_set(x_44, 1, x_28); +lean_ctor_set(x_44, 2, x_29); +lean_ctor_set(x_44, 3, x_40); +lean_ctor_set(x_44, 4, x_43); +return x_44; +} +block_52: +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_nat_add(x_37, x_46); +lean_dec(x_46); +lean_dec(x_37); +if (lean_is_scalar(x_9)) { + x_48 = lean_alloc_ctor(0, 5, 0); +} else { + x_48 = x_9; +} +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_5); +lean_ctor_set(x_48, 2, x_6); +lean_ctor_set(x_48, 3, x_7); +lean_ctor_set(x_48, 4, x_30); +x_49 = lean_nat_add(x_13, x_32); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_50; +x_50 = lean_ctor_get(x_31, 0); +lean_inc(x_50); +x_39 = x_49; +x_40 = x_48; +x_41 = x_50; +goto block_45; +} +else +{ +lean_object* x_51; +x_51 = lean_unsigned_to_nat(0u); +x_39 = x_49; +x_40 = x_48; +x_41 = x_51; +goto block_45; +} +} +} +else +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; +lean_dec(x_9); +x_55 = lean_nat_add(x_13, x_14); +x_56 = lean_nat_add(x_55, x_15); +lean_dec(x_15); +x_57 = lean_nat_add(x_55, x_27); +lean_dec(x_55); +lean_inc_ref(x_7); +if (lean_is_scalar(x_26)) { + x_58 = lean_alloc_ctor(0, 5, 0); +} else { + x_58 = x_26; +} +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_5); +lean_ctor_set(x_58, 2, x_6); +lean_ctor_set(x_58, 3, x_7); +lean_ctor_set(x_58, 4, x_18); +x_59 = !lean_is_exclusive(x_7); +if (x_59 == 0) +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_60 = lean_ctor_get(x_7, 4); +lean_dec(x_60); +x_61 = lean_ctor_get(x_7, 3); +lean_dec(x_61); +x_62 = lean_ctor_get(x_7, 2); +lean_dec(x_62); +x_63 = lean_ctor_get(x_7, 1); +lean_dec(x_63); +x_64 = lean_ctor_get(x_7, 0); +lean_dec(x_64); +lean_ctor_set(x_7, 4, x_19); +lean_ctor_set(x_7, 3, x_58); +lean_ctor_set(x_7, 2, x_17); +lean_ctor_set(x_7, 1, x_16); +lean_ctor_set(x_7, 0, x_56); +return x_7; +} +else +{ +lean_object* x_65; +lean_dec(x_7); +x_65 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_65, 0, x_56); +lean_ctor_set(x_65, 1, x_16); +lean_ctor_set(x_65, 2, x_17); +lean_ctor_set(x_65, 3, x_58); +lean_ctor_set(x_65, 4, x_19); +return x_65; +} +} +} +} +else +{ +lean_object* x_66; +x_66 = lean_ctor_get(x_12, 3); +lean_inc(x_66); +if (lean_obj_tag(x_66) == 0) +{ +uint8_t x_67; +x_67 = !lean_is_exclusive(x_12); +if (x_67 == 0) +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; +x_68 = lean_ctor_get(x_12, 4); +x_69 = lean_ctor_get(x_12, 3); +lean_dec(x_69); +x_70 = lean_ctor_get(x_12, 0); +lean_dec(x_70); +x_71 = !lean_is_exclusive(x_66); +if (x_71 == 0) +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_72 = lean_ctor_get(x_66, 1); +x_73 = lean_ctor_get(x_66, 2); +x_74 = lean_ctor_get(x_66, 4); +lean_dec(x_74); +x_75 = lean_ctor_get(x_66, 3); +lean_dec(x_75); +x_76 = lean_ctor_get(x_66, 0); +lean_dec(x_76); +x_77 = lean_unsigned_to_nat(3u); +lean_inc_n(x_68, 2); +lean_ctor_set(x_66, 4, x_68); +lean_ctor_set(x_66, 3, x_68); +lean_ctor_set(x_66, 2, x_6); +lean_ctor_set(x_66, 1, x_5); +lean_ctor_set(x_66, 0, x_13); +lean_inc(x_68); +lean_ctor_set(x_12, 3, x_68); +lean_ctor_set(x_12, 0, x_13); +if (lean_is_scalar(x_9)) { + x_78 = lean_alloc_ctor(0, 5, 0); +} else { + x_78 = x_9; +} +lean_ctor_set(x_78, 0, x_77); +lean_ctor_set(x_78, 1, x_72); +lean_ctor_set(x_78, 2, x_73); +lean_ctor_set(x_78, 3, x_66); +lean_ctor_set(x_78, 4, x_12); +return x_78; +} +else +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_79 = lean_ctor_get(x_66, 1); +x_80 = lean_ctor_get(x_66, 2); +lean_inc(x_80); +lean_inc(x_79); +lean_dec(x_66); +x_81 = lean_unsigned_to_nat(3u); +lean_inc_n(x_68, 2); +x_82 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_82, 0, x_13); +lean_ctor_set(x_82, 1, x_5); +lean_ctor_set(x_82, 2, x_6); +lean_ctor_set(x_82, 3, x_68); +lean_ctor_set(x_82, 4, x_68); +lean_inc(x_68); +lean_ctor_set(x_12, 3, x_68); +lean_ctor_set(x_12, 0, x_13); +if (lean_is_scalar(x_9)) { + x_83 = lean_alloc_ctor(0, 5, 0); +} else { + x_83 = x_9; +} +lean_ctor_set(x_83, 0, x_81); +lean_ctor_set(x_83, 1, x_79); +lean_ctor_set(x_83, 2, x_80); +lean_ctor_set(x_83, 3, x_82); +lean_ctor_set(x_83, 4, x_12); +return x_83; +} +} +else +{ +lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_84 = lean_ctor_get(x_12, 4); +x_85 = lean_ctor_get(x_12, 1); +x_86 = lean_ctor_get(x_12, 2); +lean_inc(x_84); +lean_inc(x_86); +lean_inc(x_85); +lean_dec(x_12); +x_87 = lean_ctor_get(x_66, 1); +lean_inc(x_87); +x_88 = lean_ctor_get(x_66, 2); +lean_inc(x_88); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + lean_ctor_release(x_66, 2); + lean_ctor_release(x_66, 3); + lean_ctor_release(x_66, 4); + x_89 = x_66; +} else { + lean_dec_ref(x_66); + x_89 = lean_box(0); +} +x_90 = lean_unsigned_to_nat(3u); +lean_inc_n(x_84, 2); +if (lean_is_scalar(x_89)) { + x_91 = lean_alloc_ctor(0, 5, 0); +} else { + x_91 = x_89; +} +lean_ctor_set(x_91, 0, x_13); +lean_ctor_set(x_91, 1, x_5); +lean_ctor_set(x_91, 2, x_6); +lean_ctor_set(x_91, 3, x_84); +lean_ctor_set(x_91, 4, x_84); +lean_inc(x_84); +x_92 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_92, 0, x_13); +lean_ctor_set(x_92, 1, x_85); +lean_ctor_set(x_92, 2, x_86); +lean_ctor_set(x_92, 3, x_84); +lean_ctor_set(x_92, 4, x_84); +if (lean_is_scalar(x_9)) { + x_93 = lean_alloc_ctor(0, 5, 0); +} else { + x_93 = x_9; +} +lean_ctor_set(x_93, 0, x_90); +lean_ctor_set(x_93, 1, x_87); +lean_ctor_set(x_93, 2, x_88); +lean_ctor_set(x_93, 3, x_91); +lean_ctor_set(x_93, 4, x_92); +return x_93; +} +} +else +{ +lean_object* x_94; +x_94 = lean_ctor_get(x_12, 4); +lean_inc(x_94); +if (lean_obj_tag(x_94) == 0) +{ +uint8_t x_95; +x_95 = !lean_is_exclusive(x_12); +if (x_95 == 0) +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_96 = lean_ctor_get(x_12, 1); +x_97 = lean_ctor_get(x_12, 2); +x_98 = lean_ctor_get(x_12, 4); +lean_dec(x_98); +x_99 = lean_ctor_get(x_12, 3); +lean_dec(x_99); +x_100 = lean_ctor_get(x_12, 0); +lean_dec(x_100); +x_101 = lean_unsigned_to_nat(3u); +lean_ctor_set(x_12, 4, x_66); +lean_ctor_set(x_12, 2, x_6); +lean_ctor_set(x_12, 1, x_5); +lean_ctor_set(x_12, 0, x_13); +if (lean_is_scalar(x_9)) { + x_102 = lean_alloc_ctor(0, 5, 0); +} else { + x_102 = x_9; +} +lean_ctor_set(x_102, 0, x_101); +lean_ctor_set(x_102, 1, x_96); +lean_ctor_set(x_102, 2, x_97); +lean_ctor_set(x_102, 3, x_12); +lean_ctor_set(x_102, 4, x_94); +return x_102; +} +else +{ +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_103 = lean_ctor_get(x_12, 1); +x_104 = lean_ctor_get(x_12, 2); +lean_inc(x_104); +lean_inc(x_103); +lean_dec(x_12); +x_105 = lean_unsigned_to_nat(3u); +x_106 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_106, 0, x_13); +lean_ctor_set(x_106, 1, x_5); +lean_ctor_set(x_106, 2, x_6); +lean_ctor_set(x_106, 3, x_66); +lean_ctor_set(x_106, 4, x_66); +if (lean_is_scalar(x_9)) { + x_107 = lean_alloc_ctor(0, 5, 0); +} else { + x_107 = x_9; +} +lean_ctor_set(x_107, 0, x_105); +lean_ctor_set(x_107, 1, x_103); +lean_ctor_set(x_107, 2, x_104); +lean_ctor_set(x_107, 3, x_106); +lean_ctor_set(x_107, 4, x_94); +return x_107; +} +} +else +{ +lean_object* x_108; lean_object* x_109; +x_108 = lean_unsigned_to_nat(2u); +if (lean_is_scalar(x_9)) { + x_109 = lean_alloc_ctor(0, 5, 0); +} else { + x_109 = x_9; +} +lean_ctor_set(x_109, 0, x_108); +lean_ctor_set(x_109, 1, x_5); +lean_ctor_set(x_109, 2, x_6); +lean_ctor_set(x_109, 3, x_94); +lean_ctor_set(x_109, 4, x_12); +return x_109; +} +} +} +} +else +{ +lean_object* x_110; +lean_dec(x_6); +lean_dec(x_5); +if (lean_is_scalar(x_9)) { + x_110 = lean_alloc_ctor(0, 5, 0); +} else { + x_110 = x_9; +} +lean_ctor_set(x_110, 0, x_4); +lean_ctor_set(x_110, 1, x_1); +lean_ctor_set(x_110, 2, x_2); +lean_ctor_set(x_110, 3, x_7); +lean_ctor_set(x_110, 4, x_8); +return x_110; +} +} +else +{ +lean_object* x_111; lean_object* x_112; +lean_dec(x_4); +x_111 = l_Std_DTreeMap_Internal_Impl_insert___at___00__private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go_spec__1___redArg(x_1, x_2, x_7); +x_112 = lean_unsigned_to_nat(1u); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; uint8_t x_121; +x_113 = lean_ctor_get(x_8, 0); +x_114 = lean_ctor_get(x_111, 0); +lean_inc(x_114); +x_115 = lean_ctor_get(x_111, 1); +lean_inc(x_115); +x_116 = lean_ctor_get(x_111, 2); +lean_inc(x_116); +x_117 = lean_ctor_get(x_111, 3); +lean_inc(x_117); +x_118 = lean_ctor_get(x_111, 4); +lean_inc(x_118); +x_119 = lean_unsigned_to_nat(3u); +x_120 = lean_nat_mul(x_119, x_113); +x_121 = lean_nat_dec_lt(x_120, x_114); +lean_dec(x_120); +if (x_121 == 0) +{ +lean_object* x_122; lean_object* x_123; lean_object* x_124; +lean_dec(x_118); +lean_dec(x_117); +lean_dec(x_116); +lean_dec(x_115); +x_122 = lean_nat_add(x_112, x_114); +lean_dec(x_114); +x_123 = lean_nat_add(x_122, x_113); +lean_dec(x_122); +if (lean_is_scalar(x_9)) { + x_124 = lean_alloc_ctor(0, 5, 0); +} else { + x_124 = x_9; +} +lean_ctor_set(x_124, 0, x_123); +lean_ctor_set(x_124, 1, x_5); +lean_ctor_set(x_124, 2, x_6); +lean_ctor_set(x_124, 3, x_111); +lean_ctor_set(x_124, 4, x_8); +return x_124; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; uint8_t x_134; +if (lean_is_exclusive(x_111)) { + lean_ctor_release(x_111, 0); + lean_ctor_release(x_111, 1); + lean_ctor_release(x_111, 2); + lean_ctor_release(x_111, 3); + lean_ctor_release(x_111, 4); + x_125 = x_111; +} else { + lean_dec_ref(x_111); + x_125 = lean_box(0); +} +x_126 = lean_ctor_get(x_117, 0); +x_127 = lean_ctor_get(x_118, 0); +x_128 = lean_ctor_get(x_118, 1); +x_129 = lean_ctor_get(x_118, 2); +x_130 = lean_ctor_get(x_118, 3); +x_131 = lean_ctor_get(x_118, 4); +x_132 = lean_unsigned_to_nat(2u); +x_133 = lean_nat_mul(x_132, x_126); +x_134 = lean_nat_dec_lt(x_127, x_133); +lean_dec(x_133); +if (x_134 == 0) +{ +lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_145; lean_object* x_146; +lean_inc(x_131); +lean_inc(x_130); +lean_inc(x_129); +lean_inc(x_128); +if (lean_is_exclusive(x_118)) { + lean_ctor_release(x_118, 0); + lean_ctor_release(x_118, 1); + lean_ctor_release(x_118, 2); + lean_ctor_release(x_118, 3); + lean_ctor_release(x_118, 4); + x_135 = x_118; +} else { + lean_dec_ref(x_118); + x_135 = lean_box(0); +} +x_136 = lean_nat_add(x_112, x_114); +lean_dec(x_114); +x_137 = lean_nat_add(x_136, x_113); +lean_dec(x_136); +x_145 = lean_nat_add(x_112, x_126); +if (lean_obj_tag(x_130) == 0) +{ +lean_object* x_153; +x_153 = lean_ctor_get(x_130, 0); +lean_inc(x_153); +x_146 = x_153; +goto block_152; +} +else +{ +lean_object* x_154; +x_154 = lean_unsigned_to_nat(0u); +x_146 = x_154; +goto block_152; +} +block_144: +{ +lean_object* x_141; lean_object* x_142; lean_object* x_143; +x_141 = lean_nat_add(x_139, x_140); +lean_dec(x_140); +lean_dec(x_139); +if (lean_is_scalar(x_135)) { + x_142 = lean_alloc_ctor(0, 5, 0); +} else { + x_142 = x_135; +} +lean_ctor_set(x_142, 0, x_141); +lean_ctor_set(x_142, 1, x_5); +lean_ctor_set(x_142, 2, x_6); +lean_ctor_set(x_142, 3, x_131); +lean_ctor_set(x_142, 4, x_8); +if (lean_is_scalar(x_125)) { + x_143 = lean_alloc_ctor(0, 5, 0); +} else { + x_143 = x_125; +} +lean_ctor_set(x_143, 0, x_137); +lean_ctor_set(x_143, 1, x_128); +lean_ctor_set(x_143, 2, x_129); +lean_ctor_set(x_143, 3, x_138); +lean_ctor_set(x_143, 4, x_142); +return x_143; +} +block_152: +{ +lean_object* x_147; lean_object* x_148; lean_object* x_149; +x_147 = lean_nat_add(x_145, x_146); +lean_dec(x_146); +lean_dec(x_145); +if (lean_is_scalar(x_9)) { + x_148 = lean_alloc_ctor(0, 5, 0); +} else { + x_148 = x_9; +} +lean_ctor_set(x_148, 0, x_147); +lean_ctor_set(x_148, 1, x_115); +lean_ctor_set(x_148, 2, x_116); +lean_ctor_set(x_148, 3, x_117); +lean_ctor_set(x_148, 4, x_130); +x_149 = lean_nat_add(x_112, x_113); +if (lean_obj_tag(x_131) == 0) +{ +lean_object* x_150; +x_150 = lean_ctor_get(x_131, 0); +lean_inc(x_150); +x_138 = x_148; +x_139 = x_149; +x_140 = x_150; +goto block_144; +} +else +{ +lean_object* x_151; +x_151 = lean_unsigned_to_nat(0u); +x_138 = x_148; +x_139 = x_149; +x_140 = x_151; +goto block_144; +} +} +} +else +{ +lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; uint8_t x_160; +lean_dec(x_9); +x_155 = lean_nat_add(x_112, x_114); +lean_dec(x_114); +x_156 = lean_nat_add(x_155, x_113); +lean_dec(x_155); +x_157 = lean_nat_add(x_112, x_113); +x_158 = lean_nat_add(x_157, x_127); +lean_dec(x_157); +lean_inc_ref(x_8); +if (lean_is_scalar(x_125)) { + x_159 = lean_alloc_ctor(0, 5, 0); +} else { + x_159 = x_125; +} +lean_ctor_set(x_159, 0, x_158); +lean_ctor_set(x_159, 1, x_5); +lean_ctor_set(x_159, 2, x_6); +lean_ctor_set(x_159, 3, x_118); +lean_ctor_set(x_159, 4, x_8); +x_160 = !lean_is_exclusive(x_8); +if (x_160 == 0) +{ +lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; +x_161 = lean_ctor_get(x_8, 4); +lean_dec(x_161); +x_162 = lean_ctor_get(x_8, 3); +lean_dec(x_162); +x_163 = lean_ctor_get(x_8, 2); +lean_dec(x_163); +x_164 = lean_ctor_get(x_8, 1); +lean_dec(x_164); +x_165 = lean_ctor_get(x_8, 0); +lean_dec(x_165); +lean_ctor_set(x_8, 4, x_159); +lean_ctor_set(x_8, 3, x_117); +lean_ctor_set(x_8, 2, x_116); +lean_ctor_set(x_8, 1, x_115); +lean_ctor_set(x_8, 0, x_156); +return x_8; +} +else +{ +lean_object* x_166; +lean_dec(x_8); +x_166 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_166, 0, x_156); +lean_ctor_set(x_166, 1, x_115); +lean_ctor_set(x_166, 2, x_116); +lean_ctor_set(x_166, 3, x_117); +lean_ctor_set(x_166, 4, x_159); +return x_166; +} +} +} +} +else +{ +lean_object* x_167; +x_167 = lean_ctor_get(x_111, 3); +lean_inc(x_167); +if (lean_obj_tag(x_167) == 0) +{ +uint8_t x_168; +x_168 = !lean_is_exclusive(x_111); +if (x_168 == 0) +{ +lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; +x_169 = lean_ctor_get(x_111, 4); +x_170 = lean_ctor_get(x_111, 1); +x_171 = lean_ctor_get(x_111, 2); +x_172 = lean_ctor_get(x_111, 3); +lean_dec(x_172); +x_173 = lean_ctor_get(x_111, 0); +lean_dec(x_173); +x_174 = lean_unsigned_to_nat(3u); +lean_inc(x_169); +lean_ctor_set(x_111, 3, x_169); +lean_ctor_set(x_111, 2, x_6); +lean_ctor_set(x_111, 1, x_5); +lean_ctor_set(x_111, 0, x_112); +if (lean_is_scalar(x_9)) { + x_175 = lean_alloc_ctor(0, 5, 0); +} else { + x_175 = x_9; +} +lean_ctor_set(x_175, 0, x_174); +lean_ctor_set(x_175, 1, x_170); +lean_ctor_set(x_175, 2, x_171); +lean_ctor_set(x_175, 3, x_167); +lean_ctor_set(x_175, 4, x_111); +return x_175; +} +else +{ +lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; +x_176 = lean_ctor_get(x_111, 4); +x_177 = lean_ctor_get(x_111, 1); +x_178 = lean_ctor_get(x_111, 2); +lean_inc(x_176); +lean_inc(x_178); +lean_inc(x_177); +lean_dec(x_111); +x_179 = lean_unsigned_to_nat(3u); +lean_inc(x_176); +x_180 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_180, 0, x_112); +lean_ctor_set(x_180, 1, x_5); +lean_ctor_set(x_180, 2, x_6); +lean_ctor_set(x_180, 3, x_176); +lean_ctor_set(x_180, 4, x_176); +if (lean_is_scalar(x_9)) { + x_181 = lean_alloc_ctor(0, 5, 0); +} else { + x_181 = x_9; +} +lean_ctor_set(x_181, 0, x_179); +lean_ctor_set(x_181, 1, x_177); +lean_ctor_set(x_181, 2, x_178); +lean_ctor_set(x_181, 3, x_167); +lean_ctor_set(x_181, 4, x_180); +return x_181; +} +} +else +{ +lean_object* x_182; +x_182 = lean_ctor_get(x_111, 4); +lean_inc(x_182); +if (lean_obj_tag(x_182) == 0) +{ +uint8_t x_183; +x_183 = !lean_is_exclusive(x_111); +if (x_183 == 0) +{ +lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; uint8_t x_189; +x_184 = lean_ctor_get(x_111, 1); +x_185 = lean_ctor_get(x_111, 2); +x_186 = lean_ctor_get(x_111, 4); +lean_dec(x_186); +x_187 = lean_ctor_get(x_111, 3); +lean_dec(x_187); +x_188 = lean_ctor_get(x_111, 0); +lean_dec(x_188); +x_189 = !lean_is_exclusive(x_182); +if (x_189 == 0) +{ +lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; +x_190 = lean_ctor_get(x_182, 1); +x_191 = lean_ctor_get(x_182, 2); +x_192 = lean_ctor_get(x_182, 4); +lean_dec(x_192); +x_193 = lean_ctor_get(x_182, 3); +lean_dec(x_193); +x_194 = lean_ctor_get(x_182, 0); +lean_dec(x_194); +x_195 = lean_unsigned_to_nat(3u); +lean_ctor_set(x_182, 4, x_167); +lean_ctor_set(x_182, 3, x_167); +lean_ctor_set(x_182, 2, x_185); +lean_ctor_set(x_182, 1, x_184); +lean_ctor_set(x_182, 0, x_112); +lean_ctor_set(x_111, 4, x_167); +lean_ctor_set(x_111, 2, x_6); +lean_ctor_set(x_111, 1, x_5); +lean_ctor_set(x_111, 0, x_112); +if (lean_is_scalar(x_9)) { + x_196 = lean_alloc_ctor(0, 5, 0); +} else { + x_196 = x_9; +} +lean_ctor_set(x_196, 0, x_195); +lean_ctor_set(x_196, 1, x_190); +lean_ctor_set(x_196, 2, x_191); +lean_ctor_set(x_196, 3, x_182); +lean_ctor_set(x_196, 4, x_111); +return x_196; +} +else +{ +lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; +x_197 = lean_ctor_get(x_182, 1); +x_198 = lean_ctor_get(x_182, 2); +lean_inc(x_198); +lean_inc(x_197); +lean_dec(x_182); +x_199 = lean_unsigned_to_nat(3u); +x_200 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_200, 0, x_112); +lean_ctor_set(x_200, 1, x_184); +lean_ctor_set(x_200, 2, x_185); +lean_ctor_set(x_200, 3, x_167); +lean_ctor_set(x_200, 4, x_167); +lean_ctor_set(x_111, 4, x_167); +lean_ctor_set(x_111, 2, x_6); +lean_ctor_set(x_111, 1, x_5); +lean_ctor_set(x_111, 0, x_112); +if (lean_is_scalar(x_9)) { + x_201 = lean_alloc_ctor(0, 5, 0); +} else { + x_201 = x_9; +} +lean_ctor_set(x_201, 0, x_199); +lean_ctor_set(x_201, 1, x_197); +lean_ctor_set(x_201, 2, x_198); +lean_ctor_set(x_201, 3, x_200); +lean_ctor_set(x_201, 4, x_111); +return x_201; +} +} +else +{ +lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; +x_202 = lean_ctor_get(x_111, 1); +x_203 = lean_ctor_get(x_111, 2); +lean_inc(x_203); +lean_inc(x_202); +lean_dec(x_111); +x_204 = lean_ctor_get(x_182, 1); +lean_inc(x_204); +x_205 = lean_ctor_get(x_182, 2); +lean_inc(x_205); +if (lean_is_exclusive(x_182)) { + lean_ctor_release(x_182, 0); + lean_ctor_release(x_182, 1); + lean_ctor_release(x_182, 2); + lean_ctor_release(x_182, 3); + lean_ctor_release(x_182, 4); + x_206 = x_182; +} else { + lean_dec_ref(x_182); + x_206 = lean_box(0); +} +x_207 = lean_unsigned_to_nat(3u); +if (lean_is_scalar(x_206)) { + x_208 = lean_alloc_ctor(0, 5, 0); +} else { + x_208 = x_206; +} +lean_ctor_set(x_208, 0, x_112); +lean_ctor_set(x_208, 1, x_202); +lean_ctor_set(x_208, 2, x_203); +lean_ctor_set(x_208, 3, x_167); +lean_ctor_set(x_208, 4, x_167); +x_209 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_209, 0, x_112); +lean_ctor_set(x_209, 1, x_5); +lean_ctor_set(x_209, 2, x_6); +lean_ctor_set(x_209, 3, x_167); +lean_ctor_set(x_209, 4, x_167); +if (lean_is_scalar(x_9)) { + x_210 = lean_alloc_ctor(0, 5, 0); +} else { + x_210 = x_9; +} +lean_ctor_set(x_210, 0, x_207); +lean_ctor_set(x_210, 1, x_204); +lean_ctor_set(x_210, 2, x_205); +lean_ctor_set(x_210, 3, x_208); +lean_ctor_set(x_210, 4, x_209); +return x_210; +} +} +else +{ +lean_object* x_211; lean_object* x_212; +x_211 = lean_unsigned_to_nat(2u); +if (lean_is_scalar(x_9)) { + x_212 = lean_alloc_ctor(0, 5, 0); +} else { + x_212 = x_9; +} +lean_ctor_set(x_212, 0, x_211); +lean_ctor_set(x_212, 1, x_5); +lean_ctor_set(x_212, 2, x_6); +lean_ctor_set(x_212, 3, x_111); +lean_ctor_set(x_212, 4, x_182); +return x_212; +} +} +} +} +} +else +{ +lean_object* x_213; lean_object* x_214; +x_213 = lean_unsigned_to_nat(1u); +x_214 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_214, 0, x_213); +lean_ctor_set(x_214, 1, x_1); +lean_ctor_set(x_214, 2, x_2); +lean_ctor_set(x_214, 3, x_3); +lean_ctor_set(x_214, 4, x_3); +return x_214; +} +} +} +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_insert___at___00__private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go_spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = l_Std_DTreeMap_Internal_Impl_insert___at___00__private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go_spec__1___redArg(x_2, x_3, x_4); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; uint8_t x_6; +x_5 = lean_array_get_size(x_1); +x_6 = lean_nat_dec_lt(x_3, x_5); +if (x_6 == 0) +{ +lean_dec(x_3); +lean_dec_ref(x_2); +return x_4; +} +else +{ +lean_object* x_7; +x_7 = lean_array_fget_borrowed(x_1, x_3); +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_8 = lean_ctor_get(x_7, 0); +lean_inc_ref(x_2); +lean_inc(x_3); +x_9 = lean_apply_1(x_2, x_3); +x_10 = lean_unbox(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_add(x_3, x_11); +lean_dec(x_3); +x_3 = x_12; +goto _start; +} +else +{ +lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_14 = lean_unsigned_to_nat(1u); +x_15 = lean_nat_add(x_3, x_14); +lean_dec(x_3); +x_16 = l_Std_DTreeMap_Internal_Impl_contains___at___00__private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go_spec__0___redArg(x_8, x_4); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_box(0); +lean_inc(x_8); +x_18 = l_Std_DTreeMap_Internal_Impl_insert___at___00__private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go_spec__1___redArg(x_8, x_17, x_4); +x_3 = x_15; +x_4 = x_18; +goto _start; +} +else +{ +x_3 = x_15; +goto _start; +} +} +} +else +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_unsigned_to_nat(1u); +x_22 = lean_nat_add(x_3, x_21); +lean_dec(x_3); +x_3 = x_22; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_contains___at___00__private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = l_Std_DTreeMap_Internal_Impl_contains___at___00__private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go_spec__0(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +x_5 = lean_box(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_contains___at___00__private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_DTreeMap_Internal_Impl_contains___at___00__private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go_spec__0___redArg(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go(x_1, x_2, x_3, x_4); +lean_dec_ref(x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go_match__1_splitter(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_5; lean_object* x_6; +lean_dec(x_4); +x_5 = lean_ctor_get(x_2, 0); +lean_inc(x_5); +lean_dec_ref(x_2); +x_6 = lean_apply_1(x_3, x_5); +return x_6; +} +else +{ +lean_object* x_7; +lean_dec(x_3); +x_7 = lean_apply_2(x_4, x_2, lean_box(0)); +return x_7; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go_match__1_splitter___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_4; lean_object* x_5; +lean_dec(x_3); +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +lean_dec_ref(x_1); +x_5 = lean_apply_1(x_2, x_4); +return x_5; +} +else +{ +lean_object* x_6; +lean_dec(x_2); +x_6 = lean_apply_2(x_3, x_1, lean_box(0)); +return x_6; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_unsigned_to_nat(0u); +x_5 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go(x_1, x_2, x_4, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed(x_1, x_2, x_3); +lean_dec_ref(x_1); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___lam__1(uint8_t x_1, lean_object* x_2) { +_start: +{ +return x_1; +} +} +LEAN_EXPORT uint8_t l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___lam__3(lean_object* x_1) { +_start: +{ +uint8_t x_2; +x_2 = 1; +return x_2; +} +} +LEAN_EXPORT uint8_t l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___lam__0(lean_object* x_1, uint8_t x_2, uint8_t x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_5 = lean_box(0); +x_6 = lean_array_get_borrowed(x_5, x_1, x_4); +x_7 = l_Lean_IR_IRType_isScalarOrStruct(x_6); +if (x_7 == 0) +{ +return x_2; +} +else +{ +return x_3; +} +} +} +LEAN_EXPORT uint8_t l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___lam__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_4 = l_Lean_IR_Decl_params(x_1); +x_5 = lean_array_get(x_2, x_4, x_3); +lean_dec_ref(x_4); +x_6 = lean_ctor_get(x_5, 1); +lean_inc(x_6); +lean_dec(x_5); +x_7 = l_Lean_IR_IRType_isScalarOrStruct(x_6); +lean_dec(x_6); +if (x_7 == 0) +{ +uint8_t x_8; +x_8 = 1; +return x_8; +} +else +{ +uint8_t x_9; +x_9 = 0; +return x_9; +} +} +} +LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; 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_4 = l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__0; +x_5 = l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__1; +x_6 = l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__2; +x_7 = l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__3; +x_8 = l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__4; +x_9 = l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__5; +x_10 = l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__6; +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_4); +lean_ctor_set(x_11, 1, x_5); +x_12 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_6); +lean_ctor_set(x_12, 2, x_7); +lean_ctor_set(x_12, 3, x_8); +lean_ctor_set(x_12, 4, x_9); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_10); +lean_inc_ref(x_13); +x_14 = lean_alloc_closure((void*)(l_StateT_instMonad___redArg___lam__1), 6, 1); +lean_closure_set(x_14, 0, x_13); +lean_inc_ref(x_13); +x_15 = lean_alloc_closure((void*)(l_StateT_instMonad___redArg___lam__4), 6, 1); +lean_closure_set(x_15, 0, x_13); +lean_inc_ref(x_13); +x_16 = lean_alloc_closure((void*)(l_StateT_instMonad___redArg___lam__7), 6, 1); +lean_closure_set(x_16, 0, x_13); +lean_inc_ref(x_13); +x_17 = lean_alloc_closure((void*)(l_StateT_instMonad___redArg___lam__9), 6, 1); +lean_closure_set(x_17, 0, x_13); +lean_inc_ref(x_13); +x_18 = lean_alloc_closure((void*)(l_StateT_map), 8, 3); +lean_closure_set(x_18, 0, lean_box(0)); +lean_closure_set(x_18, 1, lean_box(0)); +lean_closure_set(x_18, 2, x_13); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_14); +lean_inc_ref(x_13); +x_20 = lean_alloc_closure((void*)(l_StateT_pure), 6, 3); +lean_closure_set(x_20, 0, lean_box(0)); +lean_closure_set(x_20, 1, lean_box(0)); +lean_closure_set(x_20, 2, x_13); +x_21 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +lean_ctor_set(x_21, 2, x_15); +lean_ctor_set(x_21, 3, x_16); +lean_ctor_set(x_21, 4, x_17); +x_22 = lean_alloc_closure((void*)(l_StateT_bind), 8, 3); +lean_closure_set(x_22, 0, lean_box(0)); +lean_closure_set(x_22, 1, lean_box(0)); +lean_closure_set(x_22, 2, x_13); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +x_24 = lean_box(0); +x_25 = l_instInhabitedOfMonad___redArg(x_23, x_24); +x_26 = lean_alloc_closure((void*)(l_instInhabitedForall___redArg___lam__0___boxed), 2, 1); +lean_closure_set(x_26, 0, x_25); +x_27 = lean_panic_fn(x_26, x_1); +x_28 = lean_apply_2(x_27, x_2, x_3); +return x_28; +} +} +static lean_object* _init_l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.IR.ExplicitBoxing.BoxParams.visitVDeclExpr", 47, 47); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__3; +x_2 = lean_unsigned_to_nat(46u); +x_3 = lean_unsigned_to_nat(395u); +x_4 = l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___closed__0; +x_5 = l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__1; +x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; uint8_t x_6; uint8_t x_7; lean_object* x_8; +x_5 = lean_unbox(x_2); +x_6 = lean_unbox(x_3); +x_7 = l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___lam__0(x_1, x_5, x_6, x_4); +lean_dec(x_4); +lean_dec_ref(x_1); +x_8 = lean_box(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___lam__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; uint8_t x_4; lean_object* x_5; +x_3 = lean_unbox(x_1); +x_4 = l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___lam__1(x_3, x_2); +lean_dec(x_2); +x_5 = lean_box(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___lam__3___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___lam__3(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___lam__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___lam__2(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec_ref(x_1); +x_5 = lean_box(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr(lean_object* x_1, 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; +switch (lean_obj_tag(x_3)) { +case 0: +{ +lean_object* x_11; lean_object* x_12; uint8_t x_13; uint8_t x_14; +x_11 = lean_ctor_get(x_3, 0); +x_12 = lean_ctor_get(x_3, 1); +x_13 = l_Lean_IR_IRType_isObj(x_2); +x_14 = 1; +if (x_13 == 0) +{ +uint8_t x_15; lean_object* x_16; lean_object* x_17; +x_15 = l_Std_DTreeMap_Internal_Impl_contains___at___00__private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go_spec__0___redArg(x_1, x_6); +if (x_15 == 0) +{ +if (lean_obj_tag(x_2) == 10) +{ +lean_object* x_23; +lean_dec_ref(x_5); +x_23 = lean_ctor_get(x_2, 1); +lean_inc_ref(x_23); +x_16 = x_23; +x_17 = x_6; +goto block_22; +} +else +{ +if (lean_obj_tag(x_2) == 11) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_24 = lean_ctor_get(x_2, 1); +x_25 = lean_ctor_get(x_11, 1); +x_26 = lean_box(0); +x_27 = lean_array_get_borrowed(x_26, x_24, x_25); +if (lean_obj_tag(x_27) == 10) +{ +lean_object* x_28; +lean_dec_ref(x_5); +x_28 = lean_ctor_get(x_27, 1); +lean_inc_ref(x_28); +x_16 = x_28; +x_17 = x_6; +goto block_22; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___closed__1; +x_30 = l_panic___at___00Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr_spec__0(x_29, x_5, x_6); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec_ref(x_30); +x_7 = x_31; +goto block_10; +} +} +else +{ +lean_dec_ref(x_5); +x_7 = x_6; +goto block_10; +} +} +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +lean_dec_ref(x_5); +x_32 = lean_box(x_14); +x_33 = lean_alloc_closure((void*)(l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___lam__1___boxed), 2, 1); +lean_closure_set(x_33, 0, x_32); +x_34 = l_Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed(x_12, x_33, x_6); +x_35 = l_Lean_IR_IRType_boxed(x_2); +lean_dec(x_2); +x_36 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_36, 0, x_1); +lean_ctor_set(x_36, 1, x_35); +lean_ctor_set(x_36, 2, x_3); +lean_ctor_set(x_36, 3, x_4); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_34); +return x_37; +} +block_22: +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_18 = lean_box(x_14); +x_19 = lean_box(x_15); +x_20 = lean_alloc_closure((void*)(l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___lam__0___boxed), 4, 3); +lean_closure_set(x_20, 0, x_16); +lean_closure_set(x_20, 1, x_18); +lean_closure_set(x_20, 2, x_19); +x_21 = l_Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed(x_12, x_20, x_17); +x_7 = x_21; +goto block_10; +} +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; +lean_dec_ref(x_5); +x_38 = lean_box(x_14); +x_39 = lean_alloc_closure((void*)(l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___lam__1___boxed), 2, 1); +lean_closure_set(x_39, 0, x_38); +x_40 = l_Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed(x_12, x_39, x_6); +x_7 = x_40; +goto block_10; +} +} +case 2: +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +lean_dec_ref(x_5); +x_41 = lean_ctor_get(x_3, 2); +x_42 = lean_alloc_closure((void*)(l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___lam__3___boxed), 1, 0); +x_43 = l_Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed(x_41, x_42, x_6); +x_44 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_44, 0, x_1); +lean_ctor_set(x_44, 1, x_2); +lean_ctor_set(x_44, 2, x_3); +lean_ctor_set(x_44, 3, x_4); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_43); +return x_45; +} +case 6: +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_57; +x_46 = lean_ctor_get(x_3, 0); +x_47 = lean_ctor_get(x_3, 1); +x_48 = lean_ctor_get(x_5, 3); +lean_inc_ref(x_48); +x_49 = lean_ctor_get(x_5, 4); +lean_inc_ref(x_49); +lean_dec_ref(x_5); +x_50 = l_Lean_IR_ExplicitBoxing_mkBoxedVersionAux___closed__0; +lean_inc(x_46); +x_57 = l_Lean_IR_findEnvDecl_x27(x_49, x_46, x_48); +lean_dec_ref(x_48); +if (lean_obj_tag(x_57) == 0) +{ +lean_object* x_58; lean_object* x_59; +x_58 = l_Lean_IR_ExplicitBoxing_transformBoxedPaps___closed__3; +x_59 = l_panic___at___00Lean_IR_ExplicitBoxing_transformBoxedPaps_spec__0(x_58); +x_51 = x_59; +goto block_56; +} +else +{ +lean_object* x_60; +x_60 = lean_ctor_get(x_57, 0); +lean_inc(x_60); +lean_dec_ref(x_57); +x_51 = x_60; +goto block_56; +} +block_56: +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_52 = lean_alloc_closure((void*)(l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___lam__2___boxed), 3, 2); +lean_closure_set(x_52, 0, x_51); +lean_closure_set(x_52, 1, x_50); +x_53 = l_Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed(x_47, x_52, x_6); +x_54 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_54, 0, x_1); +lean_ctor_set(x_54, 1, x_2); +lean_ctor_set(x_54, 2, x_3); +lean_ctor_set(x_54, 3, x_4); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_53); +return x_55; +} +} +case 7: +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +lean_dec_ref(x_5); +x_61 = lean_ctor_get(x_3, 1); +x_62 = lean_alloc_closure((void*)(l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___lam__3___boxed), 1, 0); +x_63 = l_Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed(x_61, x_62, x_6); +x_64 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_64, 0, x_1); +lean_ctor_set(x_64, 1, x_2); +lean_ctor_set(x_64, 2, x_3); +lean_ctor_set(x_64, 3, x_4); +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_64); +lean_ctor_set(x_65, 1, x_63); +return x_65; +} +case 8: +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; +lean_dec_ref(x_5); +x_66 = lean_ctor_get(x_3, 1); +x_67 = lean_alloc_closure((void*)(l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___lam__3___boxed), 1, 0); +x_68 = l_Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed(x_66, x_67, x_6); +x_69 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_69, 0, x_1); +lean_ctor_set(x_69, 1, x_2); +lean_ctor_set(x_69, 2, x_3); +lean_ctor_set(x_69, 3, x_4); +x_70 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_68); +return x_70; +} +default: +{ +lean_object* x_71; lean_object* x_72; +lean_dec_ref(x_5); +x_71 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_71, 0, x_1); +lean_ctor_set(x_71, 1, x_2); +lean_ctor_set(x_71, 2, x_3); +lean_ctor_set(x_71, 3, x_4); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set(x_72, 1, x_6); +return x_72; +} +} +block_10: +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_8, 0, x_1); +lean_ctor_set(x_8, 1, x_2); +lean_ctor_set(x_8, 2, x_3); +lean_ctor_set(x_8, 3, x_4); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_7); +return x_9; +} +} +} +static lean_object* _init_l_panic___at___00Lean_IR_ExplicitBoxing_BoxParams_boxParams_spec__2___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = l_Array_empty(lean_box(0)); +return x_1; +} +} +LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_ExplicitBoxing_BoxParams_boxParams_spec__2(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_panic___at___00Lean_IR_ExplicitBoxing_BoxParams_boxParams_spec__2___closed__0; +x_3 = lean_panic_fn(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT uint8_t l_Lean_IR_ExplicitBoxing_BoxParams_boxParams___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_array_get_borrowed(x_1, x_2, x_3); +x_5 = lean_ctor_get(x_4, 1); +x_6 = l_Lean_IR_IRType_isScalarOrStruct(x_5); +if (x_6 == 0) +{ +uint8_t x_7; +x_7 = 1; +return x_7; +} +else +{ +uint8_t x_8; +x_8 = 0; +return x_8; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_BasicAux_0__mapMonoMImp_go___at___00Lean_IR_ExplicitBoxing_BoxParams_boxParams_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_array_get_size(x_3); +x_5 = lean_nat_dec_lt(x_2, x_4); +if (x_5 == 0) +{ +lean_dec(x_2); +return x_3; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_19; uint8_t x_20; lean_object* x_21; uint8_t x_22; +x_6 = lean_array_fget_borrowed(x_3, x_2); +x_19 = lean_ctor_get(x_6, 0); +x_20 = lean_ctor_get_uint8(x_6, sizeof(void*)*2); +x_21 = lean_ctor_get(x_6, 1); +x_22 = l_Lean_IR_ExplicitBoxing_BoxParams_isExpensiveScalar(x_21); +if (x_22 == 0) +{ +lean_inc(x_6); +x_7 = x_6; +goto block_18; +} +else +{ +uint8_t x_23; +x_23 = l_Std_DTreeMap_Internal_Impl_contains___at___00__private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go_spec__0___redArg(x_19, x_1); +if (x_23 == 0) +{ +lean_inc(x_6); +x_7 = x_6; +goto block_18; +} +else +{ +lean_object* x_24; lean_object* x_25; +x_24 = l_Lean_IR_IRType_boxed(x_21); +lean_inc(x_19); +x_25 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_25, 0, x_19); +lean_ctor_set(x_25, 1, x_24); +lean_ctor_set_uint8(x_25, sizeof(void*)*2, x_20); +x_7 = x_25; +goto block_18; +} +} +block_18: +{ +size_t x_8; size_t x_9; uint8_t x_10; +x_8 = lean_ptr_addr(x_6); +x_9 = lean_ptr_addr(x_7); +x_10 = lean_usize_dec_eq(x_8, x_9); +if (x_10 == 0) { -lean_object* x_110; -x_110 = lean_ctor_get(x_108, 0); -lean_ctor_set_tag(x_1, 0); -lean_ctor_set(x_1, 3, x_110); -lean_ctor_set(x_1, 2, x_106); -lean_ctor_set(x_1, 1, x_100); -lean_ctor_set(x_1, 0, x_103); -lean_ctor_set(x_108, 0, x_1); -return x_108; +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_add(x_2, x_11); +x_13 = lean_array_fset(x_3, x_2, x_7); +lean_dec(x_2); +x_2 = x_12; +x_3 = x_13; +goto _start; } else { -lean_object* x_111; lean_object* x_112; lean_object* x_113; -x_111 = lean_ctor_get(x_108, 0); -x_112 = lean_ctor_get(x_108, 1); -lean_inc(x_112); -lean_inc(x_111); -lean_dec(x_108); -lean_ctor_set_tag(x_1, 0); -lean_ctor_set(x_1, 3, x_111); -lean_ctor_set(x_1, 2, x_106); -lean_ctor_set(x_1, 1, x_100); -lean_ctor_set(x_1, 0, x_103); -x_113 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_113, 0, x_1); -lean_ctor_set(x_113, 1, x_112); -return x_113; +lean_object* x_15; lean_object* x_16; +lean_dec_ref(x_7); +x_15 = lean_unsigned_to_nat(1u); +x_16 = lean_nat_add(x_2, x_15); +lean_dec(x_2); +x_2 = x_16; +goto _start; } } -else +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_BoxParams_boxParams___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_114; -lean_dec(x_98); -lean_free_object(x_1); -x_114 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__0(x_90, x_91, x_95, x_92, x_2, x_99); +uint8_t x_4; lean_object* x_5; +x_4 = l_Lean_IR_ExplicitBoxing_BoxParams_boxParams___lam__0(x_1, x_2, x_3); +lean_dec(x_3); lean_dec_ref(x_2); -return x_114; +x_5 = lean_box(x_4); +return x_5; } } -else +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_BoxParams_boxParams(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; uint8_t x_126; -x_115 = lean_ctor_get(x_1, 0); -x_116 = lean_ctor_get(x_1, 1); -x_117 = lean_ctor_get(x_1, 2); -x_118 = lean_ctor_get(x_1, 3); -lean_inc(x_118); -lean_inc(x_117); -lean_inc(x_116); -lean_inc(x_115); -lean_dec(x_1); +switch (lean_obj_tag(x_1)) { +case 0: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_1, 1); +lean_inc(x_5); +x_6 = lean_ctor_get(x_1, 2); +lean_inc_ref(x_6); +x_7 = lean_ctor_get(x_1, 3); +lean_inc(x_7); +lean_dec_ref(x_1); lean_inc_ref(x_2); -x_119 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_118, x_2, x_3); -x_120 = lean_ctor_get(x_119, 0); -lean_inc(x_120); -x_121 = lean_ctor_get(x_119, 1); -lean_inc(x_121); -lean_dec_ref(x_119); -x_122 = l_Lean_IR_ExplicitBoxing_getVarType(x_117, x_2, x_121); -x_123 = lean_ctor_get(x_122, 0); -lean_inc(x_123); -x_124 = lean_ctor_get(x_122, 1); -lean_inc(x_124); -lean_dec_ref(x_122); -x_125 = lean_box(5); -lean_inc(x_123); -x_126 = l_Lean_IR_ExplicitBoxing_eqvTypes(x_123, x_125); -if (x_126 == 0) +x_8 = l_Lean_IR_ExplicitBoxing_BoxParams_boxParams(x_7, x_2, x_3); +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_ref(x_8); +x_11 = l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr(x_4, x_5, x_6, x_9, x_2, x_10); +return x_11; +} +case 1: { -lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; -x_127 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(x_124); -x_128 = lean_ctor_get(x_127, 0); -lean_inc(x_128); -x_129 = lean_ctor_get(x_127, 1); -lean_inc(x_129); -lean_dec_ref(x_127); +uint8_t x_12; +x_12 = !lean_is_exclusive(x_1); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_13 = lean_ctor_get(x_1, 0); +x_14 = lean_ctor_get(x_1, 1); +x_15 = lean_ctor_get(x_1, 2); +x_16 = lean_ctor_get(x_1, 3); lean_inc_ref(x_2); -x_130 = l_Lean_IR_ExplicitBoxing_mkCast(x_117, x_123, x_125, x_2, x_129); -x_131 = lean_ctor_get(x_130, 0); -lean_inc(x_131); -x_132 = lean_ctor_get(x_130, 1); -lean_inc(x_132); -lean_dec_ref(x_130); -lean_inc(x_128); -x_133 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__0(x_115, x_116, x_120, x_128, x_2, x_132); -lean_dec_ref(x_2); -x_134 = lean_ctor_get(x_133, 0); -lean_inc(x_134); -x_135 = lean_ctor_get(x_133, 1); -lean_inc(x_135); -if (lean_is_exclusive(x_133)) { - lean_ctor_release(x_133, 0); - lean_ctor_release(x_133, 1); - x_136 = x_133; +x_17 = l_Lean_IR_ExplicitBoxing_BoxParams_boxParams(x_15, x_2, x_3); +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_ref(x_17); +x_20 = !lean_is_exclusive(x_2); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_21 = lean_ctor_get(x_2, 1); +x_22 = lean_unsigned_to_nat(0u); +x_23 = l___private_Init_Data_Array_BasicAux_0__mapMonoMImp_go___at___00Lean_IR_ExplicitBoxing_BoxParams_boxParams_spec__0(x_19, x_22, x_14); +lean_inc(x_18); +lean_inc_ref(x_23); +lean_inc(x_13); +x_24 = l_Lean_IR_LocalContext_addJP(x_21, x_13, x_23, x_18); +lean_ctor_set(x_2, 1, x_24); +x_25 = l_Lean_IR_ExplicitBoxing_BoxParams_boxParams(x_16, x_2, x_19); +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) +{ +lean_object* x_27; +x_27 = lean_ctor_get(x_25, 0); +lean_ctor_set(x_1, 3, x_27); +lean_ctor_set(x_1, 2, x_18); +lean_ctor_set(x_1, 1, x_23); +lean_ctor_set(x_25, 0, x_1); +return x_25; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_25, 0); +x_29 = lean_ctor_get(x_25, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_25); +lean_ctor_set(x_1, 3, x_28); +lean_ctor_set(x_1, 2, x_18); +lean_ctor_set(x_1, 1, x_23); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_1); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_31 = lean_ctor_get(x_2, 0); +x_32 = lean_ctor_get(x_2, 1); +x_33 = lean_ctor_get(x_2, 2); +x_34 = lean_ctor_get(x_2, 3); +x_35 = lean_ctor_get(x_2, 4); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_2); +x_36 = lean_unsigned_to_nat(0u); +x_37 = l___private_Init_Data_Array_BasicAux_0__mapMonoMImp_go___at___00Lean_IR_ExplicitBoxing_BoxParams_boxParams_spec__0(x_19, x_36, x_14); +lean_inc(x_18); +lean_inc_ref(x_37); +lean_inc(x_13); +x_38 = l_Lean_IR_LocalContext_addJP(x_32, x_13, x_37, x_18); +x_39 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_39, 0, x_31); +lean_ctor_set(x_39, 1, x_38); +lean_ctor_set(x_39, 2, x_33); +lean_ctor_set(x_39, 3, x_34); +lean_ctor_set(x_39, 4, x_35); +x_40 = l_Lean_IR_ExplicitBoxing_BoxParams_boxParams(x_16, x_39, x_19); +x_41 = lean_ctor_get(x_40, 0); +lean_inc(x_41); +x_42 = lean_ctor_get(x_40, 1); +lean_inc(x_42); +if (lean_is_exclusive(x_40)) { + lean_ctor_release(x_40, 0); + lean_ctor_release(x_40, 1); + x_43 = x_40; } else { - lean_dec_ref(x_133); - x_136 = lean_box(0); -} -x_137 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_137, 0, x_128); -lean_ctor_set(x_137, 1, x_125); -lean_ctor_set(x_137, 2, x_131); -lean_ctor_set(x_137, 3, x_134); -if (lean_is_scalar(x_136)) { - x_138 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_40); + x_43 = lean_box(0); +} +lean_ctor_set(x_1, 3, x_41); +lean_ctor_set(x_1, 2, x_18); +lean_ctor_set(x_1, 1, x_37); +if (lean_is_scalar(x_43)) { + x_44 = lean_alloc_ctor(0, 2, 0); } else { - x_138 = x_136; + x_44 = x_43; +} +lean_ctor_set(x_44, 0, x_1); +lean_ctor_set(x_44, 1, x_42); +return x_44; } -lean_ctor_set(x_138, 0, x_137); -lean_ctor_set(x_138, 1, x_135); -return x_138; } else { -lean_object* x_139; -lean_dec(x_123); -x_139 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__0(x_115, x_116, x_120, x_117, x_2, x_124); -lean_dec_ref(x_2); -return x_139; +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_45 = lean_ctor_get(x_1, 0); +x_46 = lean_ctor_get(x_1, 1); +x_47 = lean_ctor_get(x_1, 2); +x_48 = lean_ctor_get(x_1, 3); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_1); +lean_inc_ref(x_2); +x_49 = l_Lean_IR_ExplicitBoxing_BoxParams_boxParams(x_47, x_2, x_3); +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec_ref(x_49); +x_52 = lean_ctor_get(x_2, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_2, 1); +lean_inc(x_53); +x_54 = lean_ctor_get(x_2, 2); +lean_inc(x_54); +x_55 = lean_ctor_get(x_2, 3); +lean_inc_ref(x_55); +x_56 = lean_ctor_get(x_2, 4); +lean_inc_ref(x_56); +if (lean_is_exclusive(x_2)) { + lean_ctor_release(x_2, 0); + lean_ctor_release(x_2, 1); + lean_ctor_release(x_2, 2); + lean_ctor_release(x_2, 3); + lean_ctor_release(x_2, 4); + x_57 = x_2; +} else { + lean_dec_ref(x_2); + x_57 = lean_box(0); +} +x_58 = lean_unsigned_to_nat(0u); +x_59 = l___private_Init_Data_Array_BasicAux_0__mapMonoMImp_go___at___00Lean_IR_ExplicitBoxing_BoxParams_boxParams_spec__0(x_51, x_58, x_46); +lean_inc(x_50); +lean_inc_ref(x_59); +lean_inc(x_45); +x_60 = l_Lean_IR_LocalContext_addJP(x_53, x_45, x_59, x_50); +if (lean_is_scalar(x_57)) { + x_61 = lean_alloc_ctor(0, 5, 0); +} else { + x_61 = x_57; +} +lean_ctor_set(x_61, 0, x_52); +lean_ctor_set(x_61, 1, x_60); +lean_ctor_set(x_61, 2, x_54); +lean_ctor_set(x_61, 3, x_55); +lean_ctor_set(x_61, 4, x_56); +x_62 = l_Lean_IR_ExplicitBoxing_BoxParams_boxParams(x_48, x_61, x_51); +x_63 = lean_ctor_get(x_62, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_62, 1); +lean_inc(x_64); +if (lean_is_exclusive(x_62)) { + lean_ctor_release(x_62, 0); + lean_ctor_release(x_62, 1); + x_65 = x_62; +} else { + lean_dec_ref(x_62); + x_65 = lean_box(0); +} +x_66 = lean_alloc_ctor(1, 4, 0); +lean_ctor_set(x_66, 0, x_45); +lean_ctor_set(x_66, 1, x_59); +lean_ctor_set(x_66, 2, x_50); +lean_ctor_set(x_66, 3, x_63); +if (lean_is_scalar(x_65)) { + x_67 = lean_alloc_ctor(0, 2, 0); +} else { + x_67 = x_65; } +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_64); +return x_67; } } -case 5: +case 9: { -lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; uint8_t x_152; -x_140 = lean_ctor_get(x_1, 0); -lean_inc(x_140); -x_141 = lean_ctor_get(x_1, 1); -lean_inc(x_141); -x_142 = lean_ctor_get(x_1, 2); -lean_inc(x_142); -x_143 = lean_ctor_get(x_1, 3); -lean_inc(x_143); -x_144 = lean_ctor_get(x_1, 4); -lean_inc(x_144); -x_145 = lean_ctor_get(x_1, 5); -lean_inc(x_145); -lean_dec_ref(x_1); -lean_inc_ref(x_2); -x_146 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_145, x_2, x_3); -x_147 = lean_ctor_get(x_146, 0); -lean_inc(x_147); -x_148 = lean_ctor_get(x_146, 1); -lean_inc(x_148); -lean_dec_ref(x_146); -x_149 = l_Lean_IR_ExplicitBoxing_getVarType(x_143, x_2, x_148); -x_150 = lean_ctor_get(x_149, 0); -lean_inc(x_150); -x_151 = lean_ctor_get(x_149, 1); -lean_inc(x_151); -lean_dec_ref(x_149); -lean_inc(x_144); -lean_inc(x_150); -x_152 = l_Lean_IR_ExplicitBoxing_eqvTypes(x_150, x_144); -if (x_152 == 0) -{ -lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; uint8_t x_160; -x_153 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(x_151); -x_154 = lean_ctor_get(x_153, 0); -lean_inc(x_154); -x_155 = lean_ctor_get(x_153, 1); -lean_inc(x_155); -lean_dec_ref(x_153); -lean_inc_ref(x_2); -lean_inc(x_144); -x_156 = l_Lean_IR_ExplicitBoxing_mkCast(x_143, x_150, x_144, x_2, x_155); -x_157 = lean_ctor_get(x_156, 0); -lean_inc(x_157); -x_158 = lean_ctor_get(x_156, 1); -lean_inc(x_158); -lean_dec_ref(x_156); -lean_inc(x_154); -lean_inc(x_144); -x_159 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__1(x_140, x_141, x_142, x_144, x_147, x_154, x_2, x_158); -lean_dec_ref(x_2); -x_160 = !lean_is_exclusive(x_159); -if (x_160 == 0) +uint8_t x_68; +x_68 = !lean_is_exclusive(x_1); +if (x_68 == 0) +{ +lean_object* x_69; size_t x_70; size_t x_71; lean_object* x_72; uint8_t x_73; +x_69 = lean_ctor_get(x_1, 3); +x_70 = lean_array_size(x_69); +x_71 = 0; +x_72 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_BoxParams_boxParams_spec__1(x_70, x_71, x_69, x_2, x_3); +x_73 = !lean_is_exclusive(x_72); +if (x_73 == 0) { -lean_object* x_161; lean_object* x_162; -x_161 = lean_ctor_get(x_159, 0); -x_162 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_162, 0, x_154); -lean_ctor_set(x_162, 1, x_144); -lean_ctor_set(x_162, 2, x_157); -lean_ctor_set(x_162, 3, x_161); -lean_ctor_set(x_159, 0, x_162); -return x_159; +lean_object* x_74; +x_74 = lean_ctor_get(x_72, 0); +lean_ctor_set(x_1, 3, x_74); +lean_ctor_set(x_72, 0, x_1); +return x_72; } else { -lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; -x_163 = lean_ctor_get(x_159, 0); -x_164 = lean_ctor_get(x_159, 1); -lean_inc(x_164); -lean_inc(x_163); -lean_dec(x_159); -x_165 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_165, 0, x_154); -lean_ctor_set(x_165, 1, x_144); -lean_ctor_set(x_165, 2, x_157); -lean_ctor_set(x_165, 3, x_163); -x_166 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_166, 0, x_165); -lean_ctor_set(x_166, 1, x_164); -return x_166; +lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_75 = lean_ctor_get(x_72, 0); +x_76 = lean_ctor_get(x_72, 1); +lean_inc(x_76); +lean_inc(x_75); +lean_dec(x_72); +lean_ctor_set(x_1, 3, x_75); +x_77 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_77, 0, x_1); +lean_ctor_set(x_77, 1, x_76); +return x_77; } } else { -lean_object* x_167; -lean_dec(x_150); -x_167 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__1(x_140, x_141, x_142, x_144, x_147, x_143, x_2, x_151); -lean_dec_ref(x_2); -return x_167; +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; size_t x_82; size_t x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_78 = lean_ctor_get(x_1, 0); +x_79 = lean_ctor_get(x_1, 1); +x_80 = lean_ctor_get(x_1, 2); +x_81 = lean_ctor_get(x_1, 3); +lean_inc(x_81); +lean_inc(x_80); +lean_inc(x_79); +lean_inc(x_78); +lean_dec(x_1); +x_82 = lean_array_size(x_81); +x_83 = 0; +x_84 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_BoxParams_boxParams_spec__1(x_82, x_83, x_81, x_2, x_3); +x_85 = lean_ctor_get(x_84, 0); +lean_inc(x_85); +x_86 = lean_ctor_get(x_84, 1); +lean_inc(x_86); +if (lean_is_exclusive(x_84)) { + lean_ctor_release(x_84, 0); + lean_ctor_release(x_84, 1); + x_87 = x_84; +} else { + lean_dec_ref(x_84); + x_87 = lean_box(0); +} +x_88 = lean_alloc_ctor(9, 4, 0); +lean_ctor_set(x_88, 0, x_78); +lean_ctor_set(x_88, 1, x_79); +lean_ctor_set(x_88, 2, x_80); +lean_ctor_set(x_88, 3, x_85); +if (lean_is_scalar(x_87)) { + x_89 = lean_alloc_ctor(0, 2, 0); +} else { + x_89 = x_87; } +lean_ctor_set(x_89, 0, x_88); +lean_ctor_set(x_89, 1, x_86); +return x_89; } -case 9: -{ -uint8_t x_168; -x_168 = !lean_is_exclusive(x_1); -if (x_168 == 0) +} +case 10: { -lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; size_t x_173; size_t x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; uint8_t x_181; -x_169 = lean_ctor_get(x_1, 0); -x_170 = lean_ctor_get(x_1, 1); -x_171 = lean_ctor_get(x_1, 2); -x_172 = lean_ctor_get(x_1, 3); -x_173 = lean_array_size(x_172); -x_174 = 0; -lean_inc_ref(x_2); -x_175 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_visitFnBody_spec__0(x_173, x_174, x_172, x_2, x_3); -x_176 = lean_ctor_get(x_175, 0); -lean_inc(x_176); -x_177 = lean_ctor_get(x_175, 1); -lean_inc(x_177); -lean_dec_ref(x_175); -x_178 = l_Lean_IR_ExplicitBoxing_getVarType(x_170, x_2, x_177); -x_179 = lean_ctor_get(x_178, 0); -lean_inc(x_179); -x_180 = lean_ctor_get(x_178, 1); -lean_inc(x_180); -lean_dec_ref(x_178); -lean_inc(x_171); -lean_inc(x_179); -x_181 = l_Lean_IR_ExplicitBoxing_eqvTypes(x_179, x_171); -if (x_181 == 0) +lean_object* x_90; +x_90 = lean_ctor_get(x_1, 0); +if (lean_obj_tag(x_90) == 0) { -lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; uint8_t x_189; -x_182 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(x_180); -x_183 = lean_ctor_get(x_182, 0); -lean_inc(x_183); -x_184 = lean_ctor_get(x_182, 1); -lean_inc(x_184); -lean_dec_ref(x_182); -lean_inc_ref(x_2); -lean_inc(x_171); -x_185 = l_Lean_IR_ExplicitBoxing_mkCast(x_170, x_179, x_171, x_2, x_184); -x_186 = lean_ctor_get(x_185, 0); -lean_inc(x_186); -x_187 = lean_ctor_get(x_185, 1); -lean_inc(x_187); -lean_dec_ref(x_185); -lean_inc(x_183); -lean_inc(x_171); -x_188 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__2(x_169, x_171, x_176, x_183, x_2, x_187); +lean_object* x_91; lean_object* x_92; uint8_t x_93; +x_91 = lean_ctor_get(x_90, 0); +x_92 = lean_ctor_get(x_2, 2); +lean_inc(x_92); lean_dec_ref(x_2); -x_189 = !lean_is_exclusive(x_188); -if (x_189 == 0) +x_93 = l_Lean_IR_IRType_isScalarOrStruct(x_92); +lean_dec(x_92); +if (x_93 == 0) { -lean_object* x_190; -x_190 = lean_ctor_get(x_188, 0); -lean_ctor_set_tag(x_1, 0); -lean_ctor_set(x_1, 3, x_190); -lean_ctor_set(x_1, 2, x_186); -lean_ctor_set(x_1, 1, x_171); -lean_ctor_set(x_1, 0, x_183); -lean_ctor_set(x_188, 0, x_1); -return x_188; +uint8_t x_94; +x_94 = l_Std_DTreeMap_Internal_Impl_contains___at___00__private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go_spec__0___redArg(x_91, x_3); +if (x_94 == 0) +{ +lean_object* x_95; lean_object* x_96; lean_object* x_97; +x_95 = lean_box(0); +lean_inc(x_91); +x_96 = l_Std_DTreeMap_Internal_Impl_insert___at___00__private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed_go_spec__1___redArg(x_91, x_95, x_3); +x_97 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_97, 0, x_1); +lean_ctor_set(x_97, 1, x_96); +return x_97; } else { -lean_object* x_191; lean_object* x_192; lean_object* x_193; -x_191 = lean_ctor_get(x_188, 0); -x_192 = lean_ctor_get(x_188, 1); -lean_inc(x_192); -lean_inc(x_191); -lean_dec(x_188); -lean_ctor_set_tag(x_1, 0); -lean_ctor_set(x_1, 3, x_191); -lean_ctor_set(x_1, 2, x_186); -lean_ctor_set(x_1, 1, x_171); -lean_ctor_set(x_1, 0, x_183); -x_193 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_193, 0, x_1); -lean_ctor_set(x_193, 1, x_192); -return x_193; +lean_object* x_98; +x_98 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_98, 0, x_1); +lean_ctor_set(x_98, 1, x_3); +return x_98; } } else { -lean_object* x_194; -lean_dec(x_179); -lean_free_object(x_1); -x_194 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__2(x_169, x_171, x_176, x_170, x_2, x_180); -lean_dec_ref(x_2); -return x_194; +lean_object* x_99; +x_99 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_99, 0, x_1); +lean_ctor_set(x_99, 1, x_3); +return x_99; } } else { -lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; size_t x_199; size_t x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; uint8_t x_207; -x_195 = lean_ctor_get(x_1, 0); -x_196 = lean_ctor_get(x_1, 1); -x_197 = lean_ctor_get(x_1, 2); -x_198 = lean_ctor_get(x_1, 3); -lean_inc(x_198); -lean_inc(x_197); -lean_inc(x_196); -lean_inc(x_195); -lean_dec(x_1); -x_199 = lean_array_size(x_198); -x_200 = 0; -lean_inc_ref(x_2); -x_201 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_visitFnBody_spec__0(x_199, x_200, x_198, x_2, x_3); -x_202 = lean_ctor_get(x_201, 0); -lean_inc(x_202); -x_203 = lean_ctor_get(x_201, 1); -lean_inc(x_203); -lean_dec_ref(x_201); -x_204 = l_Lean_IR_ExplicitBoxing_getVarType(x_196, x_2, x_203); -x_205 = lean_ctor_get(x_204, 0); -lean_inc(x_205); -x_206 = lean_ctor_get(x_204, 1); -lean_inc(x_206); -lean_dec_ref(x_204); -lean_inc(x_197); -lean_inc(x_205); -x_207 = l_Lean_IR_ExplicitBoxing_eqvTypes(x_205, x_197); -if (x_207 == 0) -{ -lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; -x_208 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(x_206); -x_209 = lean_ctor_get(x_208, 0); -lean_inc(x_209); -x_210 = lean_ctor_get(x_208, 1); -lean_inc(x_210); -lean_dec_ref(x_208); -lean_inc_ref(x_2); -lean_inc(x_197); -x_211 = l_Lean_IR_ExplicitBoxing_mkCast(x_196, x_205, x_197, x_2, x_210); -x_212 = lean_ctor_get(x_211, 0); -lean_inc(x_212); -x_213 = lean_ctor_get(x_211, 1); -lean_inc(x_213); -lean_dec_ref(x_211); -lean_inc(x_209); -lean_inc(x_197); -x_214 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__2(x_195, x_197, x_202, x_209, x_2, x_213); +lean_object* x_100; lean_dec_ref(x_2); -x_215 = lean_ctor_get(x_214, 0); -lean_inc(x_215); -x_216 = lean_ctor_get(x_214, 1); -lean_inc(x_216); -if (lean_is_exclusive(x_214)) { - lean_ctor_release(x_214, 0); - lean_ctor_release(x_214, 1); - x_217 = x_214; -} else { - lean_dec_ref(x_214); - x_217 = lean_box(0); -} -x_218 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_218, 0, x_209); -lean_ctor_set(x_218, 1, x_197); -lean_ctor_set(x_218, 2, x_212); -lean_ctor_set(x_218, 3, x_215); -if (lean_is_scalar(x_217)) { - x_219 = lean_alloc_ctor(0, 2, 0); -} else { - x_219 = x_217; +x_100 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_100, 0, x_1); +lean_ctor_set(x_100, 1, x_3); +return x_100; } -lean_ctor_set(x_219, 0, x_218); -lean_ctor_set(x_219, 1, x_216); -return x_219; } -else +case 11: { -lean_object* x_220; -lean_dec(x_205); -x_220 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__2(x_195, x_197, x_202, x_196, x_2, x_206); +lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_110; +x_101 = lean_ctor_get(x_1, 0); +x_102 = lean_ctor_get(x_1, 1); +x_103 = lean_ctor_get(x_2, 1); +lean_inc(x_103); lean_dec_ref(x_2); -return x_220; +x_104 = l_Lean_IR_ExplicitBoxing_mkBoxedVersionAux___closed__0; +x_110 = l_Lean_IR_LocalContext_getJPParams(x_103, x_101); +lean_dec(x_103); +if (lean_obj_tag(x_110) == 0) +{ +lean_object* x_111; lean_object* x_112; +x_111 = l_Lean_IR_ExplicitBoxing_transformBoxedPaps___closed__3; +x_112 = l_panic___at___00Lean_IR_ExplicitBoxing_BoxParams_boxParams_spec__2(x_111); +x_105 = x_112; +goto block_109; } +else +{ +lean_object* x_113; +x_113 = lean_ctor_get(x_110, 0); +lean_inc(x_113); +lean_dec_ref(x_110); +x_105 = x_113; +goto block_109; +} +block_109: +{ +lean_object* x_106; lean_object* x_107; lean_object* x_108; +x_106 = lean_alloc_closure((void*)(l_Lean_IR_ExplicitBoxing_BoxParams_boxParams___lam__0___boxed), 3, 2); +lean_closure_set(x_106, 0, x_104); +lean_closure_set(x_106, 1, x_105); +x_107 = l_Lean_IR_ExplicitBoxing_BoxParams_checkArgsBoxed(x_102, x_106, x_3); +x_108 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_108, 0, x_1); +lean_ctor_set(x_108, 1, x_107); +return x_108; } } -case 10: +case 4: { -lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; -x_221 = lean_ctor_get(x_1, 0); -lean_inc(x_221); -lean_dec_ref(x_1); -x_222 = l_Lean_IR_ExplicitBoxing_getResultType(x_2, x_3); -x_223 = lean_ctor_get(x_222, 0); -lean_inc(x_223); -x_224 = lean_ctor_get(x_222, 1); -lean_inc(x_224); -lean_dec_ref(x_222); -x_225 = lean_alloc_closure((void*)(l_Lean_IR_ExplicitBoxing_visitFnBody___lam__3___boxed), 3, 0); -if (lean_obj_tag(x_221) == 0) -{ -lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; uint8_t x_230; -x_226 = lean_ctor_get(x_221, 0); -lean_inc(x_226); -lean_dec_ref(x_221); -x_227 = l_Lean_IR_ExplicitBoxing_getVarType(x_226, x_2, x_224); -x_228 = lean_ctor_get(x_227, 0); -lean_inc(x_228); -x_229 = lean_ctor_get(x_227, 1); -lean_inc(x_229); -lean_dec_ref(x_227); -lean_inc(x_223); -lean_inc(x_228); -x_230 = l_Lean_IR_ExplicitBoxing_eqvTypes(x_228, x_223); -if (x_230 == 0) -{ -lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; uint8_t x_238; -x_231 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(x_229); -x_232 = lean_ctor_get(x_231, 0); -lean_inc(x_232); -x_233 = lean_ctor_get(x_231, 1); -lean_inc(x_233); -lean_dec_ref(x_231); -lean_inc_ref(x_2); -lean_inc(x_223); -x_234 = l_Lean_IR_ExplicitBoxing_mkCast(x_226, x_228, x_223, x_2, x_233); -x_235 = lean_ctor_get(x_234, 0); -lean_inc(x_235); -x_236 = lean_ctor_get(x_234, 1); -lean_inc(x_236); -lean_dec_ref(x_234); -lean_inc(x_232); -x_237 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__4(x_225, x_232, x_2, x_236); -x_238 = !lean_is_exclusive(x_237); -if (x_238 == 0) +uint8_t x_114; +x_114 = !lean_is_exclusive(x_1); +if (x_114 == 0) { -lean_object* x_239; lean_object* x_240; -x_239 = lean_ctor_get(x_237, 0); -x_240 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_240, 0, x_232); -lean_ctor_set(x_240, 1, x_223); -lean_ctor_set(x_240, 2, x_235); -lean_ctor_set(x_240, 3, x_239); -lean_ctor_set(x_237, 0, x_240); -return x_237; -} -else +lean_object* x_115; lean_object* x_116; uint8_t x_117; +x_115 = lean_ctor_get(x_1, 4); +x_116 = l_Lean_IR_ExplicitBoxing_BoxParams_boxParams(x_115, x_2, x_3); +x_117 = !lean_is_exclusive(x_116); +if (x_117 == 0) { -lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; -x_241 = lean_ctor_get(x_237, 0); -x_242 = lean_ctor_get(x_237, 1); -lean_inc(x_242); -lean_inc(x_241); -lean_dec(x_237); -x_243 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_243, 0, x_232); -lean_ctor_set(x_243, 1, x_223); -lean_ctor_set(x_243, 2, x_235); -lean_ctor_set(x_243, 3, x_241); -x_244 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_244, 0, x_243); -lean_ctor_set(x_244, 1, x_242); -return x_244; -} +lean_object* x_118; +x_118 = lean_ctor_get(x_116, 0); +lean_ctor_set(x_1, 4, x_118); +lean_ctor_set(x_116, 0, x_1); +return x_116; } else { -lean_object* x_245; -lean_dec(x_228); -lean_dec(x_223); -x_245 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__4(x_225, x_226, x_2, x_229); -return x_245; +lean_object* x_119; lean_object* x_120; lean_object* x_121; +x_119 = lean_ctor_get(x_116, 0); +x_120 = lean_ctor_get(x_116, 1); +lean_inc(x_120); +lean_inc(x_119); +lean_dec(x_116); +lean_ctor_set(x_1, 4, x_119); +x_121 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_121, 0, x_1); +lean_ctor_set(x_121, 1, x_120); +return x_121; } } else { -lean_object* x_246; -lean_dec_ref(x_225); -lean_dec(x_223); -x_246 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__3(x_221, x_2, x_224); -lean_dec_ref(x_2); -return x_246; +lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; +x_122 = lean_ctor_get(x_1, 0); +x_123 = lean_ctor_get(x_1, 1); +x_124 = lean_ctor_get(x_1, 2); +x_125 = lean_ctor_get(x_1, 3); +x_126 = lean_ctor_get(x_1, 4); +lean_inc(x_126); +lean_inc(x_125); +lean_inc(x_124); +lean_inc(x_123); +lean_inc(x_122); +lean_dec(x_1); +x_127 = l_Lean_IR_ExplicitBoxing_BoxParams_boxParams(x_126, x_2, x_3); +x_128 = lean_ctor_get(x_127, 0); +lean_inc(x_128); +x_129 = lean_ctor_get(x_127, 1); +lean_inc(x_129); +if (lean_is_exclusive(x_127)) { + lean_ctor_release(x_127, 0); + lean_ctor_release(x_127, 1); + x_130 = x_127; +} else { + lean_dec_ref(x_127); + x_130 = lean_box(0); +} +x_131 = lean_alloc_ctor(4, 5, 0); +lean_ctor_set(x_131, 0, x_122); +lean_ctor_set(x_131, 1, x_123); +lean_ctor_set(x_131, 2, x_124); +lean_ctor_set(x_131, 3, x_125); +lean_ctor_set(x_131, 4, x_128); +if (lean_is_scalar(x_130)) { + x_132 = lean_alloc_ctor(0, 2, 0); +} else { + x_132 = x_130; } +lean_ctor_set(x_132, 0, x_131); +lean_ctor_set(x_132, 1, x_129); +return x_132; } -case 11: +} +case 5: { -uint8_t x_247; -x_247 = !lean_is_exclusive(x_1); -if (x_247 == 0) +uint8_t x_133; +x_133 = !lean_is_exclusive(x_1); +if (x_133 == 0) { -lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; uint8_t x_258; -x_248 = lean_ctor_get(x_1, 0); -x_249 = lean_ctor_get(x_1, 1); -x_250 = l_Lean_IR_ExplicitBoxing_getJPParams(x_248, x_2, x_3); -x_251 = lean_ctor_get(x_250, 0); -lean_inc(x_251); -x_252 = lean_ctor_get(x_250, 1); -lean_inc(x_252); -lean_dec_ref(x_250); -x_253 = l_Lean_IR_ExplicitBoxing_mkBoxedVersionAux___closed__0; -x_254 = lean_alloc_closure((void*)(l_Lean_IR_ExplicitBoxing_visitFnBody___lam__5___boxed), 3, 2); -lean_closure_set(x_254, 0, x_253); -lean_closure_set(x_254, 1, x_251); -x_255 = l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux(x_249, x_254, x_2, x_252); -lean_dec_ref(x_249); -x_256 = lean_ctor_get(x_255, 0); -lean_inc(x_256); -x_257 = lean_ctor_get(x_255, 1); -lean_inc(x_257); -lean_dec_ref(x_255); -x_258 = !lean_is_exclusive(x_256); -if (x_258 == 0) +lean_object* x_134; lean_object* x_135; uint8_t x_136; +x_134 = lean_ctor_get(x_1, 6); +x_135 = l_Lean_IR_ExplicitBoxing_BoxParams_boxParams(x_134, x_2, x_3); +x_136 = !lean_is_exclusive(x_135); +if (x_136 == 0) { -lean_object* x_259; lean_object* x_260; lean_object* x_261; -x_259 = lean_ctor_get(x_256, 0); -x_260 = lean_ctor_get(x_256, 1); -lean_ctor_set(x_1, 1, x_259); -x_261 = l_Lean_IR_reshape(x_260, x_1); -lean_ctor_set(x_256, 1, x_257); -lean_ctor_set(x_256, 0, x_261); -return x_256; +lean_object* x_137; +x_137 = lean_ctor_get(x_135, 0); +lean_ctor_set(x_1, 6, x_137); +lean_ctor_set(x_135, 0, x_1); +return x_135; } else { -lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; -x_262 = lean_ctor_get(x_256, 0); -x_263 = lean_ctor_get(x_256, 1); -lean_inc(x_263); -lean_inc(x_262); -lean_dec(x_256); -lean_ctor_set(x_1, 1, x_262); -x_264 = l_Lean_IR_reshape(x_263, x_1); -x_265 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_265, 0, x_264); -lean_ctor_set(x_265, 1, x_257); -return x_265; +lean_object* x_138; lean_object* x_139; lean_object* x_140; +x_138 = lean_ctor_get(x_135, 0); +x_139 = lean_ctor_get(x_135, 1); +lean_inc(x_139); +lean_inc(x_138); +lean_dec(x_135); +lean_ctor_set(x_1, 6, x_138); +x_140 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_140, 0, x_1); +lean_ctor_set(x_140, 1, x_139); +return x_140; } } else { -lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; -x_266 = lean_ctor_get(x_1, 0); -x_267 = lean_ctor_get(x_1, 1); -lean_inc(x_267); -lean_inc(x_266); +lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; +x_141 = lean_ctor_get(x_1, 0); +x_142 = lean_ctor_get(x_1, 1); +x_143 = lean_ctor_get(x_1, 2); +x_144 = lean_ctor_get(x_1, 3); +x_145 = lean_ctor_get(x_1, 4); +x_146 = lean_ctor_get(x_1, 5); +x_147 = lean_ctor_get(x_1, 6); +lean_inc(x_147); +lean_inc(x_146); +lean_inc(x_145); +lean_inc(x_144); +lean_inc(x_143); +lean_inc(x_142); +lean_inc(x_141); lean_dec(x_1); -x_268 = l_Lean_IR_ExplicitBoxing_getJPParams(x_266, x_2, x_3); -x_269 = lean_ctor_get(x_268, 0); -lean_inc(x_269); -x_270 = lean_ctor_get(x_268, 1); -lean_inc(x_270); -lean_dec_ref(x_268); -x_271 = l_Lean_IR_ExplicitBoxing_mkBoxedVersionAux___closed__0; -x_272 = lean_alloc_closure((void*)(l_Lean_IR_ExplicitBoxing_visitFnBody___lam__5___boxed), 3, 2); -lean_closure_set(x_272, 0, x_271); -lean_closure_set(x_272, 1, x_269); -x_273 = l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux(x_267, x_272, x_2, x_270); -lean_dec_ref(x_267); -x_274 = lean_ctor_get(x_273, 0); -lean_inc(x_274); -x_275 = lean_ctor_get(x_273, 1); -lean_inc(x_275); -lean_dec_ref(x_273); -x_276 = lean_ctor_get(x_274, 0); -lean_inc(x_276); -x_277 = lean_ctor_get(x_274, 1); -lean_inc(x_277); -if (lean_is_exclusive(x_274)) { - lean_ctor_release(x_274, 0); - lean_ctor_release(x_274, 1); - x_278 = x_274; +x_148 = l_Lean_IR_ExplicitBoxing_BoxParams_boxParams(x_147, x_2, x_3); +x_149 = lean_ctor_get(x_148, 0); +lean_inc(x_149); +x_150 = lean_ctor_get(x_148, 1); +lean_inc(x_150); +if (lean_is_exclusive(x_148)) { + lean_ctor_release(x_148, 0); + lean_ctor_release(x_148, 1); + x_151 = x_148; } else { - lean_dec_ref(x_274); - x_278 = lean_box(0); -} -x_279 = lean_alloc_ctor(11, 2, 0); -lean_ctor_set(x_279, 0, x_266); -lean_ctor_set(x_279, 1, x_276); -x_280 = l_Lean_IR_reshape(x_277, x_279); -if (lean_is_scalar(x_278)) { - x_281 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_148); + x_151 = lean_box(0); +} +x_152 = lean_alloc_ctor(5, 7, 0); +lean_ctor_set(x_152, 0, x_141); +lean_ctor_set(x_152, 1, x_142); +lean_ctor_set(x_152, 2, x_143); +lean_ctor_set(x_152, 3, x_144); +lean_ctor_set(x_152, 4, x_145); +lean_ctor_set(x_152, 5, x_146); +lean_ctor_set(x_152, 6, x_149); +if (lean_is_scalar(x_151)) { + x_153 = lean_alloc_ctor(0, 2, 0); } else { - x_281 = x_278; + x_153 = x_151; } -lean_ctor_set(x_281, 0, x_280); -lean_ctor_set(x_281, 1, x_275); -return x_281; +lean_ctor_set(x_153, 0, x_152); +lean_ctor_set(x_153, 1, x_150); +return x_153; } } default: { -lean_object* x_282; +lean_object* x_154; lean_dec_ref(x_2); -x_282 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_282, 0, x_1); -lean_ctor_set(x_282, 1, x_3); -return x_282; +x_154 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_154, 0, x_1); +lean_ctor_set(x_154, 1, x_3); +return x_154; } } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_visitFnBody_spec__0(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_BoxParams_boxParams_spec__1(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { uint8_t x_6; @@ -5364,7 +8664,7 @@ if (x_18 == 0) lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_8, 1); lean_inc_ref(x_4); -x_20 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_19, x_4, x_5); +x_20 = l_Lean_IR_ExplicitBoxing_BoxParams_boxParams(x_19, x_4, x_5); x_21 = lean_ctor_get(x_20, 0); lean_inc(x_21); x_22 = lean_ctor_get(x_20, 1); @@ -5384,7 +8684,7 @@ lean_inc(x_24); lean_inc(x_23); lean_dec(x_8); lean_inc_ref(x_4); -x_25 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_24, x_4, x_5); +x_25 = l_Lean_IR_ExplicitBoxing_BoxParams_boxParams(x_24, x_4, x_5); x_26 = lean_ctor_get(x_25, 0); lean_inc(x_26); x_27 = lean_ctor_get(x_25, 1); @@ -5407,7 +8707,7 @@ if (x_29 == 0) lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; x_30 = lean_ctor_get(x_8, 0); lean_inc_ref(x_4); -x_31 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_30, x_4, x_5); +x_31 = l_Lean_IR_ExplicitBoxing_BoxParams_boxParams(x_30, x_4, x_5); x_32 = lean_ctor_get(x_31, 0); lean_inc(x_32); x_33 = lean_ctor_get(x_31, 1); @@ -5425,7 +8725,7 @@ x_34 = lean_ctor_get(x_8, 0); lean_inc(x_34); lean_dec(x_8); lean_inc_ref(x_4); -x_35 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_34, x_4, x_5); +x_35 = l_Lean_IR_ExplicitBoxing_BoxParams_boxParams(x_34, x_4, x_5); x_36 = lean_ctor_get(x_35, 0); lean_inc(x_36); x_37 = lean_ctor_get(x_35, 1); @@ -5452,43 +8752,64 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_BoxParams_boxParams_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_7; -x_7 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__0(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec_ref(x_5); -return x_7; -} -} -LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_IR_ExplicitBoxing_visitFnBody___lam__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec_ref(x_7); -return x_9; +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_7 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_8 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_BoxParams_boxParams_spec__1(x_6, x_7, x_3, x_4, x_5); +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_BasicAux_0__mapMonoMImp_go___at___00Lean_IR_ExplicitBoxing_BoxParams_boxParams_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_7; -x_7 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__2(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec_ref(x_5); -return x_7; +lean_object* x_4; +x_4 = l___private_Init_Data_Array_BasicAux_0__mapMonoMImp_go___at___00Lean_IR_ExplicitBoxing_BoxParams_boxParams_spec__0(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_visitFnBody_spec__0___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_IR_ExplicitBoxing_boxedConstDecl(lean_object* x_1, lean_object* x_2) { _start: { -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_7 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_8 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_visitFnBody_spec__0(x_6, x_7, x_3, x_4, x_5); -return x_8; +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +lean_inc(x_1); +x_3 = l_Lean_IR_ExplicitBoxing_mkBoxedName(x_1); +x_4 = l_Lean_IR_ExplicitBoxing_getJPParams___closed__0; +x_5 = l_Lean_IR_IRType_boxed(x_2); +x_6 = lean_unsigned_to_nat(1u); +x_7 = lean_alloc_ctor(6, 2, 0); +lean_ctor_set(x_7, 0, x_1); +lean_ctor_set(x_7, 1, x_4); +x_8 = lean_unsigned_to_nat(2u); +lean_inc(x_2); +x_9 = lean_alloc_ctor(9, 2, 0); +lean_ctor_set(x_9, 0, x_2); +lean_ctor_set(x_9, 1, x_6); +x_10 = l_Lean_IR_ExplicitBoxing_mkCast___closed__1; +lean_inc(x_5); +x_11 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_11, 0, x_8); +lean_ctor_set(x_11, 1, x_5); +lean_ctor_set(x_11, 2, x_9); +lean_ctor_set(x_11, 3, x_10); +x_12 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_12, 0, x_6); +lean_ctor_set(x_12, 1, x_2); +lean_ctor_set(x_12, 2, x_7); +lean_ctor_set(x_12, 3, x_11); +x_13 = lean_box(0); +x_14 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_14, 0, x_3); +lean_ctor_set(x_14, 1, x_4); +lean_ctor_set(x_14, 2, x_5); +lean_ctor_set(x_14, 3, x_12); +lean_ctor_set(x_14, 4, x_13); +return x_14; } } LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ExplicitBoxing_run_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { @@ -5502,7 +8823,7 @@ lean_object* x_13; x_13 = lean_array_uget(x_3, x_4); if (lean_obj_tag(x_13) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_35; lean_object* x_36; uint8_t x_37; uint8_t x_41; x_14 = lean_ctor_get(x_13, 0); lean_inc(x_14); x_15 = lean_ctor_get(x_13, 1); @@ -5525,124 +8846,504 @@ lean_ctor_set(x_24, 1, x_22); lean_ctor_set(x_24, 2, x_23); lean_ctor_set(x_24, 3, x_19); x_25 = l_Lean_IR_LocalContext_addParams(x_21, x_15); +lean_inc_ref(x_2); +lean_inc_ref(x_1); +lean_inc(x_16); +lean_inc(x_14); +x_26 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_26, 0, x_14); +lean_ctor_set(x_26, 1, x_25); +lean_ctor_set(x_26, 2, x_16); +lean_ctor_set(x_26, 3, x_1); +lean_ctor_set(x_26, 4, x_2); +x_27 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_17, x_26, x_24); +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_27, 1); +lean_inc(x_29); +lean_dec_ref(x_27); +x_35 = lean_ctor_get(x_29, 1); +lean_inc_ref(x_35); +lean_dec(x_29); +x_36 = l_Array_append___redArg(x_6, x_35); +lean_dec_ref(x_35); +x_41 = l_Lean_IR_IRType_isStruct(x_16); +if (x_41 == 0) +{ +lean_dec_ref(x_15); +x_37 = x_41; +goto block_40; +} +else +{ +uint8_t x_42; +x_42 = l_Array_isEmpty___redArg(x_15); lean_dec_ref(x_15); +x_37 = x_42; +goto block_40; +} +block_34: +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = l_Lean_IR_Decl_updateBody_x21(x_13, x_28); +x_32 = l_Lean_IR_Decl_elimDead(x_31); +x_33 = lean_array_push(x_30, x_32); +x_7 = x_33; +goto block_11; +} +block_40: +{ +if (x_37 == 0) +{ +lean_dec(x_16); +lean_dec(x_14); +x_30 = x_36; +goto block_34; +} +else +{ +lean_object* x_38; lean_object* x_39; +x_38 = l_Lean_IR_ExplicitBoxing_boxedConstDecl(x_14, x_16); +x_39 = lean_array_push(x_36, x_38); +x_30 = x_39; +goto block_34; +} +} +} +else +{ +lean_object* x_43; +x_43 = lean_array_push(x_6, x_13); +x_7 = x_43; +goto block_11; +} +} +else +{ +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_6; +} +block_11: +{ +size_t x_8; size_t x_9; +x_8 = 1; +x_9 = lean_usize_add(x_4, x_8); +x_4 = x_9; +x_6 = x_7; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_run(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_3 = lean_unsigned_to_nat(0u); +x_4 = l_Lean_IR_ExplicitBoxing_addBoxedVersions___closed__0; +x_5 = lean_array_get_size(x_2); +x_6 = lean_nat_dec_lt(x_3, x_5); +if (x_6 == 0) +{ +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_4; +} +else +{ +uint8_t x_7; +x_7 = lean_nat_dec_le(x_5, x_5); +if (x_7 == 0) +{ +if (x_6 == 0) +{ +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_4; +} +else +{ +size_t x_8; size_t x_9; lean_object* x_10; +x_8 = 0; +x_9 = lean_usize_of_nat(x_5); +lean_inc_ref(x_2); +x_10 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ExplicitBoxing_run_spec__0(x_2, x_1, x_2, x_8, x_9, x_4); +lean_dec_ref(x_2); +return x_10; +} +} +else +{ +size_t x_11; size_t x_12; lean_object* x_13; +x_11 = 0; +x_12 = lean_usize_of_nat(x_5); lean_inc_ref(x_2); +x_13 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ExplicitBoxing_run_spec__0(x_2, x_1, x_2, x_11, x_12, x_4); +lean_dec_ref(x_2); +return x_13; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ExplicitBoxing_run_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ExplicitBoxing_run_spec__0(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec_ref(x_3); +return x_9; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_prepareBoxParams_spec__0(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; uint8_t x_6; +x_5 = lean_array_uget(x_3, x_2); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_16; +x_7 = lean_ctor_get(x_5, 1); +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_array_uset(x_3, x_2, x_8); +x_16 = l_Lean_IR_IRType_isStruct(x_7); +if (x_16 == 0) +{ +x_10 = x_5; +goto block_15; +} +else +{ +lean_object* x_17; +x_17 = l_Lean_IR_IRType_boxed(x_7); +lean_dec(x_7); +lean_ctor_set(x_5, 1, x_17); +x_10 = x_5; +goto block_15; +} +block_15: +{ +size_t x_11; size_t x_12; lean_object* x_13; +x_11 = 1; +x_12 = lean_usize_add(x_2, x_11); +x_13 = lean_array_uset(x_9, x_2, x_10); +x_2 = x_12; +x_3 = x_13; +goto _start; +} +} +else +{ +lean_object* x_18; uint8_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_29; +x_18 = lean_ctor_get(x_5, 0); +x_19 = lean_ctor_get_uint8(x_5, sizeof(void*)*2); +x_20 = lean_ctor_get(x_5, 1); +lean_inc(x_20); +lean_inc(x_18); +lean_dec(x_5); +x_21 = lean_unsigned_to_nat(0u); +x_22 = lean_array_uset(x_3, x_2, x_21); +x_29 = l_Lean_IR_IRType_isStruct(x_20); +if (x_29 == 0) +{ +lean_object* x_30; +x_30 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_30, 0, x_18); +lean_ctor_set(x_30, 1, x_20); +lean_ctor_set_uint8(x_30, sizeof(void*)*2, x_19); +x_23 = x_30; +goto block_28; +} +else +{ +lean_object* x_31; lean_object* x_32; +x_31 = l_Lean_IR_IRType_boxed(x_20); +lean_dec(x_20); +x_32 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_32, 0, x_18); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set_uint8(x_32, sizeof(void*)*2, x_19); +x_23 = x_32; +goto block_28; +} +block_28: +{ +size_t x_24; size_t x_25; lean_object* x_26; +x_24 = 1; +x_25 = lean_usize_add(x_2, x_24); +x_26 = lean_array_uset(x_22, x_2, x_23); +x_2 = x_25; +x_3 = x_26; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_prepareBoxParams_spec__1(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_usize_dec_lt(x_4, x_3); +if (x_6 == 0) +{ +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_5; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_7 = lean_array_uget(x_5, x_4); +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_array_uset(x_5, x_4, x_8); +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; uint8_t x_37; +x_16 = lean_ctor_get(x_7, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_7, 1); +lean_inc_ref(x_17); +x_18 = lean_ctor_get(x_7, 2); +lean_inc(x_18); +x_19 = lean_ctor_get(x_7, 3); +lean_inc(x_19); +x_20 = lean_ctor_get(x_7, 4); +lean_inc(x_20); +if (lean_is_exclusive(x_7)) { + lean_ctor_release(x_7, 0); + lean_ctor_release(x_7, 1); + lean_ctor_release(x_7, 2); + lean_ctor_release(x_7, 3); + lean_ctor_release(x_7, 4); + x_21 = x_7; +} else { + lean_dec_ref(x_7); + x_21 = lean_box(0); +} +lean_inc(x_16); +lean_inc_ref(x_1); +x_22 = l_Lean_isExport(x_1, x_16); +if (x_22 == 0) +{ +x_37 = x_22; +goto block_39; +} +else +{ +uint8_t x_40; +x_40 = l_Lean_IR_IRType_isStruct(x_18); +x_37 = x_40; +goto block_39; +} +block_36: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_box(1); lean_inc_ref(x_1); -x_26 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_26, 0, x_14); -lean_ctor_set(x_26, 1, x_25); -lean_ctor_set(x_26, 2, x_16); -lean_ctor_set(x_26, 3, x_1); -lean_ctor_set(x_26, 4, x_2); -x_27 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_17, x_26, x_24); -x_28 = lean_ctor_get(x_27, 1); +lean_inc_ref(x_2); +lean_inc(x_23); +lean_inc(x_16); +x_25 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_25, 0, x_16); +lean_ctor_set(x_25, 1, x_24); +lean_ctor_set(x_25, 2, x_23); +lean_ctor_set(x_25, 3, x_2); +lean_ctor_set(x_25, 4, x_1); +x_26 = l_Lean_IR_ExplicitBoxing_BoxParams_boxParams(x_19, x_25, x_24); +if (x_22 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 0); -lean_inc(x_29); -lean_dec_ref(x_27); -x_30 = lean_ctor_get(x_28, 1); -lean_inc_ref(x_30); +lean_dec_ref(x_26); +x_29 = l___private_Init_Data_Array_BasicAux_0__mapMonoMImp_go___at___00Lean_IR_ExplicitBoxing_BoxParams_boxParams_spec__0(x_28, x_8, x_17); lean_dec(x_28); -x_31 = l_Array_append___redArg(x_6, x_30); -lean_dec_ref(x_30); -x_32 = l_Lean_IR_Decl_updateBody_x21(x_13, x_29); -x_33 = l_Lean_IR_Decl_elimDead(x_32); -x_34 = lean_array_push(x_31, x_33); -x_7 = x_34; -goto block_11; +if (lean_is_scalar(x_21)) { + x_30 = lean_alloc_ctor(0, 5, 0); +} else { + x_30 = x_21; +} +lean_ctor_set(x_30, 0, x_16); +lean_ctor_set(x_30, 1, x_29); +lean_ctor_set(x_30, 2, x_23); +lean_ctor_set(x_30, 3, x_27); +lean_ctor_set(x_30, 4, x_20); +x_10 = x_30; +goto block_15; } else { -lean_object* x_35; -x_35 = lean_array_push(x_6, x_13); -x_7 = x_35; -goto block_11; +lean_object* x_31; size_t x_32; size_t x_33; lean_object* x_34; lean_object* x_35; +x_31 = lean_ctor_get(x_26, 0); +lean_inc(x_31); +lean_dec_ref(x_26); +x_32 = lean_array_size(x_17); +x_33 = 0; +x_34 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_prepareBoxParams_spec__0(x_32, x_33, x_17); +if (lean_is_scalar(x_21)) { + x_35 = lean_alloc_ctor(0, 5, 0); +} else { + x_35 = x_21; +} +lean_ctor_set(x_35, 0, x_16); +lean_ctor_set(x_35, 1, x_34); +lean_ctor_set(x_35, 2, x_23); +lean_ctor_set(x_35, 3, x_31); +lean_ctor_set(x_35, 4, x_20); +x_10 = x_35; +goto block_15; } } -else +block_39: { -lean_dec_ref(x_2); -lean_dec_ref(x_1); -return x_6; +if (x_37 == 0) +{ +x_23 = x_18; +goto block_36; } -block_11: +else { -size_t x_8; size_t x_9; -x_8 = 1; -x_9 = lean_usize_add(x_4, x_8); -x_4 = x_9; -x_6 = x_7; -goto _start; +lean_object* x_38; +x_38 = l_Lean_IR_IRType_boxed(x_18); +lean_dec(x_18); +x_23 = x_38; +goto block_36; } } } -LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_run(lean_object* x_1, lean_object* x_2) { -_start: +else { -lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; -x_3 = lean_unsigned_to_nat(0u); -x_4 = l_Lean_IR_ExplicitBoxing_addBoxedVersions___closed__0; -x_5 = lean_array_get_size(x_2); -x_6 = lean_nat_dec_lt(x_3, x_5); -if (x_6 == 0) +uint8_t x_41; +x_41 = !lean_is_exclusive(x_7); +if (x_41 == 0) { -lean_object* x_7; -lean_dec_ref(x_2); -x_7 = l_Lean_IR_ExplicitBoxing_addBoxedVersions(x_1, x_4); -return x_7; +lean_object* x_42; lean_object* x_43; size_t x_44; size_t x_45; lean_object* x_46; uint8_t x_47; +x_42 = lean_ctor_get(x_7, 1); +x_43 = lean_ctor_get(x_7, 2); +x_44 = lean_array_size(x_42); +x_45 = 0; +x_46 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_prepareBoxParams_spec__0(x_44, x_45, x_42); +x_47 = l_Lean_IR_IRType_isStruct(x_43); +if (x_47 == 0) +{ +lean_ctor_set(x_7, 1, x_46); +x_10 = x_7; +goto block_15; } else { -uint8_t x_8; -x_8 = lean_nat_dec_le(x_5, x_5); -if (x_8 == 0) +lean_object* x_48; +x_48 = l_Lean_IR_IRType_boxed(x_43); +lean_dec(x_43); +lean_ctor_set(x_7, 2, x_48); +lean_ctor_set(x_7, 1, x_46); +x_10 = x_7; +goto block_15; +} +} +else { -if (x_6 == 0) +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; size_t x_53; size_t x_54; lean_object* x_55; uint8_t x_56; +x_49 = lean_ctor_get(x_7, 0); +x_50 = lean_ctor_get(x_7, 1); +x_51 = lean_ctor_get(x_7, 2); +x_52 = lean_ctor_get(x_7, 3); +lean_inc(x_52); +lean_inc(x_51); +lean_inc(x_50); +lean_inc(x_49); +lean_dec(x_7); +x_53 = lean_array_size(x_50); +x_54 = 0; +x_55 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_prepareBoxParams_spec__0(x_53, x_54, x_50); +x_56 = l_Lean_IR_IRType_isStruct(x_51); +if (x_56 == 0) { -lean_object* x_9; -lean_dec_ref(x_2); -x_9 = l_Lean_IR_ExplicitBoxing_addBoxedVersions(x_1, x_4); -return x_9; +lean_object* x_57; +x_57 = lean_alloc_ctor(1, 4, 0); +lean_ctor_set(x_57, 0, x_49); +lean_ctor_set(x_57, 1, x_55); +lean_ctor_set(x_57, 2, x_51); +lean_ctor_set(x_57, 3, x_52); +x_10 = x_57; +goto block_15; } else { -size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; -x_10 = 0; -x_11 = lean_usize_of_nat(x_5); -lean_inc_ref(x_1); -lean_inc_ref(x_2); -x_12 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ExplicitBoxing_run_spec__0(x_2, x_1, x_2, x_10, x_11, x_4); -lean_dec_ref(x_2); -x_13 = l_Lean_IR_ExplicitBoxing_addBoxedVersions(x_1, x_12); -return x_13; +lean_object* x_58; lean_object* x_59; +x_58 = l_Lean_IR_IRType_boxed(x_51); +lean_dec(x_51); +x_59 = lean_alloc_ctor(1, 4, 0); +lean_ctor_set(x_59, 0, x_49); +lean_ctor_set(x_59, 1, x_55); +lean_ctor_set(x_59, 2, x_58); +lean_ctor_set(x_59, 3, x_52); +x_10 = x_59; +goto block_15; } } -else +} +block_15: { -size_t x_14; size_t x_15; lean_object* x_16; lean_object* x_17; -x_14 = 0; -x_15 = lean_usize_of_nat(x_5); -lean_inc_ref(x_1); +size_t x_11; size_t x_12; lean_object* x_13; +x_11 = 1; +x_12 = lean_usize_add(x_4, x_11); +x_13 = lean_array_uset(x_9, x_4, x_10); +x_4 = x_12; +x_5 = x_13; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_prepareBoxParams(lean_object* x_1, lean_object* x_2) { +_start: +{ +size_t x_3; size_t x_4; lean_object* x_5; +x_3 = lean_array_size(x_2); +x_4 = 0; lean_inc_ref(x_2); -x_16 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ExplicitBoxing_run_spec__0(x_2, x_1, x_2, x_14, x_15, x_4); -lean_dec_ref(x_2); -x_17 = l_Lean_IR_ExplicitBoxing_addBoxedVersions(x_1, x_16); -return x_17; +x_5 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_prepareBoxParams_spec__1(x_1, x_2, x_3, x_4, x_2); +return x_5; } } +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_prepareBoxParams_spec__0___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___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_prepareBoxParams_spec__0(x_4, x_5, x_3); +return x_6; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ExplicitBoxing_run_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_prepareBoxParams_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -size_t x_7; size_t x_8; lean_object* x_9; +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); x_7 = lean_unbox_usize(x_4); lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ExplicitBoxing_run_spec__0(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec_ref(x_3); -return x_9; +x_8 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_prepareBoxParams_spec__1(x_1, x_2, x_6, x_7, x_5); +return x_8; } } LEAN_EXPORT lean_object* l_Lean_IR_explicitBoxing___redArg(lean_object* x_1, lean_object* x_2) { @@ -5686,7 +9387,7 @@ lean_dec(x_2); return x_4; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_1; @@ -5694,7 +9395,7 @@ x_1 = lean_mk_string_unchecked("compiler", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_1; @@ -5702,26 +9403,26 @@ x_1 = lean_mk_string_unchecked("ir", 2, 2); return x_1; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__2_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__2_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("boxing", 6, 6); +x_1 = lean_mk_string_unchecked("box_params", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__3_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__3_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__2_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; -x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; -x_3 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__2_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +x_3 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; x_4 = l_Lean_Name_mkStr3(x_3, x_2, x_1); return x_4; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__4_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__4_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_1; @@ -5729,17 +9430,17 @@ x_1 = lean_mk_string_unchecked("_private", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__5_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__5_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__4_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__4_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; x_2 = lean_box(0); x_3 = l_Lean_Name_str___override(x_2, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__6_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__6_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_1; @@ -5747,17 +9448,17 @@ x_1 = lean_mk_string_unchecked("Lean", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__7_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__7_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__6_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; -x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__5_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__6_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__5_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; x_3 = l_Lean_Name_str___override(x_2, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__8_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__8_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_1; @@ -5765,17 +9466,17 @@ x_1 = lean_mk_string_unchecked("Compiler", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__9_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__9_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__8_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; -x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__7_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__8_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__7_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; x_3 = l_Lean_Name_str___override(x_2, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__10_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__10_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_1; @@ -5783,17 +9484,17 @@ x_1 = lean_mk_string_unchecked("IR", 2, 2); return x_1; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__11_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__11_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__10_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; -x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__9_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__10_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__9_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; x_3 = l_Lean_Name_str___override(x_2, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__12_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__12_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_1; @@ -5801,47 +9502,47 @@ x_1 = lean_mk_string_unchecked("Boxing", 6, 6); return x_1; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__13_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__13_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__12_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; -x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__11_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__12_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__11_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; x_3 = l_Lean_Name_str___override(x_2, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__14_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__14_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(0u); -x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__13_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__13_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; x_3 = l_Lean_Name_num___override(x_2, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__15_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__15_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__6_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; -x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__14_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__6_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__14_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; x_3 = l_Lean_Name_str___override(x_2, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__16_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__16_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__10_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; -x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__15_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__10_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__15_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; x_3 = l_Lean_Name_str___override(x_2, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__17_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__17_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_1; @@ -5849,17 +9550,17 @@ x_1 = lean_mk_string_unchecked("initFn", 6, 6); return x_1; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__18_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__18_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__17_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; -x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__16_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__17_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__16_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; x_3 = l_Lean_Name_str___override(x_2, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__19_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__19_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_1; @@ -5867,67 +9568,67 @@ x_1 = lean_mk_string_unchecked("_@", 2, 2); return x_1; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__20_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__20_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__19_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; -x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__18_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__19_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__18_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; x_3 = l_Lean_Name_str___override(x_2, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__21_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__21_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__6_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; -x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__20_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__6_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__20_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; x_3 = l_Lean_Name_str___override(x_2, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__22_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__22_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__8_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; -x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__21_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__8_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__21_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; x_3 = l_Lean_Name_str___override(x_2, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__23_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__23_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__10_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; -x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__22_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__10_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__22_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; x_3 = l_Lean_Name_str___override(x_2, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__24_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__24_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__12_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; -x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__23_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__12_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__23_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; x_3 = l_Lean_Name_str___override(x_2, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__25_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__25_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(3293845429u); -x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__24_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +x_1 = lean_unsigned_to_nat(1604803799u); +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__24_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; x_3 = l_Lean_Name_num___override(x_2, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__26_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__26_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_1; @@ -5935,17 +9636,17 @@ x_1 = lean_mk_string_unchecked("_hygCtx", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__27_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__27_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__26_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; -x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__25_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__26_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__25_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; x_3 = l_Lean_Name_str___override(x_2, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__28_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__28_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_1; @@ -5953,33 +9654,111 @@ x_1 = lean_mk_string_unchecked("_hyg", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__29_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__29_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__28_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; -x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__27_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__28_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__27_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; x_3 = l_Lean_Name_str___override(x_2, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__30_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__30_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(2u); -x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__29_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__29_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; x_3 = l_Lean_Name_num___override(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_() { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__3_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +x_3 = 1; +x_4 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__30_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2____boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("boxing", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +x_3 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +x_4 = l_Lean_Name_mkStr3(x_3, x_2, x_1); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__2_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(3293845429u); +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__24_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +x_3 = l_Lean_Name_num___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__3_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__26_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__2_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +x_3 = l_Lean_Name_str___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__4_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__28_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__3_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +x_3 = l_Lean_Name_str___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__5_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(2u); +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__4_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +x_3 = l_Lean_Name_num___override(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; x_3 = 1; -x_4 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__30_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; +x_4 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__5_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4); return x_5; } @@ -5992,12 +9771,91 @@ x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn_00___x40_Lean_Compil return x_2; } } +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("boxed_versions", 14, 14); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +x_3 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +x_4 = l_Lean_Name_mkStr3(x_3, x_2, x_1); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__2_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(1405309204u); +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__24_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +x_3 = l_Lean_Name_num___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__3_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__26_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__2_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_; +x_3 = l_Lean_Name_str___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__4_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__28_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__3_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_; +x_3 = l_Lean_Name_str___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__5_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(2u); +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__4_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_; +x_3 = l_Lean_Name_num___override(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_; +x_3 = 1; +x_4 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__5_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_; +x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2____boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_(); +return x_2; +} +} lean_object* initialize_Lean_Runtime(uint8_t builtin); lean_object* initialize_Lean_Compiler_ClosedTermCache(uint8_t builtin); lean_object* initialize_Lean_Compiler_IR_CompilerM(uint8_t builtin); lean_object* initialize_Lean_Compiler_IR_ElimDeadVars(uint8_t builtin); lean_object* initialize_Lean_Compiler_IR_ToIRType(uint8_t builtin); lean_object* initialize_Lean_Data_AssocList(uint8_t builtin); +lean_object* initialize_Lean_Compiler_InitAttr(uint8_t builtin); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Compiler_IR_Boxing(uint8_t builtin) { lean_object * res; @@ -6021,6 +9879,9 @@ lean_dec_ref(res); res = initialize_Lean_Data_AssocList(builtin); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Compiler_InitAttr(builtin); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); l_Lean_IR_ExplicitBoxing_requiresBoxedVersion___closed__0 = _init_l_Lean_IR_ExplicitBoxing_requiresBoxedVersion___closed__0(); lean_mark_persistent(l_Lean_IR_ExplicitBoxing_requiresBoxedVersion___closed__0); l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___at___00Lean_IR_ExplicitBoxing_mkBoxedVersionAux_spec__1___redArg___closed__0 = _init_l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___at___00Lean_IR_ExplicitBoxing_mkBoxedVersionAux_spec__1___redArg___closed__0(); @@ -6031,12 +9892,20 @@ l_Lean_IR_ExplicitBoxing_mkBoxedVersionAux___closed__1 = _init_l_Lean_IR_Explici lean_mark_persistent(l_Lean_IR_ExplicitBoxing_mkBoxedVersionAux___closed__1); l_Lean_IR_ExplicitBoxing_mkBoxedVersionAux___closed__2 = _init_l_Lean_IR_ExplicitBoxing_mkBoxedVersionAux___closed__2(); lean_mark_persistent(l_Lean_IR_ExplicitBoxing_mkBoxedVersionAux___closed__2); +l_panic___at___00Lean_IR_ExplicitBoxing_transformBoxedPaps_spec__0___closed__0 = _init_l_panic___at___00Lean_IR_ExplicitBoxing_transformBoxedPaps_spec__0___closed__0(); +lean_mark_persistent(l_panic___at___00Lean_IR_ExplicitBoxing_transformBoxedPaps_spec__0___closed__0); +l_Lean_IR_ExplicitBoxing_transformBoxedPaps___closed__2 = _init_l_Lean_IR_ExplicitBoxing_transformBoxedPaps___closed__2(); +lean_mark_persistent(l_Lean_IR_ExplicitBoxing_transformBoxedPaps___closed__2); +l_Lean_IR_ExplicitBoxing_transformBoxedPaps___closed__1 = _init_l_Lean_IR_ExplicitBoxing_transformBoxedPaps___closed__1(); +lean_mark_persistent(l_Lean_IR_ExplicitBoxing_transformBoxedPaps___closed__1); +l_Lean_IR_ExplicitBoxing_transformBoxedPaps___closed__0 = _init_l_Lean_IR_ExplicitBoxing_transformBoxedPaps___closed__0(); +lean_mark_persistent(l_Lean_IR_ExplicitBoxing_transformBoxedPaps___closed__0); +l_Lean_IR_ExplicitBoxing_transformBoxedPaps___closed__3 = _init_l_Lean_IR_ExplicitBoxing_transformBoxedPaps___closed__3(); +lean_mark_persistent(l_Lean_IR_ExplicitBoxing_transformBoxedPaps___closed__3); l_Lean_IR_ExplicitBoxing_addBoxedVersions___closed__0 = _init_l_Lean_IR_ExplicitBoxing_addBoxedVersions___closed__0(); lean_mark_persistent(l_Lean_IR_ExplicitBoxing_addBoxedVersions___closed__0); l_Lean_IR_ExplicitBoxing_getJPParams___closed__0 = _init_l_Lean_IR_ExplicitBoxing_getJPParams___closed__0(); lean_mark_persistent(l_Lean_IR_ExplicitBoxing_getJPParams___closed__0); -l_Lean_IR_ExplicitBoxing_getDecl___closed__0 = _init_l_Lean_IR_ExplicitBoxing_getDecl___closed__0(); -lean_mark_persistent(l_Lean_IR_ExplicitBoxing_getDecl___closed__0); l_Lean_IR_ExplicitBoxing_mkCast___closed__0 = _init_l_Lean_IR_ExplicitBoxing_mkCast___closed__0(); lean_mark_persistent(l_Lean_IR_ExplicitBoxing_mkCast___closed__0); l_Lean_IR_ExplicitBoxing_mkCast___closed__1 = _init_l_Lean_IR_ExplicitBoxing_mkCast___closed__1(); @@ -6045,37 +9914,112 @@ l_Lean_IR_ExplicitBoxing_mkCast___closed__2 = _init_l_Lean_IR_ExplicitBoxing_mkC lean_mark_persistent(l_Lean_IR_ExplicitBoxing_mkCast___closed__2); l_Lean_IR_ExplicitBoxing_mkCast___closed__3 = _init_l_Lean_IR_ExplicitBoxing_mkCast___closed__3(); lean_mark_persistent(l_Lean_IR_ExplicitBoxing_mkCast___closed__3); -l_Lean_IR_ExplicitBoxing_mkCast___closed__4 = _init_l_Lean_IR_ExplicitBoxing_mkCast___closed__4(); -lean_mark_persistent(l_Lean_IR_ExplicitBoxing_mkCast___closed__4); l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux___closed__0 = _init_l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux___closed__0(); lean_mark_persistent(l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux___closed__0); l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux___closed__1 = _init_l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux___closed__1(); lean_mark_persistent(l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux___closed__1); -l_Lean_IR_ExplicitBoxing_visitVDeclExpr___closed__0 = _init_l_Lean_IR_ExplicitBoxing_visitVDeclExpr___closed__0(); -lean_mark_persistent(l_Lean_IR_ExplicitBoxing_visitVDeclExpr___closed__0); -l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__0 = _init_l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__0(); -lean_mark_persistent(l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__0); -l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__1 = _init_l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__1(); -lean_mark_persistent(l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__1); -l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__2 = _init_l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__2(); -lean_mark_persistent(l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__2); -l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__3 = _init_l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__3(); -lean_mark_persistent(l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__3); -l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__4 = _init_l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__4(); -lean_mark_persistent(l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__4); -l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__5 = _init_l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__5(); -lean_mark_persistent(l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__5); -l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__6 = _init_l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__6(); -lean_mark_persistent(l_panic___at___00Lean_IR_ExplicitBoxing_tryCorrectVDeclType_spec__0___closed__6); +l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__0 = _init_l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__0(); +lean_mark_persistent(l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__0); +l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__1 = _init_l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__1(); +lean_mark_persistent(l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__1); +l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__2 = _init_l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__2(); +lean_mark_persistent(l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__2); +l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__3 = _init_l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__3(); +lean_mark_persistent(l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__3); +l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__4 = _init_l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__4(); +lean_mark_persistent(l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__4); +l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__5 = _init_l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__5(); +lean_mark_persistent(l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__5); +l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__6 = _init_l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__6(); +lean_mark_persistent(l_panic___at___00Lean_IR_ExplicitBoxing_visitCtorExpr_spec__0___closed__6); +l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__0 = _init_l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__0(); +lean_mark_persistent(l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__0); +l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__1 = _init_l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__1(); +lean_mark_persistent(l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__1); +l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__2 = _init_l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__2(); +lean_mark_persistent(l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__2); +l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__3 = _init_l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__3(); +lean_mark_persistent(l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__3); +l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__4 = _init_l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__4(); +lean_mark_persistent(l_Lean_IR_ExplicitBoxing_visitCtorExpr___closed__4); l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__0 = _init_l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__0(); lean_mark_persistent(l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__0); l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__1 = _init_l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__1(); lean_mark_persistent(l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__1); l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__2 = _init_l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__2(); lean_mark_persistent(l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__2); -l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__3 = _init_l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__3(); -lean_mark_persistent(l_Lean_IR_ExplicitBoxing_tryCorrectVDeclType___closed__3); -l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_(); +l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___closed__0 = _init_l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___closed__0(); +lean_mark_persistent(l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___closed__0); +l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___closed__1 = _init_l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___closed__1(); +lean_mark_persistent(l_Lean_IR_ExplicitBoxing_BoxParams_visitVDeclExpr___closed__1); +l_panic___at___00Lean_IR_ExplicitBoxing_BoxParams_boxParams_spec__2___closed__0 = _init_l_panic___at___00Lean_IR_ExplicitBoxing_BoxParams_boxParams_spec__2___closed__0(); +lean_mark_persistent(l_panic___at___00Lean_IR_ExplicitBoxing_BoxParams_boxParams_spec__2___closed__0); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__2_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__2_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__2_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__3_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__3_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__3_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__4_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__4_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__4_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__5_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__5_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__5_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__6_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__6_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__6_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__7_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__7_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__7_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__8_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__8_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__8_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__9_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__9_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__9_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__10_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__10_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__10_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__11_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__11_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__11_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__12_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__12_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__12_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__13_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__13_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__13_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__14_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__14_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__14_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__15_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__15_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__15_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__16_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__16_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__16_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__17_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__17_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__17_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__18_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__18_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__18_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__19_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__19_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__19_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__20_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__20_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__20_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__21_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__21_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__21_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__22_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__22_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__22_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__23_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__23_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__23_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__24_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__24_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__24_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__25_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__25_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__25_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__26_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__26_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__26_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__27_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__27_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__27_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__28_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__28_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__28_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__29_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__29_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__29_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__30_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__30_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__30_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_); +if (builtin) {res = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn_00___x40_Lean_Compiler_IR_Boxing_1604803799____hygCtx___hyg_2_(); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_(); lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_); l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_(); lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_); @@ -6087,59 +10031,24 @@ l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__4_00___x40_Lean_ lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__4_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_); l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__5_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__5_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_(); lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__5_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_); -l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__6_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__6_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__6_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_); -l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__7_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__7_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__7_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_); -l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__8_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__8_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__8_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_); -l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__9_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__9_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__9_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_); -l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__10_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__10_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__10_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_); -l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__11_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__11_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__11_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_); -l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__12_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__12_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__12_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_); -l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__13_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__13_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__13_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_); -l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__14_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__14_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__14_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_); -l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__15_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__15_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__15_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_); -l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__16_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__16_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__16_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_); -l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__17_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__17_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__17_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_); -l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__18_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__18_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__18_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_); -l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__19_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__19_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__19_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_); -l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__20_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__20_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__20_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_); -l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__21_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__21_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__21_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_); -l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__22_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__22_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__22_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_); -l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__23_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__23_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__23_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_); -l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__24_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__24_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__24_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_); -l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__25_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__25_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__25_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_); -l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__26_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__26_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__26_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_); -l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__27_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__27_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__27_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_); -l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__28_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__28_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__28_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_); -l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__29_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__29_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__29_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_); -l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__30_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__30_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__30_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_); if (builtin) {res = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn_00___x40_Lean_Compiler_IR_Boxing_3293845429____hygCtx___hyg_2_(); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +}l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__2_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__2_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__2_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__3_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__3_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__3_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__4_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__4_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__4_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__5_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__5_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn___closed__5_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_); +if (builtin) {res = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_initFn_00___x40_Lean_Compiler_IR_Boxing_1405309204____hygCtx___hyg_2_(); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Compiler/IR/Checker.c b/stage0/stdlib/Lean/Compiler/IR/Checker.c index 023441e4cf3a..77329490cd4c 100644 --- a/stage0/stdlib/Lean/Compiler/IR/Checker.c +++ b/stage0/stdlib/Lean/Compiler/IR/Checker.c @@ -13,6 +13,8 @@ #ifdef __cplusplus extern "C" { #endif +uint8_t l_Lean_IR_IRType_isScalarOrStruct(lean_object*); +static lean_object* l_Lean_IR_Checker_checkEqTypes___closed__2; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_Checker_checkFnBody_spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_Format_pretty(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_contains___at___00Lean_IR_Checker_markIndex_spec__0___boxed(lean_object*, lean_object*, lean_object*); @@ -28,6 +30,7 @@ static lean_object* l_Lean_IR_Checker_maxCtorFields___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_get_max_ctor_scalars_size(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_checkDecls_spec__0(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_Checker_checkEqTypes___closed__1; LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkObjVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkObjVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -40,13 +43,14 @@ uint8_t l_Lean_IR_instBEqIRType_beq(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Std_DTreeMap_Internal_Impl_contains___at___00Lean_IR_Checker_markIndex_spec__0(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Checker_throwCheckerError___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_get_max_ctor_fields(lean_object*); +static lean_object* l_Lean_IR_Checker_checkObjOrStructType___closed__0; +uint8_t l_Lean_IR_IRType_isObjOrStruct(lean_object*); static lean_object* l_Lean_IR_Checker_checkType___closed__0; LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_IR_Checker_throwCheckerError_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_get_max_ctor_tag(lean_object*); lean_object* l_ReaderT_instApplicativeOfMonad___redArg___lam__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_Checker_checkFullApp___closed__1; lean_object* l_Lean_stringToMessageData(lean_object*); -static lean_object* l_Lean_IR_Checker_checkType___closed__2; static lean_object* l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_Checker_throwCheckerError_spec__0_spec__0___closed__6; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_Checker_checkFnBody_spec__0(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkExpr___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -58,6 +62,7 @@ uint8_t l_Lean_IR_LocalContext_isLocalVar(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkExpr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Checker_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_Checker_throwCheckerError_spec__0_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkObjOrStructType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_IR_Checker_throwCheckerError_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Nat_reprFast(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_checkDecl(lean_object*, lean_object*, lean_object*, lean_object*); @@ -83,8 +88,8 @@ lean_object* l_instMonadEST(lean_object*, lean_object*); static lean_object* l_Lean_IR_Checker_throwCheckerError___redArg___closed__1; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_Checker_checkFnBody_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_IR_Checker_checkObjType___closed__0; -static lean_object* l_Lean_IR_Checker_checkType___closed__1; LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkJP___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkScalarOrStructType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_IR_Decl_name(lean_object*); lean_object* l_Lean_IR_LocalContext_addParam(lean_object*, lean_object*); lean_object* l_ReaderT_instApplicativeOfMonad___redArg___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -118,6 +123,7 @@ static lean_object* l_Lean_IR_Checker_checkVar___closed__1; static lean_object* l_Lean_IR_Checker_maxCtorTag___closed__0; lean_object* l_Lean_Name_toStringWithToken___at___00Lean_Name_toString_spec__0(lean_object*, uint8_t); static lean_object* l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_Checker_throwCheckerError_spec__0_spec__0___closed__1; +static lean_object* l_Lean_IR_Checker_checkExpr___closed__7; LEAN_EXPORT lean_object* l_Lean_IR_Checker_markJP___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_Checker_throwCheckerError___redArg___closed__0; LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_insert___at___00Lean_IR_Checker_markIndex_spec__1___redArg(lean_object*, lean_object*, lean_object*); @@ -126,10 +132,12 @@ static lean_object* l_Lean_IR_Checker_checkExpr___closed__5; LEAN_EXPORT lean_object* l_Lean_IR_Checker_markIndex(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_Checker_checkFullApp___closed__3; LEAN_EXPORT lean_object* l_Lean_IR_Checker_withParams___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkScalarOrStructVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkFullApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Checker_maxCtorTag; LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_contains___at___00Lean_IR_Checker_markIndex_spec__0___redArg___boxed(lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkObjOrStructVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkArgs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkVarType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); @@ -140,15 +148,16 @@ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkPartialApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Checker_markIndex___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_Checker_checkExpr___closed__9; LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkEqTypes(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkObjOrStructType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instFunctorOfMonad___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkScalarVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_Checker_maxCtorScalarsSize___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkFnBody(lean_object*, lean_object*, 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_EXPORT lean_object* l_Lean_IR_Checker_checkScalarVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkScalarOrStructVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_mul(lean_object*, lean_object*); lean_object* l_ReaderT_instApplicativeOfMonad___redArg___lam__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_IR_instToFormatIRType___private__1(lean_object*); @@ -194,7 +203,10 @@ LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe uint8_t l_Lean_IR_IRType_isScalar(lean_object*); static lean_object* l_Lean_IR_Checker_checkExpr___closed__2; LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkArgs___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkScalarOrStructType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_checkDecls_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkObjOrStructVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_Checker_checkExpr___closed__8; LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkVarType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkObjType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Checker_getMaxCtorFields___boxed(lean_object* x_1) { @@ -2126,7 +2138,23 @@ static lean_object* _init_l_Lean_IR_Checker_checkEqTypes___closed__0() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("unexpected type '{ty₁}' != '{ty₂}'", 38, 34); +x_1 = lean_mk_string_unchecked("unexpected type '", 17, 17); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_Checker_checkEqTypes___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Format_defWidth; +return x_1; +} +} +static lean_object* _init_l_Lean_IR_Checker_checkEqTypes___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("' != '", 6, 6); return x_1; } } @@ -2134,21 +2162,39 @@ LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkEqTypes(lean_object* x_1, lean_o _start: { uint8_t x_8; +lean_inc(x_2); +lean_inc(x_1); x_8 = l_Lean_IR_instBEqIRType_beq(x_1, x_2); if (x_8 == 0) { -lean_object* x_9; lean_object* x_10; +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_9 = l_Lean_IR_Checker_checkEqTypes___closed__0; -x_10 = l_Lean_IR_Checker_throwCheckerError___redArg(x_9, x_3, x_4, x_5, x_6); -return x_10; +x_10 = l_Lean_IR_instToFormatIRType___private__1(x_1); +x_11 = l_Lean_IR_Checker_checkEqTypes___closed__1; +x_12 = lean_unsigned_to_nat(0u); +x_13 = l_Std_Format_pretty(x_10, x_11, x_12, x_12); +x_14 = lean_string_append(x_9, x_13); +lean_dec_ref(x_13); +x_15 = l_Lean_IR_Checker_checkEqTypes___closed__2; +x_16 = lean_string_append(x_14, x_15); +x_17 = l_Lean_IR_instToFormatIRType___private__1(x_2); +x_18 = l_Std_Format_pretty(x_17, x_11, x_12, x_12); +x_19 = lean_string_append(x_16, x_18); +lean_dec_ref(x_18); +x_20 = l_Lean_IR_Checker_checkVar___closed__2; +x_21 = lean_string_append(x_19, x_20); +x_22 = l_Lean_IR_Checker_throwCheckerError___redArg(x_21, x_3, x_4, x_5, x_6); +return x_22; } else { -lean_object* x_11; lean_object* x_12; -x_11 = lean_box(0); -x_12 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_12, 0, x_11); -return x_12; +lean_object* x_23; lean_object* x_24; +lean_dec(x_2); +lean_dec(x_1); +x_23 = lean_box(0); +x_24 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_24, 0, x_23); +return x_24; } } } @@ -2168,22 +2214,6 @@ static lean_object* _init_l_Lean_IR_Checker_checkType___closed__0() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("unexpected type '", 17, 17); -return x_1; -} -} -static lean_object* _init_l_Lean_IR_Checker_checkType___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = l_Std_Format_defWidth; -return x_1; -} -} -static lean_object* _init_l_Lean_IR_Checker_checkType___closed__2() { -_start: -{ -lean_object* x_1; x_1 = lean_mk_string_unchecked(", ", 2, 2); return x_1; } @@ -2198,9 +2228,9 @@ x_10 = lean_unbox(x_9); if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_11 = l_Lean_IR_Checker_checkType___closed__0; +x_11 = l_Lean_IR_Checker_checkEqTypes___closed__0; x_12 = l_Lean_IR_instToFormatIRType___private__1(x_1); -x_13 = l_Lean_IR_Checker_checkType___closed__1; +x_13 = l_Lean_IR_Checker_checkEqTypes___closed__1; x_14 = lean_unsigned_to_nat(0u); x_15 = l_Std_Format_pretty(x_12, x_13, x_14, x_14); x_16 = lean_string_append(x_11, x_15); @@ -2211,7 +2241,7 @@ if (lean_obj_tag(x_3) == 1) { lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; x_19 = lean_ctor_get(x_3, 0); -x_20 = l_Lean_IR_Checker_checkType___closed__2; +x_20 = l_Lean_IR_Checker_checkType___closed__0; x_21 = lean_string_append(x_18, x_20); x_22 = lean_string_append(x_21, x_19); x_23 = l_Lean_IR_Checker_throwCheckerError___redArg(x_22, x_4, x_5, x_6, x_7); @@ -2265,16 +2295,16 @@ if (x_7 == 0) { lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; x_8 = l_Lean_IR_Checker_checkObjType___closed__0; -x_9 = l_Lean_IR_Checker_checkType___closed__0; +x_9 = l_Lean_IR_Checker_checkEqTypes___closed__0; x_10 = l_Lean_IR_instToFormatIRType___private__1(x_1); -x_11 = l_Lean_IR_Checker_checkType___closed__1; +x_11 = l_Lean_IR_Checker_checkEqTypes___closed__1; x_12 = lean_unsigned_to_nat(0u); x_13 = l_Std_Format_pretty(x_10, x_11, x_12, x_12); x_14 = lean_string_append(x_9, x_13); lean_dec_ref(x_13); x_15 = l_Lean_IR_Checker_checkVar___closed__2; x_16 = lean_string_append(x_14, x_15); -x_17 = l_Lean_IR_Checker_checkType___closed__2; +x_17 = l_Lean_IR_Checker_checkType___closed__0; x_18 = lean_string_append(x_16, x_17); x_19 = lean_string_append(x_18, x_8); x_20 = l_Lean_IR_Checker_throwCheckerError___redArg(x_19, x_2, x_3, x_4, x_5); @@ -2303,6 +2333,61 @@ lean_dec_ref(x_2); return x_7; } } +static lean_object* _init_l_Lean_IR_Checker_checkObjOrStructType___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("object or struct expected", 25, 25); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkObjOrStructType(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +uint8_t x_7; +x_7 = l_Lean_IR_IRType_isObjOrStruct(x_1); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_8 = l_Lean_IR_Checker_checkObjOrStructType___closed__0; +x_9 = l_Lean_IR_Checker_checkEqTypes___closed__0; +x_10 = l_Lean_IR_instToFormatIRType___private__1(x_1); +x_11 = l_Lean_IR_Checker_checkEqTypes___closed__1; +x_12 = lean_unsigned_to_nat(0u); +x_13 = l_Std_Format_pretty(x_10, x_11, x_12, x_12); +x_14 = lean_string_append(x_9, x_13); +lean_dec_ref(x_13); +x_15 = l_Lean_IR_Checker_checkVar___closed__2; +x_16 = lean_string_append(x_14, x_15); +x_17 = l_Lean_IR_Checker_checkType___closed__0; +x_18 = lean_string_append(x_16, x_17); +x_19 = lean_string_append(x_18, x_8); +x_20 = l_Lean_IR_Checker_throwCheckerError___redArg(x_19, x_2, x_3, x_4, x_5); +return x_20; +} +else +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_1); +x_21 = lean_box(0); +x_22 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_22, 0, x_21); +return x_22; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkObjOrStructType___boxed(lean_object* x_1, 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_IR_Checker_checkObjOrStructType(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec(x_3); +lean_dec_ref(x_2); +return x_7; +} +} static lean_object* _init_l_Lean_IR_Checker_checkScalarType___closed__0() { _start: { @@ -2320,16 +2405,16 @@ if (x_7 == 0) { lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; x_8 = l_Lean_IR_Checker_checkScalarType___closed__0; -x_9 = l_Lean_IR_Checker_checkType___closed__0; +x_9 = l_Lean_IR_Checker_checkEqTypes___closed__0; x_10 = l_Lean_IR_instToFormatIRType___private__1(x_1); -x_11 = l_Lean_IR_Checker_checkType___closed__1; +x_11 = l_Lean_IR_Checker_checkEqTypes___closed__1; x_12 = lean_unsigned_to_nat(0u); x_13 = l_Std_Format_pretty(x_10, x_11, x_12, x_12); x_14 = lean_string_append(x_9, x_13); lean_dec_ref(x_13); x_15 = l_Lean_IR_Checker_checkVar___closed__2; x_16 = lean_string_append(x_14, x_15); -x_17 = l_Lean_IR_Checker_checkType___closed__2; +x_17 = l_Lean_IR_Checker_checkType___closed__0; x_18 = lean_string_append(x_16, x_17); x_19 = lean_string_append(x_18, x_8); x_20 = l_Lean_IR_Checker_throwCheckerError___redArg(x_19, x_2, x_3, x_4, x_5); @@ -2358,6 +2443,53 @@ lean_dec_ref(x_2); return x_7; } } +LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkScalarOrStructType(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +uint8_t x_7; +x_7 = l_Lean_IR_IRType_isScalarOrStruct(x_1); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_8 = l_Lean_IR_Checker_checkScalarType___closed__0; +x_9 = l_Lean_IR_Checker_checkEqTypes___closed__0; +x_10 = l_Lean_IR_instToFormatIRType___private__1(x_1); +x_11 = l_Lean_IR_Checker_checkEqTypes___closed__1; +x_12 = lean_unsigned_to_nat(0u); +x_13 = l_Std_Format_pretty(x_10, x_11, x_12, x_12); +x_14 = lean_string_append(x_9, x_13); +lean_dec_ref(x_13); +x_15 = l_Lean_IR_Checker_checkVar___closed__2; +x_16 = lean_string_append(x_14, x_15); +x_17 = l_Lean_IR_Checker_checkType___closed__0; +x_18 = lean_string_append(x_16, x_17); +x_19 = lean_string_append(x_18, x_8); +x_20 = l_Lean_IR_Checker_throwCheckerError___redArg(x_19, x_2, x_3, x_4, x_5); +return x_20; +} +else +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_1); +x_21 = lean_box(0); +x_22 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_22, 0, x_21); +return x_22; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkScalarOrStructType___boxed(lean_object* x_1, 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_IR_Checker_checkScalarOrStructType(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec(x_3); +lean_dec_ref(x_2); +return x_7; +} +} LEAN_EXPORT lean_object* l_Lean_IR_Checker_getType(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { @@ -2434,9 +2566,9 @@ if (x_13 == 0) { lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_free_object(x_9); -x_14 = l_Lean_IR_Checker_checkType___closed__0; +x_14 = l_Lean_IR_Checker_checkEqTypes___closed__0; x_15 = l_Lean_IR_instToFormatIRType___private__1(x_11); -x_16 = l_Lean_IR_Checker_checkType___closed__1; +x_16 = l_Lean_IR_Checker_checkEqTypes___closed__1; x_17 = lean_unsigned_to_nat(0u); x_18 = l_Std_Format_pretty(x_15, x_16, x_17, x_17); x_19 = lean_string_append(x_14, x_18); @@ -2447,7 +2579,7 @@ if (lean_obj_tag(x_3) == 1) { lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; x_22 = lean_ctor_get(x_3, 0); -x_23 = l_Lean_IR_Checker_checkType___closed__2; +x_23 = l_Lean_IR_Checker_checkType___closed__0; x_24 = lean_string_append(x_21, x_23); x_25 = lean_string_append(x_24, x_22); x_26 = l_Lean_IR_Checker_throwCheckerError___redArg(x_25, x_4, x_5, x_6, x_7); @@ -2481,9 +2613,9 @@ x_31 = lean_unbox(x_30); if (x_31 == 0) { lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_32 = l_Lean_IR_Checker_checkType___closed__0; +x_32 = l_Lean_IR_Checker_checkEqTypes___closed__0; x_33 = l_Lean_IR_instToFormatIRType___private__1(x_29); -x_34 = l_Lean_IR_Checker_checkType___closed__1; +x_34 = l_Lean_IR_Checker_checkEqTypes___closed__1; x_35 = lean_unsigned_to_nat(0u); x_36 = l_Std_Format_pretty(x_33, x_34, x_35, x_35); x_37 = lean_string_append(x_32, x_36); @@ -2494,7 +2626,7 @@ if (lean_obj_tag(x_3) == 1) { lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; x_40 = lean_ctor_get(x_3, 0); -x_41 = l_Lean_IR_Checker_checkType___closed__2; +x_41 = l_Lean_IR_Checker_checkType___closed__0; x_42 = lean_string_append(x_39, x_41); x_43 = lean_string_append(x_42, x_40); x_44 = l_Lean_IR_Checker_throwCheckerError___redArg(x_43, x_4, x_5, x_6, x_7); @@ -2572,16 +2704,16 @@ if (x_10 == 0) lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_free_object(x_7); x_11 = l_Lean_IR_Checker_checkObjType___closed__0; -x_12 = l_Lean_IR_Checker_checkType___closed__0; +x_12 = l_Lean_IR_Checker_checkEqTypes___closed__0; x_13 = l_Lean_IR_instToFormatIRType___private__1(x_9); -x_14 = l_Lean_IR_Checker_checkType___closed__1; +x_14 = l_Lean_IR_Checker_checkEqTypes___closed__1; x_15 = lean_unsigned_to_nat(0u); x_16 = l_Std_Format_pretty(x_13, x_14, x_15, x_15); x_17 = lean_string_append(x_12, x_16); lean_dec_ref(x_16); x_18 = l_Lean_IR_Checker_checkVar___closed__2; x_19 = lean_string_append(x_17, x_18); -x_20 = l_Lean_IR_Checker_checkType___closed__2; +x_20 = l_Lean_IR_Checker_checkType___closed__0; x_21 = lean_string_append(x_19, x_20); x_22 = lean_string_append(x_21, x_11); x_23 = l_Lean_IR_Checker_throwCheckerError___redArg(x_22, x_2, x_3, x_4, x_5); @@ -2607,16 +2739,16 @@ if (x_26 == 0) { lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; x_27 = l_Lean_IR_Checker_checkObjType___closed__0; -x_28 = l_Lean_IR_Checker_checkType___closed__0; +x_28 = l_Lean_IR_Checker_checkEqTypes___closed__0; x_29 = l_Lean_IR_instToFormatIRType___private__1(x_25); -x_30 = l_Lean_IR_Checker_checkType___closed__1; +x_30 = l_Lean_IR_Checker_checkEqTypes___closed__1; x_31 = lean_unsigned_to_nat(0u); x_32 = l_Std_Format_pretty(x_29, x_30, x_31, x_31); x_33 = lean_string_append(x_28, x_32); lean_dec_ref(x_32); x_34 = l_Lean_IR_Checker_checkVar___closed__2; x_35 = lean_string_append(x_33, x_34); -x_36 = l_Lean_IR_Checker_checkType___closed__2; +x_36 = l_Lean_IR_Checker_checkType___closed__0; x_37 = lean_string_append(x_35, x_36); x_38 = lean_string_append(x_37, x_27); x_39 = l_Lean_IR_Checker_throwCheckerError___redArg(x_38, x_2, x_3, x_4, x_5); @@ -2666,7 +2798,120 @@ lean_dec_ref(x_2); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkScalarVar(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_IR_Checker_checkObjOrStructVar(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_IR_Checker_getType(x_1, x_2, x_3, x_4, x_5); +if (lean_obj_tag(x_7) == 0) +{ +uint8_t x_8; +x_8 = !lean_is_exclusive(x_7); +if (x_8 == 0) +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_ctor_get(x_7, 0); +x_10 = l_Lean_IR_IRType_isObjOrStruct(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +lean_free_object(x_7); +x_11 = l_Lean_IR_Checker_checkObjOrStructType___closed__0; +x_12 = l_Lean_IR_Checker_checkEqTypes___closed__0; +x_13 = l_Lean_IR_instToFormatIRType___private__1(x_9); +x_14 = l_Lean_IR_Checker_checkEqTypes___closed__1; +x_15 = lean_unsigned_to_nat(0u); +x_16 = l_Std_Format_pretty(x_13, x_14, x_15, x_15); +x_17 = lean_string_append(x_12, x_16); +lean_dec_ref(x_16); +x_18 = l_Lean_IR_Checker_checkVar___closed__2; +x_19 = lean_string_append(x_17, x_18); +x_20 = l_Lean_IR_Checker_checkType___closed__0; +x_21 = lean_string_append(x_19, x_20); +x_22 = lean_string_append(x_21, x_11); +x_23 = l_Lean_IR_Checker_throwCheckerError___redArg(x_22, x_2, x_3, x_4, x_5); +return x_23; +} +else +{ +lean_object* x_24; +lean_dec(x_9); +x_24 = lean_box(0); +lean_ctor_set(x_7, 0, x_24); +return x_7; +} +} +else +{ +lean_object* x_25; uint8_t x_26; +x_25 = lean_ctor_get(x_7, 0); +lean_inc(x_25); +lean_dec(x_7); +x_26 = l_Lean_IR_IRType_isObjOrStruct(x_25); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_27 = l_Lean_IR_Checker_checkObjOrStructType___closed__0; +x_28 = l_Lean_IR_Checker_checkEqTypes___closed__0; +x_29 = l_Lean_IR_instToFormatIRType___private__1(x_25); +x_30 = l_Lean_IR_Checker_checkEqTypes___closed__1; +x_31 = lean_unsigned_to_nat(0u); +x_32 = l_Std_Format_pretty(x_29, x_30, x_31, x_31); +x_33 = lean_string_append(x_28, x_32); +lean_dec_ref(x_32); +x_34 = l_Lean_IR_Checker_checkVar___closed__2; +x_35 = lean_string_append(x_33, x_34); +x_36 = l_Lean_IR_Checker_checkType___closed__0; +x_37 = lean_string_append(x_35, x_36); +x_38 = lean_string_append(x_37, x_27); +x_39 = l_Lean_IR_Checker_throwCheckerError___redArg(x_38, x_2, x_3, x_4, x_5); +return x_39; +} +else +{ +lean_object* x_40; lean_object* x_41; +lean_dec(x_25); +x_40 = lean_box(0); +x_41 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_41, 0, x_40); +return x_41; +} +} +} +else +{ +uint8_t x_42; +x_42 = !lean_is_exclusive(x_7); +if (x_42 == 0) +{ +return x_7; +} +else +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_7, 0); +lean_inc(x_43); +lean_dec(x_7); +x_44 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_44, 0, x_43); +return x_44; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkObjOrStructVar___boxed(lean_object* x_1, 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_IR_Checker_checkObjOrStructVar(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec(x_3); +lean_dec_ref(x_2); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkScalarOrStructVar(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_7; @@ -2679,22 +2924,22 @@ if (x_8 == 0) { lean_object* x_9; uint8_t x_10; x_9 = lean_ctor_get(x_7, 0); -x_10 = l_Lean_IR_IRType_isScalar(x_9); +x_10 = l_Lean_IR_IRType_isScalarOrStruct(x_9); if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_free_object(x_7); x_11 = l_Lean_IR_Checker_checkScalarType___closed__0; -x_12 = l_Lean_IR_Checker_checkType___closed__0; +x_12 = l_Lean_IR_Checker_checkEqTypes___closed__0; x_13 = l_Lean_IR_instToFormatIRType___private__1(x_9); -x_14 = l_Lean_IR_Checker_checkType___closed__1; +x_14 = l_Lean_IR_Checker_checkEqTypes___closed__1; x_15 = lean_unsigned_to_nat(0u); x_16 = l_Std_Format_pretty(x_13, x_14, x_15, x_15); x_17 = lean_string_append(x_12, x_16); lean_dec_ref(x_16); x_18 = l_Lean_IR_Checker_checkVar___closed__2; x_19 = lean_string_append(x_17, x_18); -x_20 = l_Lean_IR_Checker_checkType___closed__2; +x_20 = l_Lean_IR_Checker_checkType___closed__0; x_21 = lean_string_append(x_19, x_20); x_22 = lean_string_append(x_21, x_11); x_23 = l_Lean_IR_Checker_throwCheckerError___redArg(x_22, x_2, x_3, x_4, x_5); @@ -2715,21 +2960,21 @@ lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_7, 0); lean_inc(x_25); lean_dec(x_7); -x_26 = l_Lean_IR_IRType_isScalar(x_25); +x_26 = l_Lean_IR_IRType_isScalarOrStruct(x_25); if (x_26 == 0) { lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; x_27 = l_Lean_IR_Checker_checkScalarType___closed__0; -x_28 = l_Lean_IR_Checker_checkType___closed__0; +x_28 = l_Lean_IR_Checker_checkEqTypes___closed__0; x_29 = l_Lean_IR_instToFormatIRType___private__1(x_25); -x_30 = l_Lean_IR_Checker_checkType___closed__1; +x_30 = l_Lean_IR_Checker_checkEqTypes___closed__1; x_31 = lean_unsigned_to_nat(0u); x_32 = l_Std_Format_pretty(x_29, x_30, x_31, x_31); x_33 = lean_string_append(x_28, x_32); lean_dec_ref(x_32); x_34 = l_Lean_IR_Checker_checkVar___closed__2; x_35 = lean_string_append(x_33, x_34); -x_36 = l_Lean_IR_Checker_checkType___closed__2; +x_36 = l_Lean_IR_Checker_checkType___closed__0; x_37 = lean_string_append(x_35, x_36); x_38 = lean_string_append(x_37, x_27); x_39 = l_Lean_IR_Checker_throwCheckerError___redArg(x_38, x_2, x_3, x_4, x_5); @@ -2767,11 +3012,11 @@ return x_44; } } } -LEAN_EXPORT lean_object* l_Lean_IR_Checker_checkScalarVar___boxed(lean_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_IR_Checker_checkScalarOrStructVar___boxed(lean_object* x_1, 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_IR_Checker_checkScalarVar(x_1, x_2, x_3, x_4, x_5); +x_7 = l_Lean_IR_Checker_checkScalarOrStructVar(x_1, x_2, x_3, x_4, x_5); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); @@ -3053,6 +3298,30 @@ static lean_object* _init_l_Lean_IR_Checker_checkExpr___closed__6() { _start: { lean_object* x_1; +x_1 = lean_mk_string_unchecked("expected constructor index 0 for struct", 39, 39); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_Checker_checkExpr___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid proj constructor index", 30, 30); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_Checker_checkExpr___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("expected struct inside union", 28, 28); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_Checker_checkExpr___closed__9() { +_start: +{ +lean_object* x_1; x_1 = lean_mk_string_unchecked("unexpected IR type '", 20, 20); return x_1; } @@ -3106,7 +3375,7 @@ return x_17; else { lean_object* x_18; -x_18 = l_Lean_IR_Checker_checkObjType(x_1, x_10, x_11, x_12, x_13); +x_18 = l_Lean_IR_Checker_checkObjOrStructType(x_1, x_10, x_11, x_12, x_13); if (lean_obj_tag(x_18) == 0) { lean_object* x_19; @@ -3275,761 +3544,683 @@ return x_73; } case 3: { -lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; x_76 = lean_ctor_get(x_2, 0); lean_inc(x_76); x_77 = lean_ctor_get(x_2, 1); lean_inc(x_77); +x_78 = lean_ctor_get(x_2, 2); +lean_inc(x_78); lean_dec_ref(x_2); -x_78 = l_Lean_IR_Checker_getType(x_77, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_78) == 0) -{ -uint8_t x_79; -x_79 = !lean_is_exclusive(x_78); -if (x_79 == 0) +x_79 = l_Lean_IR_Checker_getType(x_78, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_79) == 0) { -lean_object* x_80; -x_80 = lean_ctor_get(x_78, 0); +lean_object* x_80; lean_object* x_81; +x_80 = lean_ctor_get(x_79, 0); +lean_inc(x_80); +if (lean_is_exclusive(x_79)) { + lean_ctor_release(x_79, 0); + x_81 = x_79; +} else { + lean_dec_ref(x_79); + x_81 = lean_box(0); +} switch (lean_obj_tag(x_80)) { case 7: { -lean_object* x_81; -lean_free_object(x_78); +lean_object* x_82; +lean_dec(x_81); +lean_dec(x_77); lean_dec(x_76); -x_81 = l_Lean_IR_Checker_checkObjType(x_1, x_3, x_4, x_5, x_6); -return x_81; +x_82 = l_Lean_IR_Checker_checkObjOrStructType(x_1, x_3, x_4, x_5, x_6); +return x_82; } case 8: { -lean_object* x_82; -lean_free_object(x_78); +lean_object* x_83; +lean_dec(x_81); +lean_dec(x_77); lean_dec(x_76); -x_82 = l_Lean_IR_Checker_checkObjType(x_1, x_3, x_4, x_5, x_6); -return x_82; +x_83 = l_Lean_IR_Checker_checkObjOrStructType(x_1, x_3, x_4, x_5, x_6); +return x_83; } case 10: { -lean_object* x_83; lean_object* x_84; uint8_t x_85; -x_83 = lean_ctor_get(x_80, 1); -lean_inc_ref(x_83); -lean_dec_ref(x_80); -x_84 = lean_array_get_size(x_83); -x_85 = lean_nat_dec_lt(x_76, x_84); -if (x_85 == 0) -{ -lean_object* x_86; lean_object* x_87; -lean_dec_ref(x_83); -lean_free_object(x_78); -lean_dec(x_76); +lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_97; uint8_t x_98; lean_dec(x_1); -x_86 = l_Lean_IR_Checker_checkExpr___closed__5; -x_87 = l_Lean_IR_Checker_throwCheckerError___redArg(x_86, x_3, x_4, x_5, x_6); -return x_87; -} -else -{ -lean_object* x_88; uint8_t x_89; -x_88 = lean_array_fget(x_83, x_76); +x_84 = lean_ctor_get(x_80, 1); +lean_inc_ref(x_84); +lean_dec_ref(x_80); +x_97 = lean_unsigned_to_nat(0u); +x_98 = lean_nat_dec_eq(x_76, x_97); lean_dec(x_76); -lean_dec_ref(x_83); -x_89 = l_Lean_IR_instBEqIRType_beq(x_88, x_1); -if (x_89 == 0) +if (x_98 == 0) { -lean_object* x_90; lean_object* x_91; -lean_free_object(x_78); -x_90 = l_Lean_IR_Checker_checkEqTypes___closed__0; -x_91 = l_Lean_IR_Checker_throwCheckerError___redArg(x_90, x_3, x_4, x_5, x_6); -return x_91; +lean_object* x_99; lean_object* x_100; +lean_dec_ref(x_84); +lean_dec(x_81); +lean_dec(x_77); +x_99 = l_Lean_IR_Checker_checkExpr___closed__6; +x_100 = l_Lean_IR_Checker_throwCheckerError___redArg(x_99, x_3, x_4, x_5, x_6); +return x_100; } else { -lean_object* x_92; -x_92 = lean_box(0); -lean_ctor_set(x_78, 0, x_92); -return x_78; +x_85 = x_3; +x_86 = x_4; +x_87 = x_5; +x_88 = x_6; +x_89 = lean_box(0); +goto block_96; } -} -} -case 11: +block_96: { -lean_object* x_93; lean_object* x_94; uint8_t x_95; -x_93 = lean_ctor_get(x_80, 1); -lean_inc_ref(x_93); -lean_dec_ref(x_80); -x_94 = lean_array_get_size(x_93); -x_95 = lean_nat_dec_lt(x_76, x_94); -if (x_95 == 0) +lean_object* x_90; uint8_t x_91; +x_90 = lean_array_get_size(x_84); +lean_dec_ref(x_84); +x_91 = lean_nat_dec_le(x_90, x_77); +lean_dec(x_77); +if (x_91 == 0) { -lean_object* x_96; lean_object* x_97; -lean_dec_ref(x_93); -lean_free_object(x_78); -lean_dec(x_76); -lean_dec(x_1); -x_96 = l_Lean_IR_Checker_checkExpr___closed__5; -x_97 = l_Lean_IR_Checker_throwCheckerError___redArg(x_96, x_3, x_4, x_5, x_6); -return x_97; +lean_object* x_92; lean_object* x_93; +x_92 = lean_box(0); +if (lean_is_scalar(x_81)) { + x_93 = lean_alloc_ctor(0, 1, 0); +} else { + x_93 = x_81; } -else -{ -lean_object* x_98; uint8_t x_99; -x_98 = lean_array_fget(x_93, x_76); -lean_dec(x_76); -lean_dec_ref(x_93); -x_99 = l_Lean_IR_instBEqIRType_beq(x_98, x_1); -if (x_99 == 0) -{ -lean_object* x_100; lean_object* x_101; -lean_free_object(x_78); -x_100 = l_Lean_IR_Checker_checkEqTypes___closed__0; -x_101 = l_Lean_IR_Checker_throwCheckerError___redArg(x_100, x_3, x_4, x_5, x_6); -return x_101; +lean_ctor_set(x_93, 0, x_92); +return x_93; } else { -lean_object* x_102; -x_102 = lean_box(0); -lean_ctor_set(x_78, 0, x_102); -return x_78; +lean_object* x_94; lean_object* x_95; +lean_dec(x_81); +x_94 = l_Lean_IR_Checker_checkExpr___closed__5; +x_95 = l_Lean_IR_Checker_throwCheckerError___redArg(x_94, x_85, x_86, x_87, x_88); +return x_95; } } } -case 12: +case 11: { -lean_object* x_103; -lean_dec(x_76); +lean_object* x_101; lean_object* x_102; uint8_t x_103; lean_dec(x_1); -x_103 = lean_box(0); -lean_ctor_set(x_78, 0, x_103); -return x_78; -} -default: +x_101 = lean_ctor_get(x_80, 1); +lean_inc_ref(x_101); +lean_dec_ref(x_80); +x_102 = lean_array_get_size(x_101); +x_103 = lean_nat_dec_lt(x_76, x_102); +if (x_103 == 0) { -lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; -lean_free_object(x_78); +lean_object* x_104; lean_object* x_105; +lean_dec_ref(x_101); +lean_dec(x_81); +lean_dec(x_77); lean_dec(x_76); -lean_dec(x_1); -x_104 = l_Lean_IR_Checker_checkExpr___closed__6; -x_105 = l_Lean_IR_instToFormatIRType___private__1(x_80); -x_106 = l_Lean_IR_Checker_checkType___closed__1; -x_107 = lean_unsigned_to_nat(0u); -x_108 = l_Std_Format_pretty(x_105, x_106, x_107, x_107); -x_109 = lean_string_append(x_104, x_108); -lean_dec_ref(x_108); -x_110 = l_Lean_IR_Checker_checkVar___closed__2; -x_111 = lean_string_append(x_109, x_110); -x_112 = l_Lean_IR_Checker_throwCheckerError___redArg(x_111, x_3, x_4, x_5, x_6); -return x_112; -} -} +x_104 = l_Lean_IR_Checker_checkExpr___closed__7; +x_105 = l_Lean_IR_Checker_throwCheckerError___redArg(x_104, x_3, x_4, x_5, x_6); +return x_105; } else { -lean_object* x_113; -x_113 = lean_ctor_get(x_78, 0); -lean_inc(x_113); -lean_dec(x_78); -switch (lean_obj_tag(x_113)) { -case 7: -{ -lean_object* x_114; +lean_object* x_106; +x_106 = lean_array_fget(x_101, x_76); lean_dec(x_76); -x_114 = l_Lean_IR_Checker_checkObjType(x_1, x_3, x_4, x_5, x_6); -return x_114; -} -case 8: -{ -lean_object* x_115; -lean_dec(x_76); -x_115 = l_Lean_IR_Checker_checkObjType(x_1, x_3, x_4, x_5, x_6); -return x_115; -} -case 10: -{ -lean_object* x_116; lean_object* x_117; uint8_t x_118; -x_116 = lean_ctor_get(x_113, 1); -lean_inc_ref(x_116); -lean_dec_ref(x_113); -x_117 = lean_array_get_size(x_116); -x_118 = lean_nat_dec_lt(x_76, x_117); -if (x_118 == 0) +lean_dec_ref(x_101); +if (lean_obj_tag(x_106) == 10) +{ +lean_object* x_107; lean_object* x_108; uint8_t x_109; +x_107 = lean_ctor_get(x_106, 1); +lean_inc_ref(x_107); +lean_dec_ref(x_106); +x_108 = lean_array_get_size(x_107); +lean_dec_ref(x_107); +x_109 = lean_nat_dec_le(x_108, x_77); +lean_dec(x_77); +if (x_109 == 0) { -lean_object* x_119; lean_object* x_120; -lean_dec_ref(x_116); -lean_dec(x_76); -lean_dec(x_1); -x_119 = l_Lean_IR_Checker_checkExpr___closed__5; -x_120 = l_Lean_IR_Checker_throwCheckerError___redArg(x_119, x_3, x_4, x_5, x_6); -return x_120; +lean_object* x_110; lean_object* x_111; +x_110 = lean_box(0); +if (lean_is_scalar(x_81)) { + x_111 = lean_alloc_ctor(0, 1, 0); +} else { + x_111 = x_81; } -else -{ -lean_object* x_121; uint8_t x_122; -x_121 = lean_array_fget(x_116, x_76); -lean_dec(x_76); -lean_dec_ref(x_116); -x_122 = l_Lean_IR_instBEqIRType_beq(x_121, x_1); -if (x_122 == 0) -{ -lean_object* x_123; lean_object* x_124; -x_123 = l_Lean_IR_Checker_checkEqTypes___closed__0; -x_124 = l_Lean_IR_Checker_throwCheckerError___redArg(x_123, x_3, x_4, x_5, x_6); -return x_124; +lean_ctor_set(x_111, 0, x_110); +return x_111; } else { -lean_object* x_125; lean_object* x_126; -x_125 = lean_box(0); -x_126 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_126, 0, x_125); -return x_126; -} -} -} -case 11: -{ -lean_object* x_127; lean_object* x_128; uint8_t x_129; -x_127 = lean_ctor_get(x_113, 1); -lean_inc_ref(x_127); -lean_dec_ref(x_113); -x_128 = lean_array_get_size(x_127); -x_129 = lean_nat_dec_lt(x_76, x_128); -if (x_129 == 0) -{ -lean_object* x_130; lean_object* x_131; -lean_dec_ref(x_127); -lean_dec(x_76); -lean_dec(x_1); -x_130 = l_Lean_IR_Checker_checkExpr___closed__5; -x_131 = l_Lean_IR_Checker_throwCheckerError___redArg(x_130, x_3, x_4, x_5, x_6); -return x_131; +lean_object* x_112; lean_object* x_113; +lean_dec(x_81); +x_112 = l_Lean_IR_Checker_checkExpr___closed__5; +x_113 = l_Lean_IR_Checker_throwCheckerError___redArg(x_112, x_3, x_4, x_5, x_6); +return x_113; } -else -{ -lean_object* x_132; uint8_t x_133; -x_132 = lean_array_fget(x_127, x_76); -lean_dec(x_76); -lean_dec_ref(x_127); -x_133 = l_Lean_IR_instBEqIRType_beq(x_132, x_1); -if (x_133 == 0) -{ -lean_object* x_134; lean_object* x_135; -x_134 = l_Lean_IR_Checker_checkEqTypes___closed__0; -x_135 = l_Lean_IR_Checker_throwCheckerError___redArg(x_134, x_3, x_4, x_5, x_6); -return x_135; } else { -lean_object* x_136; lean_object* x_137; -x_136 = lean_box(0); -x_137 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_137, 0, x_136); -return x_137; +lean_object* x_114; lean_object* x_115; +lean_dec(x_106); +lean_dec(x_81); +lean_dec(x_77); +x_114 = l_Lean_IR_Checker_checkExpr___closed__8; +x_115 = l_Lean_IR_Checker_throwCheckerError___redArg(x_114, x_3, x_4, x_5, x_6); +return x_115; } } } case 12: { -lean_object* x_138; lean_object* x_139; +lean_object* x_116; lean_object* x_117; +lean_dec(x_77); lean_dec(x_76); lean_dec(x_1); -x_138 = lean_box(0); -x_139 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_139, 0, x_138); -return x_139; +x_116 = lean_box(0); +if (lean_is_scalar(x_81)) { + x_117 = lean_alloc_ctor(0, 1, 0); +} else { + x_117 = x_81; +} +lean_ctor_set(x_117, 0, x_116); +return x_117; } default: { -lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; +lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; +lean_dec(x_81); +lean_dec(x_77); lean_dec(x_76); lean_dec(x_1); -x_140 = l_Lean_IR_Checker_checkExpr___closed__6; -x_141 = l_Lean_IR_instToFormatIRType___private__1(x_113); -x_142 = l_Lean_IR_Checker_checkType___closed__1; -x_143 = lean_unsigned_to_nat(0u); -x_144 = l_Std_Format_pretty(x_141, x_142, x_143, x_143); -x_145 = lean_string_append(x_140, x_144); -lean_dec_ref(x_144); -x_146 = l_Lean_IR_Checker_checkVar___closed__2; -x_147 = lean_string_append(x_145, x_146); -x_148 = l_Lean_IR_Checker_throwCheckerError___redArg(x_147, x_3, x_4, x_5, x_6); -return x_148; -} +x_118 = l_Lean_IR_Checker_checkExpr___closed__9; +x_119 = l_Lean_IR_instToFormatIRType___private__1(x_80); +x_120 = l_Lean_IR_Checker_checkEqTypes___closed__1; +x_121 = lean_unsigned_to_nat(0u); +x_122 = l_Std_Format_pretty(x_119, x_120, x_121, x_121); +x_123 = lean_string_append(x_118, x_122); +lean_dec_ref(x_122); +x_124 = l_Lean_IR_Checker_checkVar___closed__2; +x_125 = lean_string_append(x_123, x_124); +x_126 = l_Lean_IR_Checker_throwCheckerError___redArg(x_125, x_3, x_4, x_5, x_6); +return x_126; } } } else { -uint8_t x_149; +uint8_t x_127; +lean_dec(x_77); lean_dec(x_76); lean_dec(x_1); -x_149 = !lean_is_exclusive(x_78); -if (x_149 == 0) +x_127 = !lean_is_exclusive(x_79); +if (x_127 == 0) { -return x_78; +return x_79; } else { -lean_object* x_150; lean_object* x_151; -x_150 = lean_ctor_get(x_78, 0); -lean_inc(x_150); -lean_dec(x_78); -x_151 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_151, 0, x_150); -return x_151; +lean_object* x_128; lean_object* x_129; +x_128 = lean_ctor_get(x_79, 0); +lean_inc(x_128); +lean_dec(x_79); +x_129 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_129, 0, x_128); +return x_129; } } } case 4: { -lean_object* x_152; lean_object* x_153; -x_152 = lean_ctor_get(x_2, 1); -lean_inc(x_152); +lean_object* x_130; lean_object* x_131; +x_130 = lean_ctor_get(x_2, 2); +lean_inc(x_130); lean_dec_ref(x_2); -x_153 = l_Lean_IR_Checker_checkObjVar(x_152, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_153) == 0) +x_131 = l_Lean_IR_Checker_checkObjOrStructVar(x_130, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_131) == 0) { -uint8_t x_154; -x_154 = !lean_is_exclusive(x_153); -if (x_154 == 0) +uint8_t x_132; +x_132 = !lean_is_exclusive(x_131); +if (x_132 == 0) { -lean_object* x_155; lean_object* x_156; uint8_t x_157; -x_155 = lean_ctor_get(x_153, 0); -lean_dec(x_155); -x_156 = lean_box(5); +lean_object* x_133; lean_object* x_134; uint8_t x_135; +x_133 = lean_ctor_get(x_131, 0); +lean_dec(x_133); +x_134 = lean_box(5); lean_inc(x_1); -x_157 = l_Lean_IR_instBEqIRType_beq(x_1, x_156); -if (x_157 == 0) -{ -lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; -lean_free_object(x_153); -x_158 = l_Lean_IR_Checker_checkType___closed__0; -x_159 = l_Lean_IR_instToFormatIRType___private__1(x_1); -x_160 = l_Lean_IR_Checker_checkType___closed__1; -x_161 = lean_unsigned_to_nat(0u); -x_162 = l_Std_Format_pretty(x_159, x_160, x_161, x_161); -x_163 = lean_string_append(x_158, x_162); -lean_dec_ref(x_162); -x_164 = l_Lean_IR_Checker_checkVar___closed__2; -x_165 = lean_string_append(x_163, x_164); -x_166 = l_Lean_IR_Checker_throwCheckerError___redArg(x_165, x_3, x_4, x_5, x_6); -return x_166; -} -else -{ -lean_object* x_167; +x_135 = l_Lean_IR_instBEqIRType_beq(x_1, x_134); +if (x_135 == 0) +{ +lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; +lean_free_object(x_131); +x_136 = l_Lean_IR_Checker_checkEqTypes___closed__0; +x_137 = l_Lean_IR_instToFormatIRType___private__1(x_1); +x_138 = l_Lean_IR_Checker_checkEqTypes___closed__1; +x_139 = lean_unsigned_to_nat(0u); +x_140 = l_Std_Format_pretty(x_137, x_138, x_139, x_139); +x_141 = lean_string_append(x_136, x_140); +lean_dec_ref(x_140); +x_142 = l_Lean_IR_Checker_checkVar___closed__2; +x_143 = lean_string_append(x_141, x_142); +x_144 = l_Lean_IR_Checker_throwCheckerError___redArg(x_143, x_3, x_4, x_5, x_6); +return x_144; +} +else +{ +lean_object* x_145; lean_dec(x_1); -x_167 = lean_box(0); -lean_ctor_set(x_153, 0, x_167); -return x_153; +x_145 = lean_box(0); +lean_ctor_set(x_131, 0, x_145); +return x_131; } } else { -lean_object* x_168; uint8_t x_169; -lean_dec(x_153); -x_168 = lean_box(5); +lean_object* x_146; uint8_t x_147; +lean_dec(x_131); +x_146 = lean_box(5); lean_inc(x_1); -x_169 = l_Lean_IR_instBEqIRType_beq(x_1, x_168); -if (x_169 == 0) -{ -lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; -x_170 = l_Lean_IR_Checker_checkType___closed__0; -x_171 = l_Lean_IR_instToFormatIRType___private__1(x_1); -x_172 = l_Lean_IR_Checker_checkType___closed__1; -x_173 = lean_unsigned_to_nat(0u); -x_174 = l_Std_Format_pretty(x_171, x_172, x_173, x_173); -x_175 = lean_string_append(x_170, x_174); -lean_dec_ref(x_174); -x_176 = l_Lean_IR_Checker_checkVar___closed__2; -x_177 = lean_string_append(x_175, x_176); -x_178 = l_Lean_IR_Checker_throwCheckerError___redArg(x_177, x_3, x_4, x_5, x_6); -return x_178; +x_147 = l_Lean_IR_instBEqIRType_beq(x_1, x_146); +if (x_147 == 0) +{ +lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; +x_148 = l_Lean_IR_Checker_checkEqTypes___closed__0; +x_149 = l_Lean_IR_instToFormatIRType___private__1(x_1); +x_150 = l_Lean_IR_Checker_checkEqTypes___closed__1; +x_151 = lean_unsigned_to_nat(0u); +x_152 = l_Std_Format_pretty(x_149, x_150, x_151, x_151); +x_153 = lean_string_append(x_148, x_152); +lean_dec_ref(x_152); +x_154 = l_Lean_IR_Checker_checkVar___closed__2; +x_155 = lean_string_append(x_153, x_154); +x_156 = l_Lean_IR_Checker_throwCheckerError___redArg(x_155, x_3, x_4, x_5, x_6); +return x_156; } else { -lean_object* x_179; lean_object* x_180; +lean_object* x_157; lean_object* x_158; lean_dec(x_1); -x_179 = lean_box(0); -x_180 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_180, 0, x_179); -return x_180; +x_157 = lean_box(0); +x_158 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_158, 0, x_157); +return x_158; } } } else { lean_dec(x_1); -return x_153; +return x_131; } } case 5: { -lean_object* x_181; lean_object* x_182; -x_181 = lean_ctor_get(x_2, 2); -lean_inc(x_181); +lean_object* x_159; lean_object* x_160; +x_159 = lean_ctor_get(x_2, 3); +lean_inc(x_159); lean_dec_ref(x_2); -x_182 = l_Lean_IR_Checker_checkObjVar(x_181, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_182) == 0) +x_160 = l_Lean_IR_Checker_checkObjOrStructVar(x_159, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_160) == 0) { -lean_object* x_183; -lean_dec_ref(x_182); -x_183 = l_Lean_IR_Checker_checkScalarType(x_1, x_3, x_4, x_5, x_6); -return x_183; +lean_object* x_161; +lean_dec_ref(x_160); +x_161 = l_Lean_IR_Checker_checkScalarType(x_1, x_3, x_4, x_5, x_6); +return x_161; } else { lean_dec(x_1); -return x_182; +return x_160; } } case 6: { -lean_object* x_184; lean_object* x_185; lean_object* x_186; +lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_dec(x_1); -x_184 = lean_ctor_get(x_2, 0); -lean_inc(x_184); -x_185 = lean_ctor_get(x_2, 1); -lean_inc_ref(x_185); +x_162 = lean_ctor_get(x_2, 0); +lean_inc(x_162); +x_163 = lean_ctor_get(x_2, 1); +lean_inc_ref(x_163); lean_dec_ref(x_2); -x_186 = l_Lean_IR_Checker_checkFullApp(x_184, x_185, x_3, x_4, x_5, x_6); -lean_dec_ref(x_185); -return x_186; +x_164 = l_Lean_IR_Checker_checkFullApp(x_162, x_163, x_3, x_4, x_5, x_6); +lean_dec_ref(x_163); +return x_164; } case 7: { -lean_object* x_187; lean_object* x_188; lean_object* x_189; -x_187 = lean_ctor_get(x_2, 0); -lean_inc(x_187); -x_188 = lean_ctor_get(x_2, 1); -lean_inc_ref(x_188); +lean_object* x_165; lean_object* x_166; lean_object* x_167; +x_165 = lean_ctor_get(x_2, 0); +lean_inc(x_165); +x_166 = lean_ctor_get(x_2, 1); +lean_inc_ref(x_166); lean_dec_ref(x_2); -x_189 = l_Lean_IR_Checker_checkPartialApp(x_187, x_188, x_3, x_4, x_5, x_6); -lean_dec_ref(x_188); -if (lean_obj_tag(x_189) == 0) +x_167 = l_Lean_IR_Checker_checkPartialApp(x_165, x_166, x_3, x_4, x_5, x_6); +lean_dec_ref(x_166); +if (lean_obj_tag(x_167) == 0) { -lean_object* x_190; -lean_dec_ref(x_189); -x_190 = l_Lean_IR_Checker_checkObjType(x_1, x_3, x_4, x_5, x_6); -return x_190; +lean_object* x_168; +lean_dec_ref(x_167); +x_168 = l_Lean_IR_Checker_checkObjType(x_1, x_3, x_4, x_5, x_6); +return x_168; } else { lean_dec(x_1); -return x_189; +return x_167; } } case 8: { -lean_object* x_191; lean_object* x_192; lean_object* x_193; -x_191 = lean_ctor_get(x_2, 0); -lean_inc(x_191); -x_192 = lean_ctor_get(x_2, 1); -lean_inc_ref(x_192); +lean_object* x_169; lean_object* x_170; lean_object* x_171; +x_169 = lean_ctor_get(x_2, 0); +lean_inc(x_169); +x_170 = lean_ctor_get(x_2, 1); +lean_inc_ref(x_170); lean_dec_ref(x_2); -x_193 = l_Lean_IR_Checker_checkObjVar(x_191, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_193) == 0) +x_171 = l_Lean_IR_Checker_checkObjVar(x_169, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_171) == 0) { -lean_object* x_194; -lean_dec_ref(x_193); -x_194 = l_Lean_IR_Checker_checkArgs(x_192, x_3, x_4, x_5, x_6); -lean_dec_ref(x_192); -if (lean_obj_tag(x_194) == 0) +lean_object* x_172; +lean_dec_ref(x_171); +x_172 = l_Lean_IR_Checker_checkArgs(x_170, x_3, x_4, x_5, x_6); +lean_dec_ref(x_170); +if (lean_obj_tag(x_172) == 0) { -lean_object* x_195; -lean_dec_ref(x_194); -x_195 = l_Lean_IR_Checker_checkObjType(x_1, x_3, x_4, x_5, x_6); -return x_195; +lean_object* x_173; +lean_dec_ref(x_172); +x_173 = l_Lean_IR_Checker_checkObjType(x_1, x_3, x_4, x_5, x_6); +return x_173; } else { lean_dec(x_1); -return x_194; +return x_172; } } else { -lean_dec_ref(x_192); +lean_dec_ref(x_170); lean_dec(x_1); -return x_193; +return x_171; } } case 9: { -lean_object* x_196; lean_object* x_197; lean_object* x_198; -x_196 = lean_ctor_get(x_2, 0); -lean_inc(x_196); -x_197 = lean_ctor_get(x_2, 1); -lean_inc(x_197); +lean_object* x_174; lean_object* x_175; lean_object* x_176; +x_174 = lean_ctor_get(x_2, 0); +lean_inc(x_174); +x_175 = lean_ctor_get(x_2, 1); +lean_inc(x_175); lean_dec_ref(x_2); -x_198 = l_Lean_IR_Checker_checkObjType(x_1, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_198) == 0) -{ -lean_object* x_199; -lean_dec_ref(x_198); -lean_inc(x_197); -x_199 = l_Lean_IR_Checker_checkScalarVar(x_197, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_199) == 0) -{ -lean_object* x_200; -lean_dec_ref(x_199); -x_200 = l_Lean_IR_Checker_getType(x_197, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_200) == 0) -{ -uint8_t x_201; -x_201 = !lean_is_exclusive(x_200); -if (x_201 == 0) -{ -lean_object* x_202; uint8_t x_203; -x_202 = lean_ctor_get(x_200, 0); -lean_inc(x_202); -x_203 = l_Lean_IR_instBEqIRType_beq(x_202, x_196); -if (x_203 == 0) -{ -lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; -lean_free_object(x_200); -x_204 = l_Lean_IR_Checker_checkType___closed__0; -x_205 = l_Lean_IR_instToFormatIRType___private__1(x_202); -x_206 = l_Lean_IR_Checker_checkType___closed__1; -x_207 = lean_unsigned_to_nat(0u); -x_208 = l_Std_Format_pretty(x_205, x_206, x_207, x_207); -x_209 = lean_string_append(x_204, x_208); -lean_dec_ref(x_208); -x_210 = l_Lean_IR_Checker_checkVar___closed__2; -x_211 = lean_string_append(x_209, x_210); -x_212 = l_Lean_IR_Checker_throwCheckerError___redArg(x_211, x_3, x_4, x_5, x_6); -return x_212; +x_176 = l_Lean_IR_Checker_checkObjOrStructType(x_1, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_176) == 0) +{ +lean_object* x_177; +lean_dec_ref(x_176); +lean_inc(x_175); +x_177 = l_Lean_IR_Checker_checkScalarOrStructVar(x_175, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_177) == 0) +{ +lean_object* x_178; +lean_dec_ref(x_177); +x_178 = l_Lean_IR_Checker_getType(x_175, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_178) == 0) +{ +uint8_t x_179; +x_179 = !lean_is_exclusive(x_178); +if (x_179 == 0) +{ +lean_object* x_180; uint8_t x_181; +x_180 = lean_ctor_get(x_178, 0); +lean_inc(x_180); +x_181 = l_Lean_IR_instBEqIRType_beq(x_180, x_174); +if (x_181 == 0) +{ +lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; +lean_free_object(x_178); +x_182 = l_Lean_IR_Checker_checkEqTypes___closed__0; +x_183 = l_Lean_IR_instToFormatIRType___private__1(x_180); +x_184 = l_Lean_IR_Checker_checkEqTypes___closed__1; +x_185 = lean_unsigned_to_nat(0u); +x_186 = l_Std_Format_pretty(x_183, x_184, x_185, x_185); +x_187 = lean_string_append(x_182, x_186); +lean_dec_ref(x_186); +x_188 = l_Lean_IR_Checker_checkVar___closed__2; +x_189 = lean_string_append(x_187, x_188); +x_190 = l_Lean_IR_Checker_throwCheckerError___redArg(x_189, x_3, x_4, x_5, x_6); +return x_190; } else { -lean_object* x_213; -lean_dec(x_202); -x_213 = lean_box(0); -lean_ctor_set(x_200, 0, x_213); -return x_200; +lean_object* x_191; +lean_dec(x_180); +x_191 = lean_box(0); +lean_ctor_set(x_178, 0, x_191); +return x_178; } } else { -lean_object* x_214; uint8_t x_215; -x_214 = lean_ctor_get(x_200, 0); -lean_inc(x_214); -lean_dec(x_200); -lean_inc(x_214); -x_215 = l_Lean_IR_instBEqIRType_beq(x_214, x_196); -if (x_215 == 0) +lean_object* x_192; uint8_t x_193; +x_192 = lean_ctor_get(x_178, 0); +lean_inc(x_192); +lean_dec(x_178); +lean_inc(x_192); +x_193 = l_Lean_IR_instBEqIRType_beq(x_192, x_174); +if (x_193 == 0) { -lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; -x_216 = l_Lean_IR_Checker_checkType___closed__0; -x_217 = l_Lean_IR_instToFormatIRType___private__1(x_214); -x_218 = l_Lean_IR_Checker_checkType___closed__1; -x_219 = lean_unsigned_to_nat(0u); -x_220 = l_Std_Format_pretty(x_217, x_218, x_219, x_219); -x_221 = lean_string_append(x_216, x_220); -lean_dec_ref(x_220); -x_222 = l_Lean_IR_Checker_checkVar___closed__2; -x_223 = lean_string_append(x_221, x_222); -x_224 = l_Lean_IR_Checker_throwCheckerError___redArg(x_223, x_3, x_4, x_5, x_6); -return x_224; +lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; +x_194 = l_Lean_IR_Checker_checkEqTypes___closed__0; +x_195 = l_Lean_IR_instToFormatIRType___private__1(x_192); +x_196 = l_Lean_IR_Checker_checkEqTypes___closed__1; +x_197 = lean_unsigned_to_nat(0u); +x_198 = l_Std_Format_pretty(x_195, x_196, x_197, x_197); +x_199 = lean_string_append(x_194, x_198); +lean_dec_ref(x_198); +x_200 = l_Lean_IR_Checker_checkVar___closed__2; +x_201 = lean_string_append(x_199, x_200); +x_202 = l_Lean_IR_Checker_throwCheckerError___redArg(x_201, x_3, x_4, x_5, x_6); +return x_202; } else { -lean_object* x_225; lean_object* x_226; -lean_dec(x_214); -x_225 = lean_box(0); -x_226 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_226, 0, x_225); -return x_226; +lean_object* x_203; lean_object* x_204; +lean_dec(x_192); +x_203 = lean_box(0); +x_204 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_204, 0, x_203); +return x_204; } } } else { -uint8_t x_227; -lean_dec(x_196); -x_227 = !lean_is_exclusive(x_200); -if (x_227 == 0) +uint8_t x_205; +lean_dec(x_174); +x_205 = !lean_is_exclusive(x_178); +if (x_205 == 0) { -return x_200; +return x_178; } else { -lean_object* x_228; lean_object* x_229; -x_228 = lean_ctor_get(x_200, 0); -lean_inc(x_228); -lean_dec(x_200); -x_229 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_229, 0, x_228); -return x_229; +lean_object* x_206; lean_object* x_207; +x_206 = lean_ctor_get(x_178, 0); +lean_inc(x_206); +lean_dec(x_178); +x_207 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_207, 0, x_206); +return x_207; } } } else { -lean_dec(x_197); -lean_dec(x_196); -return x_199; +lean_dec(x_175); +lean_dec(x_174); +return x_177; } } else { -lean_dec(x_197); -lean_dec(x_196); -return x_198; +lean_dec(x_175); +lean_dec(x_174); +return x_176; } } case 10: { -lean_object* x_230; lean_object* x_231; -x_230 = lean_ctor_get(x_2, 0); -lean_inc(x_230); +lean_object* x_208; lean_object* x_209; +x_208 = lean_ctor_get(x_2, 0); +lean_inc(x_208); lean_dec_ref(x_2); -x_231 = l_Lean_IR_Checker_checkScalarType(x_1, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_231) == 0) +x_209 = l_Lean_IR_Checker_checkScalarOrStructType(x_1, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_209) == 0) { -lean_object* x_232; -lean_dec_ref(x_231); -x_232 = l_Lean_IR_Checker_checkObjVar(x_230, x_3, x_4, x_5, x_6); -return x_232; +lean_object* x_210; +lean_dec_ref(x_209); +x_210 = l_Lean_IR_Checker_checkObjVar(x_208, x_3, x_4, x_5, x_6); +return x_210; } else { -lean_dec(x_230); -return x_231; +lean_dec(x_208); +return x_209; } } case 11: { -uint8_t x_233; -x_233 = !lean_is_exclusive(x_2); -if (x_233 == 0) +uint8_t x_211; +x_211 = !lean_is_exclusive(x_2); +if (x_211 == 0) { -lean_object* x_234; -x_234 = lean_ctor_get(x_2, 0); -if (lean_obj_tag(x_234) == 1) +lean_object* x_212; +x_212 = lean_ctor_get(x_2, 0); +if (lean_obj_tag(x_212) == 1) { -lean_object* x_235; +lean_object* x_213; lean_free_object(x_2); -lean_dec_ref(x_234); -x_235 = l_Lean_IR_Checker_checkObjType(x_1, x_3, x_4, x_5, x_6); -return x_235; +lean_dec_ref(x_212); +x_213 = l_Lean_IR_Checker_checkObjType(x_1, x_3, x_4, x_5, x_6); +return x_213; } else { -lean_object* x_236; -lean_dec_ref(x_234); +lean_object* x_214; +lean_dec_ref(x_212); lean_dec(x_1); -x_236 = lean_box(0); +x_214 = lean_box(0); lean_ctor_set_tag(x_2, 0); -lean_ctor_set(x_2, 0, x_236); +lean_ctor_set(x_2, 0, x_214); return x_2; } } else { -lean_object* x_237; -x_237 = lean_ctor_get(x_2, 0); -lean_inc(x_237); +lean_object* x_215; +x_215 = lean_ctor_get(x_2, 0); +lean_inc(x_215); lean_dec(x_2); -if (lean_obj_tag(x_237) == 1) +if (lean_obj_tag(x_215) == 1) { -lean_object* x_238; -lean_dec_ref(x_237); -x_238 = l_Lean_IR_Checker_checkObjType(x_1, x_3, x_4, x_5, x_6); -return x_238; +lean_object* x_216; +lean_dec_ref(x_215); +x_216 = l_Lean_IR_Checker_checkObjType(x_1, x_3, x_4, x_5, x_6); +return x_216; } else { -lean_object* x_239; lean_object* x_240; -lean_dec_ref(x_237); +lean_object* x_217; lean_object* x_218; +lean_dec_ref(x_215); lean_dec(x_1); -x_239 = lean_box(0); -x_240 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_240, 0, x_239); -return x_240; +x_217 = lean_box(0); +x_218 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_218, 0, x_217); +return x_218; } } } default: { -lean_object* x_241; lean_object* x_242; -x_241 = lean_ctor_get(x_2, 0); -lean_inc(x_241); +lean_object* x_219; lean_object* x_220; +x_219 = lean_ctor_get(x_2, 0); +lean_inc(x_219); lean_dec_ref(x_2); -x_242 = l_Lean_IR_Checker_checkObjVar(x_241, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_242) == 0) +x_220 = l_Lean_IR_Checker_checkObjVar(x_219, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_220) == 0) { -uint8_t x_243; -x_243 = !lean_is_exclusive(x_242); -if (x_243 == 0) +uint8_t x_221; +x_221 = !lean_is_exclusive(x_220); +if (x_221 == 0) { -lean_object* x_244; lean_object* x_245; uint8_t x_246; -x_244 = lean_ctor_get(x_242, 0); -lean_dec(x_244); -x_245 = lean_box(1); +lean_object* x_222; lean_object* x_223; uint8_t x_224; +x_222 = lean_ctor_get(x_220, 0); +lean_dec(x_222); +x_223 = lean_box(1); lean_inc(x_1); -x_246 = l_Lean_IR_instBEqIRType_beq(x_1, x_245); -if (x_246 == 0) -{ -lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; -lean_free_object(x_242); -x_247 = l_Lean_IR_Checker_checkType___closed__0; -x_248 = l_Lean_IR_instToFormatIRType___private__1(x_1); -x_249 = l_Lean_IR_Checker_checkType___closed__1; -x_250 = lean_unsigned_to_nat(0u); -x_251 = l_Std_Format_pretty(x_248, x_249, x_250, x_250); -x_252 = lean_string_append(x_247, x_251); -lean_dec_ref(x_251); -x_253 = l_Lean_IR_Checker_checkVar___closed__2; -x_254 = lean_string_append(x_252, x_253); -x_255 = l_Lean_IR_Checker_throwCheckerError___redArg(x_254, x_3, x_4, x_5, x_6); -return x_255; -} -else -{ -lean_object* x_256; +x_224 = l_Lean_IR_instBEqIRType_beq(x_1, x_223); +if (x_224 == 0) +{ +lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; +lean_free_object(x_220); +x_225 = l_Lean_IR_Checker_checkEqTypes___closed__0; +x_226 = l_Lean_IR_instToFormatIRType___private__1(x_1); +x_227 = l_Lean_IR_Checker_checkEqTypes___closed__1; +x_228 = lean_unsigned_to_nat(0u); +x_229 = l_Std_Format_pretty(x_226, x_227, x_228, x_228); +x_230 = lean_string_append(x_225, x_229); +lean_dec_ref(x_229); +x_231 = l_Lean_IR_Checker_checkVar___closed__2; +x_232 = lean_string_append(x_230, x_231); +x_233 = l_Lean_IR_Checker_throwCheckerError___redArg(x_232, x_3, x_4, x_5, x_6); +return x_233; +} +else +{ +lean_object* x_234; lean_dec(x_1); -x_256 = lean_box(0); -lean_ctor_set(x_242, 0, x_256); -return x_242; +x_234 = lean_box(0); +lean_ctor_set(x_220, 0, x_234); +return x_220; } } else { -lean_object* x_257; uint8_t x_258; -lean_dec(x_242); -x_257 = lean_box(1); +lean_object* x_235; uint8_t x_236; +lean_dec(x_220); +x_235 = lean_box(1); lean_inc(x_1); -x_258 = l_Lean_IR_instBEqIRType_beq(x_1, x_257); -if (x_258 == 0) +x_236 = l_Lean_IR_instBEqIRType_beq(x_1, x_235); +if (x_236 == 0) { -lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; -x_259 = l_Lean_IR_Checker_checkType___closed__0; -x_260 = l_Lean_IR_instToFormatIRType___private__1(x_1); -x_261 = l_Lean_IR_Checker_checkType___closed__1; -x_262 = lean_unsigned_to_nat(0u); -x_263 = l_Std_Format_pretty(x_260, x_261, x_262, x_262); -x_264 = lean_string_append(x_259, x_263); -lean_dec_ref(x_263); -x_265 = l_Lean_IR_Checker_checkVar___closed__2; -x_266 = lean_string_append(x_264, x_265); -x_267 = l_Lean_IR_Checker_throwCheckerError___redArg(x_266, x_3, x_4, x_5, x_6); -return x_267; +lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; +x_237 = l_Lean_IR_Checker_checkEqTypes___closed__0; +x_238 = l_Lean_IR_instToFormatIRType___private__1(x_1); +x_239 = l_Lean_IR_Checker_checkEqTypes___closed__1; +x_240 = lean_unsigned_to_nat(0u); +x_241 = l_Std_Format_pretty(x_238, x_239, x_240, x_240); +x_242 = lean_string_append(x_237, x_241); +lean_dec_ref(x_241); +x_243 = l_Lean_IR_Checker_checkVar___closed__2; +x_244 = lean_string_append(x_242, x_243); +x_245 = l_Lean_IR_Checker_throwCheckerError___redArg(x_244, x_3, x_4, x_5, x_6); +return x_245; } else { -lean_object* x_268; lean_object* x_269; +lean_object* x_246; lean_object* x_247; lean_dec(x_1); -x_268 = lean_box(0); -x_269 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_269, 0, x_268); -return x_269; +x_246 = lean_box(0); +x_247 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_247, 0, x_246); +return x_247; } } } else { lean_dec(x_1); -return x_242; +return x_220; } } } @@ -5000,9 +5191,9 @@ case 4: lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; x_75 = lean_ctor_get(x_1, 0); lean_inc(x_75); -x_76 = lean_ctor_get(x_1, 2); +x_76 = lean_ctor_get(x_1, 3); lean_inc(x_76); -x_77 = lean_ctor_get(x_1, 3); +x_77 = lean_ctor_get(x_1, 4); lean_inc(x_77); lean_dec_ref(x_1); x_78 = l_Lean_IR_Checker_checkVar(x_75, x_2, x_3, x_4, x_5); @@ -5037,9 +5228,9 @@ case 5: lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; x_81 = lean_ctor_get(x_1, 0); lean_inc(x_81); -x_82 = lean_ctor_get(x_1, 3); +x_82 = lean_ctor_get(x_1, 4); lean_inc(x_82); -x_83 = lean_ctor_get(x_1, 5); +x_83 = lean_ctor_get(x_1, 6); lean_inc(x_83); lean_dec_ref(x_1); x_84 = l_Lean_IR_Checker_checkVar(x_81, x_2, x_3, x_4, x_5); @@ -5857,14 +6048,16 @@ l_Lean_IR_Checker_checkJP___closed__1 = _init_l_Lean_IR_Checker_checkJP___closed lean_mark_persistent(l_Lean_IR_Checker_checkJP___closed__1); l_Lean_IR_Checker_checkEqTypes___closed__0 = _init_l_Lean_IR_Checker_checkEqTypes___closed__0(); lean_mark_persistent(l_Lean_IR_Checker_checkEqTypes___closed__0); +l_Lean_IR_Checker_checkEqTypes___closed__1 = _init_l_Lean_IR_Checker_checkEqTypes___closed__1(); +lean_mark_persistent(l_Lean_IR_Checker_checkEqTypes___closed__1); +l_Lean_IR_Checker_checkEqTypes___closed__2 = _init_l_Lean_IR_Checker_checkEqTypes___closed__2(); +lean_mark_persistent(l_Lean_IR_Checker_checkEqTypes___closed__2); l_Lean_IR_Checker_checkType___closed__0 = _init_l_Lean_IR_Checker_checkType___closed__0(); lean_mark_persistent(l_Lean_IR_Checker_checkType___closed__0); -l_Lean_IR_Checker_checkType___closed__1 = _init_l_Lean_IR_Checker_checkType___closed__1(); -lean_mark_persistent(l_Lean_IR_Checker_checkType___closed__1); -l_Lean_IR_Checker_checkType___closed__2 = _init_l_Lean_IR_Checker_checkType___closed__2(); -lean_mark_persistent(l_Lean_IR_Checker_checkType___closed__2); l_Lean_IR_Checker_checkObjType___closed__0 = _init_l_Lean_IR_Checker_checkObjType___closed__0(); lean_mark_persistent(l_Lean_IR_Checker_checkObjType___closed__0); +l_Lean_IR_Checker_checkObjOrStructType___closed__0 = _init_l_Lean_IR_Checker_checkObjOrStructType___closed__0(); +lean_mark_persistent(l_Lean_IR_Checker_checkObjOrStructType___closed__0); l_Lean_IR_Checker_checkScalarType___closed__0 = _init_l_Lean_IR_Checker_checkScalarType___closed__0(); lean_mark_persistent(l_Lean_IR_Checker_checkScalarType___closed__0); l_Lean_IR_Checker_checkFullApp___closed__0 = _init_l_Lean_IR_Checker_checkFullApp___closed__0(); @@ -5895,6 +6088,12 @@ l_Lean_IR_Checker_checkExpr___closed__5 = _init_l_Lean_IR_Checker_checkExpr___cl lean_mark_persistent(l_Lean_IR_Checker_checkExpr___closed__5); l_Lean_IR_Checker_checkExpr___closed__6 = _init_l_Lean_IR_Checker_checkExpr___closed__6(); lean_mark_persistent(l_Lean_IR_Checker_checkExpr___closed__6); +l_Lean_IR_Checker_checkExpr___closed__7 = _init_l_Lean_IR_Checker_checkExpr___closed__7(); +lean_mark_persistent(l_Lean_IR_Checker_checkExpr___closed__7); +l_Lean_IR_Checker_checkExpr___closed__8 = _init_l_Lean_IR_Checker_checkExpr___closed__8(); +lean_mark_persistent(l_Lean_IR_Checker_checkExpr___closed__8); +l_Lean_IR_Checker_checkExpr___closed__9 = _init_l_Lean_IR_Checker_checkExpr___closed__9(); +lean_mark_persistent(l_Lean_IR_Checker_checkExpr___closed__9); l_Lean_IR_Checker_withParams___closed__0 = _init_l_Lean_IR_Checker_withParams___closed__0(); lean_mark_persistent(l_Lean_IR_Checker_withParams___closed__0); l_Lean_IR_Checker_withParams___closed__1 = _init_l_Lean_IR_Checker_withParams___closed__1(); diff --git a/stage0/stdlib/Lean/Compiler/IR/EmitC.c b/stage0/stdlib/Lean/Compiler/IR/EmitC.c index 2e429d799b1d..2495f2fd220a 100644 --- a/stage0/stdlib/Lean/Compiler/IR/EmitC.c +++ b/stage0/stdlib/Lean/Compiler/IR/EmitC.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Compiler.IR.EmitC -// Imports: public import Lean.Compiler.NameMangling public import Lean.Compiler.IR.EmitUtil public import Lean.Compiler.IR.NormIds public import Lean.Compiler.IR.SimpCase public import Lean.Compiler.IR.Boxing public import Lean.Compiler.ModPkgExt +// Imports: public import Lean.Compiler.NameMangling public import Lean.Compiler.IR.EmitUtil public import Lean.Compiler.IR.NormIds public import Lean.Compiler.IR.SimpCase public import Lean.Compiler.IR.Boxing public import Lean.Compiler.IR.StructRC public import Lean.Compiler.ModPkgExt #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -13,35 +13,52 @@ #ifdef __cplusplus extern "C" { #endif +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitFullAppArg_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDecOfType_spec__0___redArg___closed__0; extern lean_object* l_Lean_IR_instInhabitedArg_default; +uint8_t l_Lean_IR_IRType_isScalarOrStruct(lean_object*); +static lean_object* l_Lean_IR_EmitC_emitUSet___closed__4; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitTailCall___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitInc___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLns___redArg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_hasMainFn(lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitStructIncDecFn___closed__3; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructIncDecFn_spec__0_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitFileHeader___closed__1; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__1; static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__25; LEAN_EXPORT lean_object* l_List_forM___at___00Lean_IR_EmitC_emitInitFn_spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitInc___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitFileHeader___closed__0; +LEAN_EXPORT lean_object* l_panic___at___00Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0_spec__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitStructBox___closed__1; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFnDecls(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLit___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___at___00Lean_IR_EmitC_emitSimpleExternalCall_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLns___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitDecl(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_IR_EmitC_emitInc___redArg___closed__0; lean_object* lean_mk_empty_array_with_capacity(lean_object*); +static lean_object* l_Lean_IR_EmitC_emitSSet___closed__9; lean_object* l_List_forM___redArg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg___closed__0; lean_object* l_EStateM_instMonad___lam__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_EmitC_emitUSet_spec__0(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitSSet___closed__10; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructIncDecFn_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__7; static lean_object* l_Lean_IR_EmitC_emitBlock___closed__0; lean_object* lean_string_utf8_next_fast(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitArg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUnionSwitch(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitUnionSwitch_spec__0___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitPartialApp_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructBoxFn___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitCtor_spec__0___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_uint32_to_nat(uint32_t); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitDec___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_IR_IRType_isVoid(lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitBoxFn___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitBoxFn___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLhs(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__22; lean_object* l_EStateM_instMonad___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -49,134 +66,178 @@ static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__34; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__24; static lean_object* l_Lean_IR_EmitC_emitInitFn___closed__3; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__45; static lean_object* l_Lean_IR_EmitC_emitReuse___closed__1; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLit___redArg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentEnvExtension_getState___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUnboxFn___boxed(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_IR_instBEqJoinPointId_beq(lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_type(lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFullApp_spec__0(lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_structBoxFnPrefix; LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00__private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0___redArg(lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__6; lean_object* l_Lean_IR_findEnvDecl(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitCase_spec__4(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_toStringArgs(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreadArgs_spec__0(uint8_t, lean_object*, size_t, size_t, uint8_t, lean_object*, lean_object*); +lean_object* l_Lean_IR_collectStructTypes(lean_object*); static lean_object* l_Lean_IR_EmitC_emitDeclInit___closed__3; LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Array_isEmpty___redArg(lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitArg___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_IR_EmitC_emitSSet___redArg___closed__5; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__1___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUProj(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_declareParams_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__1___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUProj(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitStructBox___closed__5; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructIncDecFn_spec__0(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitFileHeader___closed__4; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitCtor_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_IR_instBEqVarId_beq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_isIf(lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitUnionSwitch_spec__0___redArg___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_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at___00Lean_IR_EmitC_emitFnDecls_spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitUnionSwitchWithImpossible_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_emitC___closed__0; static lean_object* l_Lean_IR_EmitC_emitInitFn___closed__4; static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__17; -static lean_object* l_Lean_IR_EmitC_declareVar___redArg___closed__0; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDecls_spec__1___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__10; +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructReshapeFn___lam__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_IR_EmitC_emitNumLit(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLns(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitStructIncDecFn___lam__0___closed__0; +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructBoxFn___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDecOfType_spec__0___redArg(lean_object*, uint8_t, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructIncDecFn___lam__0(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_argToCString___closed__1; static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__47; static lean_object* l_Lean_IR_EmitC_emitFileHeader___closed__16; static lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__0___redArg___closed__0; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_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_IR_EmitC_emitLns___redArg___closed__7; static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__23; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFullAppArgs___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitMainFn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitReset(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSProj(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSProj(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitCtorSetArgs_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_structUnboxFnPrefix; static lean_object* l_Lean_IR_EmitC_emitFnDeclAux___closed__4; static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__19; static lean_object* l_Lean_IR_EmitC_emitFileHeader___closed__21; LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitAllocCtor___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitReuse(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__4; +static lean_object* l_Lean_IR_EmitC_emitStructIncDecFn___closed__1; uint8_t lean_usize_dec_eq(size_t, size_t); +lean_object* l_Lean_IR_instBEqIRTypeApprox___lam__0___boxed(lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitLns___redArg___closed__8; static lean_object* l_Lean_IR_EmitC_emitMarkPersistent___closed__0; +static lean_object* l_Lean_IR_EmitC_emitStructDefn___closed__8; LEAN_EXPORT lean_object* l_List_forM___at___00Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0_spec__0(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSpreadArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_toCName___closed__0; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__0___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_IR_EmitC_emitJmp___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitStructBoxFn___closed__0; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__1___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitSSet___closed__8; +LEAN_EXPORT lean_object* l_Array_isEqvAux___at___00Lean_IR_EmitC_compatibleReshape_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__0; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_EmitC_needsRC_spec__0___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__1(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at___00Lean_IR_EmitC_emitFnDecls_spec__0(lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitCtor___closed__2; lean_object* l_Lean_getExternEntryFor(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFileFooter___redArg(lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__44; uint8_t l_Lean_IR_instBEqIRType_beq(lean_object*, lean_object*); -static lean_object* l_Lean_IR_EmitC_emitSSet___redArg___closed__1; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDecls_spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitPartialApp_spec__0___redArg___closed__0; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitCtor_spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_declareParams(lean_object*, lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__3; static lean_object* l_Lean_IR_EmitC_throwInvalidExportName___redArg___closed__0; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0___redArg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitFnDeclAux___closed__5; +static lean_object* l_Lean_IR_EmitC_emitSpreadArg___closed__0; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitStructDecls_spec__0___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_IR_EmitC_emitExternDeclAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__2___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitAllocCtor___redArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitLns___redArg___closed__0; static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__4; uint8_t l_Lean_IR_IRType_isErased(lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at___00Lean_IR_EmitC_emitFnDecls_spec__0___boxed(lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitReset___closed__3; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_declareParams_spec__0(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitSSet___closed__4; uint8_t l_Lean_IR_FnBody_isTerminal(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSetTag___redArg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_structBoxFnPrefix___closed__0; +static lean_object* l_Lean_IR_EmitC_emitSpreadArg___closed__1; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFnDeclAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_IR_EmitC_emitApp___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_getJPParams(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_IR_EmitC_emitDec___redArg___closed__0; +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSpreadValue___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitStructReshapeFn___closed__1; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFnDeclAux(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitReset_spec__0___redArg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_get_export_name_for(lean_object*, lean_object*); uint8_t l_Lean_isClosedTermName(lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__12; -static lean_object* l_Lean_IR_EmitC_toCType___closed__12; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitFullAppArgs_spec__0___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUnbox(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitBox___redArg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget_borrowed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitDeclInit___closed__1; -static lean_object* l_Lean_IR_EmitC_emitSProj___redArg___closed__5; static lean_object* l_List_forM___at___00Lean_IR_EmitC_emitFnDecls_spec__3___closed__0; -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitProj(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitUnionSwitchWithImpossible_spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitProj(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_get_symbol_stem(lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_structUnboxFnPrefix___closed__0; static lean_object* l_Lean_IR_EmitC_emitFileHeader___closed__15; -static lean_object* l_Lean_IR_EmitC_emitSProj___redArg___closed__4; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructUnboxFn_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitDecl___closed__0; LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitCtorSetArgs_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__2(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_IR_EmitC_argToCString___closed__0; -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_declareVar___redArg(lean_object*, lean_object*, lean_object*); lean_object* l_EStateM_bind(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSSet(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Array_isEqvAux___at___00Lean_IR_EmitC_compatibleReshape_spec__0___redArg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSSet(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitStructDefn___closed__3; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFnDecl(lean_object*, uint8_t, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_toCType___closed__4; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCName(lean_object*, lean_object*, lean_object*); +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__3; static lean_object* l_Lean_IR_EmitC_emitInitFn___closed__5; static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__20; -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_IR_EmitC_emitBoxFn___redArg___closed__1; -static lean_object* l_Lean_IR_EmitC_emitUSet___redArg___closed__0; static lean_object* l_Lean_IR_EmitC_emitInitFn___closed__0; +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSpreadArgs___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_structIncFnPrefix___closed__0; lean_object* lean_nat_shiftr(lean_object*, lean_object*); lean_object* l_Array_unzip___redArg(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreadArgs_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitFileHeader___closed__19; static lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_EmitC_emitInitFn_spec__0___redArg___closed__1; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitStructDecls_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__32; LEAN_EXPORT uint8_t l_Option_instBEq_beq___at___00Lean_IR_EmitC_emitMainFn_spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitBox(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitBox(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_utf8_byte_size(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitReset_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSetTag(lean_object*, lean_object*, lean_object*, lean_object*); @@ -185,60 +246,88 @@ lean_object* lean_string_push(lean_object*, uint32_t); static lean_object* l_Lean_IR_EmitC_toCName___closed__1; static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__15; lean_object* l_Lean_Expr_appArg_x21(lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); +lean_object* l_Array_ofFn___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__0___redArg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instMonad___redArg(lean_object*); +static lean_object* l_Lean_IR_EmitC_emitStructUnboxFn___closed__2; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_isIf___boxed(lean_object*); uint8_t l_Lean_hasInitAttr(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUSet___redArg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_structType(lean_object*); +static lean_object* l_Lean_IR_EmitC_emitStructReshapeFn___closed__2; static lean_object* l_Lean_IR_EmitC_emitInitFn___closed__9; +static lean_object* l_Lean_IR_EmitC_emitStructDefn___closed__7; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSimpleExternalCall___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__5; +static lean_object* l_Lean_IR_EmitC_emitStructIncDecFn___lam__0___closed__1; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_declareVars(lean_object*, uint8_t, lean_object*, lean_object*); uint32_t lean_string_utf8_get_fast(lean_object*, lean_object*); lean_object* l_Lean_IR_Decl_resultType(lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_structReshapeFnPrefix___closed__0; lean_object* l_Lean_IR_ensureHasDefault(lean_object*); +static lean_object* l_Lean_IR_EmitC_emitStructUnboxFn___closed__5; +LEAN_EXPORT uint8_t l_Lean_IR_EmitC_compatibleReshape(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructUnboxFn_spec__0___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_IR_EmitC_emitMainFn___closed__38; static lean_object* l_Lean_IR_EmitC_emitIf___closed__0; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0_spec__0(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitInitFn_spec__2_spec__2___redArg(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitFullAppArg_spec__0___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_IR_EmitC_emitIncOfType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__2; +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructUnboxFn(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSet___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFnBody(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFullAppArg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); lean_object* l_Nat_reprFast(lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_lookupStruct___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitSSet___closed__6; LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitPartialApp_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_toHexDigit(lean_object*); static lean_object* l_Lean_IR_EmitC_emitSet___redArg___closed__0; -static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__56; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructUnboxFn_spec__0___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__53; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_hasMainFn___lam__0___boxed(lean_object*); static lean_object* l_panic___at___00Lean_IR_EmitC_emitMainFn_spec__2___closed__0; size_t lean_usize_of_nat(lean_object*); +static lean_object* l_Lean_IR_EmitC_emitSpreadValue___closed__1; static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__26; -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFullApp_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_IR_EmitC_toCType___closed__8; -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCtor(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCtor(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_toCType___closed__0; LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFileFooter___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitIsShared(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructBox_spec__0___redArg___closed__1; LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitPartialApp_spec__0___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forM___at___00Lean_IR_EmitC_emitFnDecls_spec__3___closed__1; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitStructBox___closed__4; +static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__1; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_IR_EmitC_emitDec___redArg___closed__1; +static lean_object* l_Lean_IR_EmitC_emitBoxFn___closed__2; static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__18; -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_declareVar___redArg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_foldrM___at___00Lean_IR_EmitC_emitFnDecls_spec__2(lean_object*, lean_object*); extern lean_object* l___private_Lean_Compiler_ModPkgExt_0__Lean_modPkgExt; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitNumLit___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_IR_EmitC_emitInc___redArg___closed__4; +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructUnboxFn___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitStructUnboxFn___closed__3; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructBox_spec__0___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDecls_spec__2___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_getModName(lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitNumLit___redArg___closed__4; static lean_object* l_Lean_IR_EmitC_emitFileHeader___closed__3; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitReuse___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitUProj___closed__2; +static lean_object* l_Lean_IR_EmitC_emitStructDefn___closed__0; +static lean_object* l_Lean_IR_EmitC_emitStructBox___closed__2; static lean_object* l_Lean_IR_EmitC_leanMainFn___closed__0; +static lean_object* l_Lean_IR_EmitC_emitStructBox___closed__6; uint8_t l_Lean_IR_usesModuleFrom(lean_object*, lean_object*); extern lean_object* l_Lean_IR_instInhabitedParam_default; +uint8_t l_Lean_IR_IRType_compatibleWith(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLhs___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitIsShared___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_IR_Decl_params(lean_object*); @@ -248,77 +337,111 @@ lean_object* lean_get_init_fn_name_for(lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__2; LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg(lean_object*, lean_object*, lean_object*, lean_object*); uint32_t l_Nat_digitChar(lean_object*); -static lean_object* l_Lean_IR_EmitC_emitSSet___redArg___closed__2; +static lean_object* l_Lean_IR_EmitC_emitStructDefn___closed__5; uint64_t lean_uint64_shift_right(uint64_t, uint64_t); +static lean_object* l_Lean_IR_EmitC_lookupStruct___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emit___redArg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_quoteString(lean_object*); -static lean_object* l_Lean_IR_EmitC_emitBoxFn___redArg___closed__0; -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_toCType(lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitBoxFn(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_toCType(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitBoxFn(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at___00Lean_IR_EmitC_emitFnDecls_spec__1(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_isIOUnitBuiltinInitFn(lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitLns___redArg___closed__9; +static lean_object* l_Lean_IR_EmitC_emitIncOfType___closed__4; +static lean_object* l_Lean_IR_EmitC_emitSProj___closed__5; +static lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg___closed__1; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitFullAppArgs_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitStructDefn___closed__1; +static lean_object* l_Lean_IR_EmitC_emitUSet___closed__6; static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__3; extern lean_object* l_Lean_closureMaxArgs; -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUSet(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitUSet___closed__7; +static lean_object* l_Lean_IR_EmitC_emitSProj___closed__3; +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUSet(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitIsShared___redArg___closed__0; +lean_object* l_Lean_IR_IRType_hashApprox___boxed(lean_object*); LEAN_EXPORT uint8_t l_Lean_IR_EmitC_overwriteParam(lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitSetTag___redArg___closed__0; +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSpreads(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructIncDecFn_spec__0_spec__0___redArg(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitSProj___closed__0; static lean_object* l_Lean_IR_EmitC_emitTailCall___closed__1; +static lean_object* l_Lean_IR_EmitC_emitSSet___closed__5; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructIncDecFn_spec__0_spec__0(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__51; lean_object* l_EStateM_instMonad___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_getEnv(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUnbox___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUnbox___redArg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_IR_EmitC_emitFnDeclAux___closed__8; static lean_object* l_Lean_IR_EmitC_getJPParams___closed__0; static lean_object* l_Lean_IR_EmitC_emitTailCall___closed__2; uint8_t l_Lean_isExternC(lean_object*, lean_object*); -static lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0___redArg___closed__0; +static lean_object* l_Lean_IR_EmitC_emitSSet___closed__3; LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_EmitC_emitMainFn_spec__2(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00__private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1_spec__1___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitDeclInit___closed__4; static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__50; static lean_object* l_Lean_IR_EmitC_emitLns___redArg___closed__6; +static lean_object* l_Lean_IR_EmitC_emitBoxFn___closed__5; lean_object* l_Lean_IR_Decl_name(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFileFooter(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_compatibleReshape___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructReshapeFn___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_EmitC_0__Lean_IR_EmitC_emitSpreadArg_match__10_splitter(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at___00Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0_spec__0___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUSet___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_argToCString(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_declareParams___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_IR_EmitC_emitInc___redArg___closed__1; extern lean_object* l_Lean_instInhabitedConstantInfo_default; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitCtor_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitReuse___closed__0; -static lean_object* l_Lean_IR_EmitC_emitFullApp___closed__0; -static lean_object* l_Lean_IR_EmitC_emitInc___redArg___closed__3; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitFullAppArgs_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitNumLit___redArg___closed__1; +LEAN_EXPORT uint8_t l_Lean_IR_EmitC_needsRC(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLn(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitCtor_spec__0___redArg___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_IR_EmitC_emitApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitSSet___closed__7; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitTag___redArg___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_instInhabitedForall___redArg___lam__0___boxed(lean_object*, lean_object*); uint8_t l_Lean_IR_IRType_isObj(lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructUnboxFn___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitNumLit___redArg___closed__2; +static lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg___closed__2; static lean_object* l_Lean_IR_EmitC_emitCase___closed__1; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitExternDeclAux(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_EmitC_emitInitFn_spec__0___redArg___closed__0; +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__0; lean_object* l_Lean_IR_FnBody_body(lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitCtor_spec__1___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitIncOfType___closed__0; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___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_IR_EmitC_emitFullAppArgs(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadArg_spec__0___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_IR_EmitC_emitArgs___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructIncDecFn___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*); -static lean_object* l_Lean_IR_EmitC_emitFnDeclAux___closed__7; static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__37; LEAN_EXPORT lean_object* l_List_forM___at___00Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0_spec__0___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitCtorSetArgs_spec__0___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_IR_StructRC_visitDecl(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCase(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__1; static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__30; -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_declareParams_spec__0___redArg(lean_object*, size_t, size_t, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitNumLit___redArg___closed__3; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDecls_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_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0_spec__0___redArg(lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitDel___redArg___closed__0; static lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitReset_spec__0___redArg___closed__0; -static lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___closed__0; +static lean_object* l_Lean_IR_EmitC_emitDeclInit___closed__5; static lean_object* l_Lean_IR_EmitC_emitFileHeader___closed__12; static lean_object* l_Lean_IR_EmitC_emitFileHeader___closed__23; +static lean_object* l_Lean_IR_EmitC_emitIncOfType___closed__2; +static lean_object* l_Lean_IR_EmitC_emitCtor___closed__1; +static lean_object* l_Lean_IR_EmitC_emitStructBox___closed__0; static lean_object* l_Lean_IR_EmitC_emitInitFn___closed__8; +lean_object* l_EStateM_instInhabited___redArg(lean_object*); static lean_object* l_Lean_IR_EmitC_emitNumLit___redArg___closed__6; LEAN_EXPORT lean_object* l_List_mapTR_loop___at___00Lean_IR_EmitC_toStringArgs_spec__0(lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__46; @@ -328,6 +451,7 @@ LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at static lean_object* l_Lean_IR_EmitC_emitCtor___closed__0; LEAN_EXPORT lean_object* l_List_forM___at___00Lean_IR_EmitC_emitInitFn_spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_getModInitFn___closed__0; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitJmp___closed__0; LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_throwInvalidExportName(lean_object*, lean_object*, lean_object*, lean_object*); @@ -335,88 +459,122 @@ static lean_object* l_Lean_IR_EmitC_emitInitFn___closed__10; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitMarkPersistent___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_EmitC_emitInitFn_spec__0(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_IR_EmitC_hasMainFn___lam__0(lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_constName_x3f(lean_object*); static lean_object* l_Lean_IR_EmitC_emitApp___closed__1; static lean_object* l_Lean_IR_EmitC_emitDeclAux___closed__1; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLhs___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCase___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_throwUnknownVar___redArg(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_EmitC_needsRC_spec__0(lean_object*, size_t, size_t); static lean_object* l_Lean_IR_EmitC_emitTailCall___closed__0; -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUProj___redArg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructIncDecFn_spec__0_spec__0___redArg___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_IR_EmitC_emitExternCall___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUnionSwitchWithImpossible___lam__0___boxed(lean_object*); static lean_object* l_Lean_IR_EmitC_emitJmp___closed__1; static lean_object* l_Lean_IR_EmitC_toCInitName___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSet(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__36; +LEAN_EXPORT lean_object* l_panic___at___00Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0_spec__1___redArg(lean_object*, lean_object*); +static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__0; +static lean_object* l_Lean_IR_EmitC_emitStructDefn___closed__6; +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__0___redArg___closed__0; static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__6; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLn___redArg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitDeclAux___closed__0; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructBox_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitUnionSwitch_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructUnboxFn_spec__0___redArg___closed__0; static lean_object* l_Lean_IR_EmitC_getDecl___closed__0; static lean_object* l_Lean_IR_EmitC_emitFileHeader___closed__17; uint8_t lean_name_eq(lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_toCType___closed__1; +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitDecOfType(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); +static lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__4; +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__2; LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitIncOfType___closed__1; +static lean_object* l_Lean_IR_EmitC_emitStructDefn___closed__4; +static lean_object* l_Lean_IR_EmitC_emitBoxFn___closed__1; static lean_object* l_Lean_IR_EmitC_emitFileHeader___closed__24; static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__9; LEAN_EXPORT lean_object* l_List_forM___at___00Lean_IR_EmitC_emitFnDecls_spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitBox___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUnionSwitchWithImpossible___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitTailCall___closed__3; lean_object* l_Lean_expandExternPattern(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructBox___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitInitFn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLn___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint64_t l_Lean_IR_IRType_hashApprox(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitIf___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__42; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_declareParams_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitBlock(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitProj___closed__0; +lean_object* l_Lean_IR_ExplicitBoxing_mkBoxedName(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitReset___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitStructBox___closed__3; static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__8; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitPartialApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitOffset(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_IR_EmitC_toCType___closed__9; LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___at___00Lean_IR_EmitC_emitSimpleExternalCall_spec__0___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__1___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_structReshapeFnPrefix; +static lean_object* l_Lean_IR_EmitC_emitSProj___closed__4; +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__1___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitFnDeclAux___closed__0; -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_toCInitName(lean_object*, lean_object*, lean_object*); -static lean_object* l_panic___at___00Lean_IR_EmitC_toCType_spec__0___closed__0; static lean_object* l_Lean_IR_EmitC_emitLns___redArg___closed__2; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitVDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_declareVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitBlock___closed__1; -static lean_object* l_Lean_IR_EmitC_emitInc___redArg___closed__2; -static lean_object* l_Lean_IR_EmitC_emitSSet___redArg___closed__3; static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__41; +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructDeclsFor(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_overwriteParam___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitSProj___closed__6; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitStructBoxFn___closed__2; +static lean_object* l_Lean_IR_EmitC_emitUSet___closed__2; +static lean_object* l_Lean_IR_EmitC_emitUProj___closed__0; lean_object* l_Lean_getBuiltinInitFnNameFor_x3f(lean_object*, lean_object*); lean_object* l_Lean_getExternNameFor(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitFullAppArgs___closed__0; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitCase_spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitUnionSwitchWithImpossible___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSimpleExternalCall(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_isTailCall(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitReshapeFn(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFullAppArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFullApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_NameSet_insert(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitIf(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__39; +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadArg_spec__0___redArg___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emit___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_IR_EmitC_emitFnDeclAux___closed__9; lean_object* lean_array_get_borrowed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSSet___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadArg_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitPartialApp_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitDeclInit___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_getEnv___boxed(lean_object*, lean_object*); lean_object* l_EStateM_map(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__1(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_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_EmitC_emitInitFn_spec__0___redArg(lean_object*, size_t, size_t, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__33; -static lean_object* l_Lean_IR_EmitC_emitBoxFn___redArg___closed__3; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructIncDecFn_spec__0___redArg(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructBox_spec__0___redArg___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitJmp(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitFullAppArgs_spec__0___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSpreadArgs(lean_object*, uint8_t, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__0___redArg___closed__1; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitFullAppArg_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitFileHeader___closed__2; +static lean_object* l_Lean_IR_EmitC_structDecFnPrefix___closed__0; static lean_object* l_Lean_IR_EmitC_emitFileHeader___closed__8; static lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__1___redArg___closed__0; static lean_object* l_Lean_IR_EmitC_emitNumLit___redArg___closed__0; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_toStringWithToken___at___00Lean_Name_toString_spec__0(lean_object*, uint8_t); static lean_object* l_Lean_IR_EmitC_emitOffset___redArg___closed__1; LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -424,15 +582,25 @@ static lean_object* l_Lean_IR_EmitC_quoteString___closed__0; static lean_object* l_Lean_IR_EmitC_emitApp___closed__4; static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__55; LEAN_EXPORT uint8_t l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__0___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_IR_EmitC_emitBoxFn___redArg___closed__2; +static lean_object* l_Lean_IR_EmitC_emitUnionSwitch___closed__0; +static lean_object* l_Lean_IR_EmitC_emitTag___redArg___closed__1; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitDeclAux(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitTag(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitStructUnboxFn___closed__1; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitIf___closed__1; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_declareVars___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitIncOfType___closed__3; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitFileHeader___closed__11; +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_EmitC_0__Lean_IR_EmitC_emitSpreadArg_match__4_splitter___redArg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitLns___redArg___closed__4; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitCtor_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_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSet___redArg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_IR_EmitC_emitStructReshapeFn___lam__0(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00__private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1_spec__1___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitSProj___closed__2; LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_leanMainFn; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); @@ -440,94 +608,103 @@ lean_object* l_EStateM_pure(lean_object*, lean_object*, lean_object*, lean_objec LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCtorScalarSize___redArg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_getModName___boxed(lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitInitFn___closed__2; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLns___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitOffset___redArg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitPartialApp___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__0___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitTag___redArg(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_IR_EmitC_toCType___closed__10; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_EmitC_emitInitFn_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_main___closed__0; static lean_object* l_Lean_IR_EmitC_emitInitFn___closed__7; extern lean_object* l_Lean_NameSet_empty; LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitCtorSetArgs_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSProj___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_IR_EmitC_emitSProj___redArg___closed__1; +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__1___redArg___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCtorScalarSize(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitInitFn_spec__2_spec__2___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitDeclInit(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__31; lean_object* lean_string_length(lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSSet___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitProj___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitDec___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitStructUnboxFn___closed__4; +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitProj___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_isTailCall___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitInitFn_spec__2_spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_IR_EmitC_emitSProj___redArg___closed__2; lean_object* l_Lean_Expr_getForallBody(lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitArgs_spec__0___redArg(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__1___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSSet___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSSet___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__11; -static lean_object* l_Lean_IR_EmitC_emitDeclAux___closed__3; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -static lean_object* l_Lean_IR_EmitC_emitSProj___redArg___closed__3; +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructBoxFn(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitDel___redArg(lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_toCType___closed__2; lean_object* lean_nat_mod(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitUnionSwitchWithImpossible_spec__0(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitLit___redArg___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_throwUnknownVar(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadArg_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__21; +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_structIncFnPrefix; uint8_t lean_uint32_dec_eq(uint32_t, uint32_t); static lean_object* l_Lean_IR_EmitC_emitInitFn___closed__6; LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__0___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitAllocCtor(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_getJPParams___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitLns___redArg___closed__10; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDecls_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_IR_EmitC_emitMainFn___closed__28; lean_object* l_Lean_IR_getDecls(lean_object*); +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___lam__0___closed__0; static lean_object* l_Lean_IR_EmitC_toCType___closed__5; lean_object* l_Lean_Environment_getModuleIdx_x3f(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_toCType___closed__6; +static lean_object* l_Lean_IR_EmitC_emitStructIncDecFn___closed__0; static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__13; +LEAN_EXPORT uint8_t l_Array_isEqvAux___at___00Lean_IR_EmitC_compatibleReshape_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_IR_isTailCallTo(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_foldrM___at___00Lean_IR_EmitC_emitFnDecls_spec__2___boxed(lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitFnDeclAux___closed__6; LEAN_EXPORT lean_object* l_Option_instBEq_beq___at___00Lean_IR_EmitC_emitMainFn_spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitDel(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitStructReshapeFn___closed__0; static lean_object* l_Lean_IR_EmitC_emitFileFooter___redArg___closed__1; uint8_t l_Lean_IR_ExplicitBoxing_isBoxedName(lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__1___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUnionSwitchWithImpossible___lam__0(lean_object*); lean_object* l_Lean_IR_getUnboxOpName(lean_object*); -static lean_object* l_Lean_IR_EmitC_emitBoxFn___redArg___closed__4; static lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__2___redArg___closed__0; -static lean_object* l_Lean_IR_EmitC_emitSProj___redArg___closed__0; LEAN_EXPORT uint8_t l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__1___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__1; LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_NameSet_contains(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCInitName(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitInc(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*); +static lean_object* l_panic___at___00Lean_IR_EmitC_emitUSet_spec__0___closed__0; lean_object* l_Lean_IR_Alt_body(lean_object*); uint64_t lean_uint64_xor(uint64_t, uint64_t); static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__40; static lean_object* l_Lean_IR_EmitC_emitLn___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDecOfType_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitPartialApp___closed__1; static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__54; LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitArgs_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -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_emitSSet___redArg___closed__6; -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructBox_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructBox_spec__0___redArg___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_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitUnionSwitch_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_reverse___redArg(lean_object*); +static lean_object* l_Lean_IR_EmitC_emitUProj___closed__3; 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*); +static lean_object* l_Lean_IR_EmitC_emitStructUnboxFn___lam__0___closed__0; lean_object* lean_nat_sub(lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitStructUnboxFn___closed__0; static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__16; lean_object* l_Lean_IR_mkVarJPMaps(lean_object*); static lean_object* l_Lean_IR_EmitC_emitJmp___closed__2; @@ -536,6 +713,9 @@ LEAN_EXPORT lean_object* l_Lean_IR_EmitC_toHexDigit___boxed(lean_object*); static lean_object* l_Lean_IR_EmitC_emitReset___closed__2; static lean_object* l_Lean_IR_EmitC_emitFileFooter___redArg___closed__0; static lean_object* l_Lean_IR_EmitC_emitFileHeader___closed__22; +static lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg___closed__3; +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUnboxFn(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_lookupStruct(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitLns___redArg___closed__1; LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitReset_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitDeclAux___closed__2; @@ -544,25 +724,29 @@ LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__48; static lean_object* l_Lean_IR_EmitC_emitFileHeader___closed__10; -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUProj___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0_spec__0___redArg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at___00Lean_IR_EmitC_emitFns_spec__0(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitIf___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_EmitC_0__Lean_IR_EmitC_emitSpreadArg_match__4_splitter(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__2___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitBoxFn___redArg(lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitFileHeader___closed__13; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_throwUnknownVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_main___lam__0(lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitFileHeader___closed__9; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_main(lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitUSet___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFileHeader(lean_object*, lean_object*); static lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitInitFn_spec__2___closed__0; static lean_object* l_Lean_IR_EmitC_emitApp___closed__2; lean_object* l_Lean_Environment_imports(lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitProj___redArg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitTailCall_spec__3(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0___redArg___boxed(lean_object*, lean_object*); size_t lean_usize_sub(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructDecls(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_EStateM_seqRight(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_IR_IRType_isStruct(lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__52; static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__1; lean_object* l_Lean_IR_declToString(lean_object*); @@ -570,130 +754,187 @@ LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__5; uint64_t l_Lean_IR_instHashableJoinPointId_hash(lean_object*); static lean_object* l_Lean_IR_EmitC_emitApp___closed__3; -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSProj___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSProj___boxed(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_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitInitFn_spec__2_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitFullAppArgs_spec__0___redArg___lam__0(uint8_t, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitFnDeclAux___closed__1; static lean_object* l_Lean_IR_EmitC_emitPartialApp___closed__0; +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructBox(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); lean_object* l_mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__1; +static lean_object* l_Lean_IR_EmitC_emitStructDecls___redArg___closed__0; +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructIncDecFn(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitInitFn_spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitNumLit___redArg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitArgs_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructReshapeFn(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitArg___redArg(lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__27; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFns(lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_lookupStruct___closed__1; lean_object* lean_array_uget(lean_object*, size_t); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitArgs(lean_object*, lean_object*, lean_object*); size_t lean_array_size(lean_object*); -static lean_object* l_Lean_IR_EmitC_emitBoxFn___redArg___closed__5; +static lean_object* l_Lean_IR_EmitC_emitSSet___closed__2; static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__49; +static lean_object* l_Lean_IR_EmitC_emitStructReshapeFn___lam__1___closed__1; static lean_object* l_Lean_IR_EmitC_emitFnDeclAux___closed__3; -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitBox___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitReshapeFn___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDecOfType_spec__0___redArg___closed__1; +extern lean_object* l_Lean_IR_instInhabitedStructTypeInfo_default; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0(lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitBox___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUnionSwitchWithImpossible(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitReset___closed__0; -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitInc___redArg(lean_object*, lean_object*, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructIncDecFn_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr1(lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitUnionSwitch_spec__0___redArg___closed__0; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDecls_spec__1___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_toCName(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0_spec__0___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_toCType___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSpreadValue(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitUnionSwitchWithImpossible_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_toCType___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitExternCall(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__3; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitDeclInit___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitTailCall(lean_object*, lean_object*, lean_object*); +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__4; static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__35; +static lean_object* l_Lean_IR_EmitC_emitStructReshapeFn___closed__4; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0_spec__0(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitFileHeader___closed__14; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_EmitC_emitInitFn_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSpreadArg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__2; +static lean_object* l_Lean_IR_EmitC_emitUProj___closed__1; static lean_object* l_Lean_IR_EmitC_emitCtorScalarSize___redArg___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitDel___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitBoxFn___closed__3; lean_object* lean_string_append(lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitFnDeclAux___closed__2; static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__14; -static lean_object* l_Lean_IR_EmitC_emitUProj___redArg___closed__0; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructDecls___redArg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_structType___closed__0; static lean_object* l_Lean_IR_EmitC_emitLns___redArg___closed__5; static lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitCase_spec__4___closed__1; static lean_object* l_Lean_IR_EmitC_emitNumLit___redArg___closed__5; lean_object* lean_array_get_size(lean_object*); +static lean_object* l_Lean_IR_EmitC_emitUSet___closed__3; +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructDefn(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitUSet___closed__1; +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructIncDecFn_spec__0___redArg___closed__0; +static lean_object* l_Lean_IR_EmitC_emitStructReshapeFn___lam__1___closed__0; static lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__0; -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCtor___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_IR_EmitC_emitProj___redArg___closed__0; -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitBoxFn___redArg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCtor___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_isIOUnitInitFn(lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitExternCall___closed__0; uint8_t lean_nat_dec_le(lean_object*, lean_object*); +lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitDeclInit___closed__2; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_throwInvalidExportName___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructReshapeFn___lam__1(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitOffset___redArg___closed__0; static lean_object* l_Lean_IR_EmitC_emitTag___redArg___closed__0; uint8_t lean_usize_dec_lt(size_t, size_t); +static lean_object* l_Lean_IR_EmitC_emitBoxFn___closed__0; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitCtor_spec__1___redArg___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_IR_EmitC_emitApp___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_String_Slice_positions(lean_object*); static lean_object* l_Lean_IR_EmitC_throwUnknownVar___redArg___closed__0; +uint64_t l_Lean_IR_instHashableVarId_hash(lean_object*); lean_object* l_Lean_IR_collectUsedDecls(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_IR_EmitC_emitSimpleExternalCall___closed__0; +static lean_object* l_Lean_IR_EmitC_emitStructDefn___closed__9; LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___at___00Lean_IR_EmitC_emitSimpleExternalCall_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00__private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1_spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_isEqvAux___at___00Lean_IR_EmitC_compatibleReshape_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_List_any___redArg(lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_getDecl___closed__1; static lean_object* l_Lean_IR_EmitC_emitInitFn___closed__1; lean_object* l_Lean_IR_Decl_normalizeIds(lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDecls_spec__2___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadArg_spec__0___redArg(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitSProj___closed__1; static lean_object* l_Lean_IR_EmitC_overwriteParam___closed__0; static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__43; +static lean_object* l_Lean_IR_EmitC_emitSSet___closed__1; LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitUSet___closed__5; static lean_object* l_Lean_IR_EmitC_emitLns___redArg___closed__3; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSetTag___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_IR_EmitC_emitSSet___redArg___closed__4; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructUnboxFn_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitFileHeader___closed__6; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitJPs(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitReset___closed__1; +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_structDecFnPrefix; +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitDecOfType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__0___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___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0_spec__0___redArg___boxed(lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitAllocCtor___redArg___closed__0; -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitDec___redArg(lean_object*, lean_object*, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitFullAppArg_spec__0___redArg(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at___00Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0_spec__0___redArg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitStructBoxFn___closed__1; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitIsShared___redArg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCtorScalarSize___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_IR_EmitC_emitSSet___redArg___closed__0; -LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_EmitC_toCType_spec__0(lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDecls_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_needsRC___boxed(lean_object*); +static lean_object* l_Lean_IR_EmitC_emitSSet___closed__0; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDecOfType_spec__0(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_getDecl(lean_object*, lean_object*, lean_object*); lean_object* l_Array_zip___redArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSpreads___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitMainFnIfNeeded(lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__0; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0___redArg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitBoxFn___closed__4; static lean_object* l_Lean_IR_EmitC_emitFileHeader___closed__7; -static lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__0___redArg___closed__2; +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitIncOfType(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitFullAppArgs_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructUnboxFn_spec__0___redArg___closed__1; static lean_object* l_Lean_IR_EmitC_emitFileHeader___closed__5; +static lean_object* l_Lean_IR_EmitC_emitStructDefn___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_EmitC_0__Lean_IR_EmitC_emitSpreadArg_match__10_splitter___redArg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__0; +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructIncDecFn___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_IR_IRType_isScalar(lean_object*); static lean_object* l_Lean_IR_EmitC_emitJPs___closed__0; -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSProj___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitTailCall_spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDecOfType_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_toCType___closed__7; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_getModInitFn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitOffset___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Environment_getModulePackageByIdx_x3f(lean_object*, lean_object*); +static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitInitFn_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitArgs_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_emitC(lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitStructReshapeFn___closed__3; size_t lean_usize_land(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUnionSwitch___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitCtor___closed__3; LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFnDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUnbox___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__29; static lean_object* l_Lean_IR_emitC___closed__1; static lean_object* l_Lean_IR_EmitC_emitFileHeader___closed__18; +static lean_object* l_Lean_IR_EmitC_declareVar___closed__0; +static lean_object* l_Lean_IR_EmitC_emitStructIncDecFn___closed__2; static lean_object* l_Lean_IR_EmitC_emitFileHeader___closed__20; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_paramEqArg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCtorSetArgs___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_IR_EmitC_emitTailCall___closed__4; LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitReset_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_toCType___closed__3; +static lean_object* l_Lean_IR_EmitC_emitSpreadValue___closed__0; static lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitCase_spec__4___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_declareVar(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* _init_l_Lean_IR_EmitC_leanMainFn___closed__0() { @@ -1160,130 +1401,247 @@ lean_dec_ref(x_2); return x_4; } } -static lean_object* _init_l_panic___at___00Lean_IR_EmitC_toCType_spec__0___closed__0() { +static lean_object* _init_l_Lean_IR_EmitC_lookupStruct___closed__0() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("", 0, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_IR_instBEqIRTypeApprox___lam__0___boxed), 2, 0); return x_1; } } -LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_EmitC_toCType_spec__0(lean_object* x_1) { +static lean_object* _init_l_Lean_IR_EmitC_lookupStruct___closed__1() { _start: { -lean_object* x_2; lean_object* x_3; -x_2 = l_panic___at___00Lean_IR_EmitC_toCType_spec__0___closed__0; -x_3 = lean_panic_fn(x_2, x_1); -return x_3; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_IR_IRType_hashApprox___boxed), 1, 0); +return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_toCType___closed__0() { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_lookupStruct(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("double", 6, 6); -return x_1; +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_4 = lean_ctor_get(x_2, 6); +x_5 = l_Lean_IR_EmitC_lookupStruct___closed__0; +x_6 = l_Lean_IR_EmitC_lookupStruct___closed__1; +x_7 = lean_unsigned_to_nat(0u); +x_8 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___redArg(x_5, x_6, x_7, x_4, x_1); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_3); +return x_9; } } -static lean_object* _init_l_Lean_IR_EmitC_toCType___closed__1() { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_lookupStruct___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_IR_EmitC_lookupStruct(x_1, x_2, x_3); +lean_dec_ref(x_2); +return x_4; +} +} +static lean_object* _init_l_Lean_IR_EmitC_structType___closed__0() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("uint8_t", 7, 7); +x_1 = lean_mk_string_unchecked("struct l_s", 10, 10); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_toCType___closed__2() { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_structType(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Lean_IR_EmitC_structType___closed__0; +x_3 = l_Nat_reprFast(x_1); +x_4 = lean_string_append(x_2, x_3); +lean_dec_ref(x_3); +return x_4; +} +} +static lean_object* _init_l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg___closed__2() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("uint16_t", 8, 8); +x_1 = lean_mk_string_unchecked("key is not present in hash table", 32, 32); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_toCType___closed__3() { +static lean_object* _init_l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("uint32_t", 8, 8); +x_1 = lean_mk_string_unchecked("Std.DHashMap.Internal.AssocList.get!", 36, 36); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_toCType___closed__4() { +static lean_object* _init_l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg___closed__0() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("uint64_t", 8, 8); +x_1 = lean_mk_string_unchecked("Std.Data.DHashMap.Internal.AssocList.Basic", 42, 42); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_toCType___closed__5() { +static lean_object* _init_l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg___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_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg___closed__2; +x_2 = lean_unsigned_to_nat(11u); +x_3 = lean_unsigned_to_nat(163u); +x_4 = l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg___closed__1; +x_5 = l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg___closed__0; +x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg(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_object* x_5; +x_4 = l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg___closed__3; +x_5 = lean_panic_fn(x_1, x_4); +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; uint8_t x_10; +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 = 1; +x_10 = l_Lean_IR_IRType_compatibleWith(x_6, x_2, x_9); +if (x_10 == 0) +{ +x_3 = x_8; +goto _start; +} +else +{ +lean_dec(x_1); +lean_inc(x_7); +return x_7; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg(x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; uint64_t 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; size_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_array_get_size(x_4); +x_6 = l_Lean_IR_IRType_hashApprox(x_3); +x_7 = 32; +x_8 = lean_uint64_shift_right(x_6, x_7); +x_9 = lean_uint64_xor(x_6, x_8); +x_10 = 16; +x_11 = lean_uint64_shift_right(x_9, x_10); +x_12 = lean_uint64_xor(x_9, x_11); +x_13 = lean_uint64_to_usize(x_12); +x_14 = lean_usize_of_nat(x_5); +x_15 = 1; +x_16 = lean_usize_sub(x_14, x_15); +x_17 = lean_usize_land(x_13, x_16); +x_18 = lean_array_uget(x_4, x_17); +x_19 = l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg(x_1, x_3, x_18); +lean_dec(x_18); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0___redArg(x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_panic___at___00Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0_spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_panic_fn(x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_panic___at___00Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0_spec__1___redArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_panic_fn(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_IR_EmitC_toCType___closed__0() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("size_t", 6, 6); +x_1 = lean_mk_string_unchecked("double", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_toCType___closed__6() { +static lean_object* _init_l_Lean_IR_EmitC_toCType___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("float", 5, 5); +x_1 = lean_mk_string_unchecked("uint8_t", 7, 7); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_toCType___closed__7() { +static lean_object* _init_l_Lean_IR_EmitC_toCType___closed__2() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Compiler.IR.EmitC", 22, 22); +x_1 = lean_mk_string_unchecked("uint16_t", 8, 8); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_toCType___closed__8() { +static lean_object* _init_l_Lean_IR_EmitC_toCType___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.IR.EmitC.toCType", 21, 21); +x_1 = lean_mk_string_unchecked("uint32_t", 8, 8); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_toCType___closed__9() { +static lean_object* _init_l_Lean_IR_EmitC_toCType___closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("not implemented yet", 19, 19); +x_1 = lean_mk_string_unchecked("uint64_t", 8, 8); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_toCType___closed__10() { +static lean_object* _init_l_Lean_IR_EmitC_toCType___closed__5() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Lean_IR_EmitC_toCType___closed__9; -x_2 = lean_unsigned_to_nat(25u); -x_3 = lean_unsigned_to_nat(76u); -x_4 = l_Lean_IR_EmitC_toCType___closed__8; -x_5 = l_Lean_IR_EmitC_toCType___closed__7; -x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); -return x_6; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("size_t", 6, 6); +return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_toCType___closed__11() { +static lean_object* _init_l_Lean_IR_EmitC_toCType___closed__6() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Lean_IR_EmitC_toCType___closed__9; -x_2 = lean_unsigned_to_nat(25u); -x_3 = lean_unsigned_to_nat(77u); -x_4 = l_Lean_IR_EmitC_toCType___closed__8; -x_5 = l_Lean_IR_EmitC_toCType___closed__7; -x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); -return x_6; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("float", 5, 5); +return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_toCType___closed__12() { +static lean_object* _init_l_Lean_IR_EmitC_toCType___closed__7() { _start: { lean_object* x_1; @@ -1291,82 +1649,157 @@ x_1 = lean_mk_string_unchecked("lean_object*", 12, 12); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_toCType(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_toCType(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { switch (lean_obj_tag(x_1)) { case 0: { -lean_object* x_2; -x_2 = l_Lean_IR_EmitC_toCType___closed__0; -return x_2; +lean_object* x_4; lean_object* x_5; +x_4 = l_Lean_IR_EmitC_toCType___closed__0; +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_4); +lean_ctor_set(x_5, 1, x_3); +return x_5; } case 1: { -lean_object* x_3; -x_3 = l_Lean_IR_EmitC_toCType___closed__1; -return x_3; +lean_object* x_6; lean_object* x_7; +x_6 = l_Lean_IR_EmitC_toCType___closed__1; +x_7 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_3); +return x_7; } case 2: { -lean_object* x_4; -x_4 = l_Lean_IR_EmitC_toCType___closed__2; -return x_4; +lean_object* x_8; lean_object* x_9; +x_8 = l_Lean_IR_EmitC_toCType___closed__2; +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_3); +return x_9; } case 3: { -lean_object* x_5; -x_5 = l_Lean_IR_EmitC_toCType___closed__3; -return x_5; +lean_object* x_10; lean_object* x_11; +x_10 = l_Lean_IR_EmitC_toCType___closed__3; +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_3); +return x_11; } case 4: { -lean_object* x_6; -x_6 = l_Lean_IR_EmitC_toCType___closed__4; -return x_6; +lean_object* x_12; lean_object* x_13; +x_12 = l_Lean_IR_EmitC_toCType___closed__4; +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_3); +return x_13; } case 5: { -lean_object* x_7; -x_7 = l_Lean_IR_EmitC_toCType___closed__5; -return x_7; +lean_object* x_14; lean_object* x_15; +x_14 = l_Lean_IR_EmitC_toCType___closed__5; +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_3); +return x_15; } case 9: { -lean_object* x_8; -x_8 = l_Lean_IR_EmitC_toCType___closed__6; -return x_8; +lean_object* x_16; lean_object* x_17; +x_16 = l_Lean_IR_EmitC_toCType___closed__6; +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_3); +return x_17; } case 10: { -lean_object* x_9; lean_object* x_10; -x_9 = l_Lean_IR_EmitC_toCType___closed__10; -x_10 = l_panic___at___00Lean_IR_EmitC_toCType_spec__0(x_9); -return x_10; +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_18 = lean_ctor_get(x_2, 6); +x_19 = lean_unsigned_to_nat(0u); +x_20 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0___redArg(x_19, x_18, x_1); +x_21 = l_Lean_IR_EmitC_structType(x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_3); +return x_22; } case 11: { -lean_object* x_11; lean_object* x_12; -x_11 = l_Lean_IR_EmitC_toCType___closed__11; -x_12 = l_panic___at___00Lean_IR_EmitC_toCType_spec__0(x_11); -return x_12; +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_2, 6); +x_24 = lean_unsigned_to_nat(0u); +x_25 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0___redArg(x_24, x_23, x_1); +x_26 = l_Lean_IR_EmitC_structType(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_3); +return x_27; } default: { -lean_object* x_13; -x_13 = l_Lean_IR_EmitC_toCType___closed__12; -return x_13; +lean_object* x_28; lean_object* x_29; +x_28 = l_Lean_IR_EmitC_toCType___closed__7; +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_3); +return x_29; } } } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_toCType___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_2; -x_2 = l_Lean_IR_EmitC_toCType(x_1); +lean_object* x_5; +x_5 = l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0(x_1, x_2, x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0(x_1, x_2, x_3, x_4); +lean_dec(x_4); +lean_dec_ref(x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0___redArg(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec_ref(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_toCType___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_IR_EmitC_toCType(x_1, x_2, x_3); +lean_dec_ref(x_2); lean_dec(x_1); -return x_2; +return x_4; } } static lean_object* _init_l_Lean_IR_EmitC_throwInvalidExportName___redArg___closed__0() { @@ -1703,7 +2136,15 @@ return x_18; } } } -static lean_object* _init_l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___closed__0() { +static lean_object* _init_l_Lean_IR_EmitC_emitSpreadArg___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" ", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitSpreadArg___closed__1() { _start: { lean_object* x_1; @@ -1711,10164 +2152,18304 @@ x_1 = lean_mk_string_unchecked(", ", 2, 2); return x_1; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg(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_IR_EmitC_emitSpreadArg(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_6; uint8_t x_7; -x_6 = lean_unsigned_to_nat(0u); -x_7 = lean_nat_dec_eq(x_4, x_6); -if (x_7 == 1) +lean_object* x_6; lean_object* x_11; lean_object* x_12; lean_object* x_22; lean_object* x_23; +if (lean_obj_tag(x_1) == 10) { -lean_object* x_8; lean_object* x_9; -lean_dec(x_4); -x_8 = lean_box(0); -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_8); -lean_ctor_set(x_9, 1, x_5); -return x_9; +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_1, 1); +x_30 = lean_ctor_get(x_1, 2); +x_31 = lean_ctor_get(x_1, 3); +x_32 = lean_unsigned_to_nat(0u); +x_33 = lean_nat_dec_eq(x_30, x_32); +if (x_33 == 0) +{ +x_22 = x_4; +x_23 = x_5; +goto block_28; } else { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_21; -x_10 = lean_unsigned_to_nat(1u); -x_11 = lean_nat_sub(x_4, x_10); -lean_dec(x_4); -x_12 = lean_nat_sub(x_3, x_11); -x_13 = lean_nat_sub(x_12, x_10); -lean_dec(x_12); -x_21 = lean_nat_dec_lt(x_2, x_13); -if (x_21 == 0) +uint8_t x_34; +x_34 = lean_nat_dec_eq(x_31, x_32); +if (x_34 == 0) { -x_14 = x_5; -goto block_20; +x_22 = x_4; +x_23 = x_5; +goto block_28; } else { -lean_object* x_22; lean_object* x_23; -x_22 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___closed__0; -x_23 = lean_string_append(x_5, x_22); -x_14 = x_23; -goto block_20; -} -block_20: -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_15 = lean_array_fget_borrowed(x_1, x_13); -lean_dec(x_13); -x_16 = lean_ctor_get(x_15, 1); -x_17 = l_Lean_IR_EmitC_toCType(x_16); -x_18 = lean_string_append(x_14, x_17); -lean_dec_ref(x_17); -x_4 = x_11; -x_5 = x_18; -goto _start; -} +lean_object* x_35; lean_object* x_36; +x_35 = lean_array_get_size(x_29); +x_36 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadArg_spec__0___redArg(x_35, x_29, x_2, x_32, x_3, x_4, x_5); +return x_36; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: +else { -lean_object* x_8; -x_8 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg(x_1, x_2, x_3, x_4, x_7); -return x_8; -} +x_22 = x_4; +x_23 = x_5; +goto block_28; } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__1(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { -_start: +block_10: { -lean_object* x_5; uint8_t x_10; -x_10 = lean_usize_dec_eq(x_2, x_3); -if (x_10 == 0) +uint8_t x_7; lean_object* x_8; lean_object* x_9; +x_7 = 0; +x_8 = lean_box(x_7); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_6); +return x_9; +} +block_21: { -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = lean_array_uget(x_1, x_2); -x_12 = lean_ctor_get(x_11, 1); -lean_inc(x_12); -x_13 = l_Lean_IR_IRType_isErased(x_12); -lean_dec(x_12); -if (x_13 == 0) +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = l_Lean_IR_EmitC_toCType(x_1, x_11, x_12); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec_ref(x_13); +x_16 = lean_string_append(x_15, x_14); +lean_dec(x_14); +if (lean_obj_tag(x_2) == 1) { -lean_object* x_14; -x_14 = lean_array_push(x_4, x_11); -x_5 = x_14; -goto block_9; +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = lean_ctor_get(x_2, 0); +lean_inc(x_17); +lean_dec_ref(x_2); +x_18 = l_Lean_IR_EmitC_emitSpreadArg___closed__0; +x_19 = lean_string_append(x_16, x_18); +x_20 = lean_string_append(x_19, x_17); +lean_dec(x_17); +x_6 = x_20; +goto block_10; } else { -lean_dec(x_11); -x_5 = x_4; -goto block_9; +lean_dec(x_2); +x_6 = x_16; +goto block_10; } } -else +block_28: { -return x_4; -} -block_9: +if (lean_obj_tag(x_1) == 13) { -size_t x_6; size_t x_7; -x_6 = 1; -x_7 = lean_usize_add(x_2, x_6); -x_2 = x_7; -x_4 = x_5; -goto _start; -} -} +lean_object* x_24; lean_object* x_25; +lean_dec(x_2); +x_24 = lean_box(x_3); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_23); +return x_25; } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__2(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; uint8_t x_10; -x_10 = lean_usize_dec_eq(x_2, x_3); -if (x_10 == 0) +else { -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = lean_array_uget(x_1, x_2); -x_12 = lean_ctor_get(x_11, 1); -lean_inc(x_12); -x_13 = l_Lean_IR_IRType_isVoid(x_12); -lean_dec(x_12); -if (x_13 == 0) +if (x_3 == 0) { -lean_object* x_14; -x_14 = lean_array_push(x_4, x_11); -x_5 = x_14; -goto block_9; +lean_object* x_26; lean_object* x_27; +x_26 = l_Lean_IR_EmitC_emitSpreadArg___closed__1; +x_27 = lean_string_append(x_23, x_26); +x_11 = x_22; +x_12 = x_27; +goto block_21; } else { -lean_dec(x_11); -x_5 = x_4; -goto block_9; -} +x_11 = x_22; +x_12 = x_23; +goto block_21; } -else -{ -return x_4; } -block_9: -{ -size_t x_6; size_t x_7; -x_6 = 1; -x_7 = lean_usize_add(x_2, x_6); -x_2 = x_7; -x_4 = x_5; -goto _start; } } } -static lean_object* _init_l_Lean_IR_EmitC_emitFnDeclAux___closed__0() { +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadArg_spec__0___redArg___closed__0() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked(";", 1, 1); +x_1 = lean_mk_string_unchecked("_", 1, 1); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitFnDeclAux___closed__1() { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadArg_spec__0___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked(")", 1, 1); -return x_1; -} +uint8_t x_8; lean_object* x_9; uint8_t x_14; +x_14 = lean_nat_dec_lt(x_4, x_1); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +lean_dec(x_4); +lean_dec(x_3); +x_15 = lean_box(x_5); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_7); +return x_16; } -static lean_object* _init_l_Lean_IR_EmitC_emitFnDeclAux___closed__2() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_object**", 13, 13); -return x_1; +lean_object* x_17; lean_object* x_18; +x_17 = lean_array_fget_borrowed(x_2, x_4); +if (lean_obj_tag(x_17) == 6) +{ +x_8 = x_5; +x_9 = x_7; +goto block_13; } +else +{ +if (lean_obj_tag(x_3) == 0) +{ +x_18 = x_3; +goto block_23; } -static lean_object* _init_l_Lean_IR_EmitC_emitFnDeclAux___closed__3() { -_start: +else { -lean_object* x_1; -x_1 = l_Lean_closureMaxArgs; -return x_1; +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_24 = lean_ctor_get(x_3, 0); +x_25 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadArg_spec__0___redArg___closed__0; +lean_inc(x_24); +x_26 = lean_string_append(x_24, x_25); +lean_inc(x_4); +x_27 = l_Nat_reprFast(x_4); +x_28 = lean_string_append(x_26, x_27); +lean_dec_ref(x_27); +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, x_28); +x_18 = x_29; +goto block_23; } } -static lean_object* _init_l_Lean_IR_EmitC_emitFnDeclAux___closed__4() { -_start: +block_23: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked(" ", 1, 1); -return x_1; +lean_object* x_19; +x_19 = l_Lean_IR_EmitC_emitSpreadArg(x_17, x_18, x_5, x_6, x_7); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec_ref(x_19); +x_22 = lean_unbox(x_20); +lean_dec(x_20); +x_8 = x_22; +x_9 = x_21; +goto block_13; +} +else +{ +lean_dec(x_4); +lean_dec(x_3); +return x_19; } } -static lean_object* _init_l_Lean_IR_EmitC_emitFnDeclAux___closed__5() { -_start: +} +block_13: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("(", 1, 1); -return x_1; +lean_object* x_10; lean_object* x_11; +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_add(x_4, x_10); +lean_dec(x_4); +x_4 = x_11; +x_5 = x_8; +x_7 = x_9; +goto _start; } } -static lean_object* _init_l_Lean_IR_EmitC_emitFnDeclAux___closed__6() { +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadArg_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(0u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; +lean_object* x_11; +x_11 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadArg_spec__0___redArg(x_1, x_2, x_3, x_6, x_7, x_9, x_10); +return x_11; } } -static lean_object* _init_l_Lean_IR_EmitC_emitFnDeclAux___closed__7() { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadArg_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("LEAN_EXPORT ", 12, 12); -return x_1; +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_7); +x_12 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadArg_spec__0(x_1, x_2, x_3, x_4, x_5, x_6, x_11, x_8, x_9, x_10); +lean_dec_ref(x_9); +lean_dec_ref(x_2); +lean_dec(x_1); +return x_12; } } -static lean_object* _init_l_Lean_IR_EmitC_emitFnDeclAux___closed__8() { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadArg_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("static ", 7, 7); -return x_1; +uint8_t x_8; lean_object* x_9; +x_8 = lean_unbox(x_5); +x_9 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadArg_spec__0___redArg(x_1, x_2, x_3, x_4, x_8, x_6, x_7); +lean_dec_ref(x_6); +lean_dec_ref(x_2); +lean_dec(x_1); +return x_9; } } -static lean_object* _init_l_Lean_IR_EmitC_emitFnDeclAux___closed__9() { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSpreadArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("extern ", 7, 7); -return x_1; +uint8_t x_6; lean_object* x_7; +x_6 = lean_unbox(x_3); +x_7 = l_Lean_IR_EmitC_emitSpreadArg(x_1, x_2, x_6, x_4, x_5); +lean_dec_ref(x_4); +lean_dec(x_1); +return x_7; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFnDeclAux(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_EmitC_0__Lean_IR_EmitC_emitSpreadArg_match__4_splitter(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_6; lean_object* x_14; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_57; lean_object* x_58; lean_object* x_59; uint8_t x_84; -x_39 = lean_ctor_get(x_4, 0); -lean_inc_ref(x_39); -x_57 = l_Lean_IR_Decl_params(x_1); -x_84 = l_Array_isEmpty___redArg(x_57); -if (x_84 == 0) +if (lean_obj_tag(x_2) == 10) { -if (x_3 == 0) +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_5 = lean_ctor_get(x_2, 0); +x_6 = lean_ctor_get(x_2, 1); +x_7 = lean_ctor_get(x_2, 2); +x_8 = lean_ctor_get(x_2, 3); +x_9 = lean_unsigned_to_nat(0u); +x_10 = lean_nat_dec_eq(x_7, x_9); +if (x_10 == 0) { -goto block_83; +lean_object* x_11; +lean_dec(x_3); +x_11 = lean_apply_2(x_4, x_2, lean_box(0)); +return x_11; } else { -if (x_84 == 0) +uint8_t x_12; +lean_inc(x_8); +lean_inc_ref(x_6); +lean_inc(x_5); +x_12 = !lean_is_exclusive(x_2); +if (x_12 == 0) { -x_58 = x_4; -x_59 = x_5; -goto block_80; -} -else +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_13 = lean_ctor_get(x_2, 3); +lean_dec(x_13); +x_14 = lean_ctor_get(x_2, 2); +lean_dec(x_14); +x_15 = lean_ctor_get(x_2, 1); +lean_dec(x_15); +x_16 = lean_ctor_get(x_2, 0); +lean_dec(x_16); +x_17 = lean_nat_dec_eq(x_8, x_9); +if (x_17 == 0) { -goto block_83; +lean_object* x_18; +lean_dec(x_3); +lean_ctor_set(x_2, 2, x_9); +x_18 = lean_apply_2(x_4, x_2, lean_box(0)); +return x_18; } +else +{ +lean_object* x_19; +lean_free_object(x_2); +lean_dec(x_8); +lean_dec(x_4); +x_19 = lean_apply_2(x_3, x_5, x_6); +return x_19; } } else { -if (x_3 == 0) -{ -lean_object* x_85; uint8_t x_86; -x_85 = l_Lean_IR_Decl_name(x_1); -lean_inc_ref(x_39); -x_86 = l_Lean_isClosedTermName(x_39, x_85); -lean_dec(x_85); -if (x_86 == 0) +uint8_t x_20; +lean_dec(x_2); +x_20 = lean_nat_dec_eq(x_8, x_9); +if (x_20 == 0) { -lean_object* x_87; lean_object* x_88; -x_87 = l_Lean_IR_EmitC_emitFnDeclAux___closed__7; -x_88 = lean_string_append(x_5, x_87); -x_58 = x_4; -x_59 = x_88; -goto block_80; +lean_object* x_21; lean_object* x_22; +lean_dec(x_3); +x_21 = lean_alloc_ctor(10, 4, 0); +lean_ctor_set(x_21, 0, x_5); +lean_ctor_set(x_21, 1, x_6); +lean_ctor_set(x_21, 2, x_9); +lean_ctor_set(x_21, 3, x_8); +x_22 = lean_apply_2(x_4, x_21, lean_box(0)); +return x_22; } else { -lean_object* x_89; lean_object* x_90; -x_89 = l_Lean_IR_EmitC_emitFnDeclAux___closed__8; -x_90 = lean_string_append(x_5, x_89); -x_58 = x_4; -x_59 = x_90; -goto block_80; +lean_object* x_23; +lean_dec(x_8); +lean_dec(x_4); +x_23 = lean_apply_2(x_3, x_5, x_6); +return x_23; +} +} } } else { -lean_object* x_91; lean_object* x_92; -x_91 = l_Lean_IR_EmitC_emitFnDeclAux___closed__9; -x_92 = lean_string_append(x_5, x_91); -x_58 = x_4; -x_59 = x_92; -goto block_80; +lean_object* x_24; +lean_dec(x_3); +x_24 = lean_apply_2(x_4, x_2, lean_box(0)); +return x_24; } } -block_13: -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_7 = l_Lean_IR_EmitC_emitFnDeclAux___closed__0; -x_8 = lean_string_append(x_6, x_7); -x_9 = l_Lean_IR_EmitC_emitLn___closed__0; -x_10 = lean_box(0); -x_11 = lean_string_append(x_8, x_9); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_10); -lean_ctor_set(x_12, 1, x_11); -return x_12; } -block_17: +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_EmitC_0__Lean_IR_EmitC_emitSpreadArg_match__4_splitter___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_15; lean_object* x_16; -x_15 = l_Lean_IR_EmitC_emitFnDeclAux___closed__1; -x_16 = lean_string_append(x_14, x_15); -x_6 = x_16; -goto block_13; -} -block_28: +if (lean_obj_tag(x_1) == 10) { -lean_dec_ref(x_18); -if (x_23 == 0) +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_4 = lean_ctor_get(x_1, 0); +x_5 = lean_ctor_get(x_1, 1); +x_6 = lean_ctor_get(x_1, 2); +x_7 = lean_ctor_get(x_1, 3); +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_nat_dec_eq(x_6, x_8); +if (x_9 == 0) { -lean_object* x_24; lean_object* x_25; -lean_inc(x_22); -x_24 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg(x_20, x_21, x_22, x_22, x_19); -lean_dec(x_22); -lean_dec_ref(x_20); -x_25 = lean_ctor_get(x_24, 1); -lean_inc(x_25); -lean_dec_ref(x_24); -x_14 = x_25; -goto block_17; +lean_object* x_10; +lean_dec(x_2); +x_10 = lean_apply_2(x_3, x_1, lean_box(0)); +return x_10; } else { -lean_object* x_26; lean_object* x_27; -lean_dec(x_22); -lean_dec_ref(x_20); -x_26 = l_Lean_IR_EmitC_emitFnDeclAux___closed__2; -x_27 = lean_string_append(x_19, x_26); -x_14 = x_27; -goto block_17; -} -} -block_38: +uint8_t x_11; +lean_inc(x_7); +lean_inc_ref(x_5); +lean_inc(x_4); +x_11 = !lean_is_exclusive(x_1); +if (x_11 == 0) { -lean_object* x_34; lean_object* x_35; uint8_t x_36; -x_34 = l_Lean_IR_EmitC_emitFnDeclAux___closed__3; -x_35 = lean_array_get_size(x_33); -x_36 = lean_nat_dec_lt(x_34, x_35); -if (x_36 == 0) +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_12 = lean_ctor_get(x_1, 3); +lean_dec(x_12); +x_13 = lean_ctor_get(x_1, 2); +lean_dec(x_13); +x_14 = lean_ctor_get(x_1, 1); +lean_dec(x_14); +x_15 = lean_ctor_get(x_1, 0); +lean_dec(x_15); +x_16 = lean_nat_dec_eq(x_7, x_8); +if (x_16 == 0) { -lean_dec(x_31); -x_18 = x_29; -x_19 = x_30; -x_20 = x_33; -x_21 = x_32; -x_22 = x_35; -x_23 = x_36; -goto block_28; +lean_object* x_17; +lean_dec(x_2); +lean_ctor_set(x_1, 2, x_8); +x_17 = lean_apply_2(x_3, x_1, lean_box(0)); +return x_17; } else { -uint8_t x_37; -x_37 = l_Lean_IR_ExplicitBoxing_isBoxedName(x_31); -lean_dec(x_31); -x_18 = x_29; -x_19 = x_30; -x_20 = x_33; -x_21 = x_32; -x_22 = x_35; -x_23 = x_37; -goto block_28; +lean_object* x_18; +lean_free_object(x_1); +lean_dec(x_7); +lean_dec(x_3); +x_18 = lean_apply_2(x_2, x_4, x_5); +return x_18; } } -block_56: +else { -lean_object* x_44; uint8_t x_45; -x_44 = l_Lean_IR_Decl_name(x_1); -lean_inc(x_44); -x_45 = l_Lean_isExternC(x_39, x_44); -if (x_45 == 0) +uint8_t x_19; +lean_dec(x_1); +x_19 = lean_nat_dec_eq(x_7, x_8); +if (x_19 == 0) { -x_29 = x_40; -x_30 = x_41; -x_31 = x_44; -x_32 = x_42; -x_33 = x_43; -goto block_38; +lean_object* x_20; lean_object* x_21; +lean_dec(x_2); +x_20 = lean_alloc_ctor(10, 4, 0); +lean_ctor_set(x_20, 0, x_4); +lean_ctor_set(x_20, 1, x_5); +lean_ctor_set(x_20, 2, x_8); +lean_ctor_set(x_20, 3, x_7); +x_21 = lean_apply_2(x_3, x_20, lean_box(0)); +return x_21; } else { -lean_object* x_46; lean_object* x_47; uint8_t x_48; -x_46 = lean_array_get_size(x_43); -x_47 = lean_mk_empty_array_with_capacity(x_42); -x_48 = lean_nat_dec_lt(x_42, x_46); -if (x_48 == 0) -{ -lean_dec_ref(x_43); -x_29 = x_40; -x_30 = x_41; -x_31 = x_44; -x_32 = x_42; -x_33 = x_47; -goto block_38; +lean_object* x_22; +lean_dec(x_7); +lean_dec(x_3); +x_22 = lean_apply_2(x_2, x_4, x_5); +return x_22; +} +} +} } else { -uint8_t x_49; -x_49 = lean_nat_dec_le(x_46, x_46); -if (x_49 == 0) +lean_object* x_23; +lean_dec(x_2); +x_23 = lean_apply_2(x_3, x_1, lean_box(0)); +return x_23; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_EmitC_0__Lean_IR_EmitC_emitSpreadArg_match__10_splitter(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -if (x_48 == 0) +if (lean_obj_tag(x_2) == 1) { -lean_dec_ref(x_43); -x_29 = x_40; -x_30 = x_41; -x_31 = x_44; -x_32 = x_42; -x_33 = x_47; -goto block_38; +lean_object* x_5; lean_object* x_6; +lean_dec(x_4); +x_5 = lean_ctor_get(x_2, 0); +lean_inc(x_5); +lean_dec_ref(x_2); +x_6 = lean_apply_1(x_3, x_5); +return x_6; } else { -size_t x_50; size_t x_51; lean_object* x_52; -x_50 = 0; -x_51 = lean_usize_of_nat(x_46); -x_52 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__1(x_43, x_50, x_51, x_47); -lean_dec_ref(x_43); -x_29 = x_40; -x_30 = x_41; -x_31 = x_44; -x_32 = x_42; -x_33 = x_52; -goto block_38; +lean_object* x_7; +lean_dec(x_3); +x_7 = lean_apply_2(x_4, x_2, lean_box(0)); +return x_7; } } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_EmitC_0__Lean_IR_EmitC_emitSpreadArg_match__10_splitter___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -size_t x_53; size_t x_54; lean_object* x_55; -x_53 = 0; -x_54 = lean_usize_of_nat(x_46); -x_55 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__1(x_43, x_53, x_54, x_47); -lean_dec_ref(x_43); -x_29 = x_40; -x_30 = x_41; -x_31 = x_44; -x_32 = x_42; -x_33 = x_55; -goto block_38; +if (lean_obj_tag(x_1) == 1) +{ +lean_object* x_4; lean_object* x_5; +lean_dec(x_3); +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +lean_dec_ref(x_1); +x_5 = lean_apply_1(x_2, x_4); +return x_5; } +else +{ +lean_object* x_6; +lean_dec(x_2); +x_6 = lean_apply_2(x_3, x_1, lean_box(0)); +return x_6; } } } -block_80: -{ -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; -x_60 = l_Lean_IR_Decl_resultType(x_1); -x_61 = l_Lean_IR_EmitC_toCType(x_60); -lean_dec(x_60); -x_62 = l_Lean_IR_EmitC_emitFnDeclAux___closed__4; -x_63 = lean_string_append(x_61, x_62); -x_64 = lean_string_append(x_63, x_2); -x_65 = lean_string_append(x_59, x_64); -lean_dec_ref(x_64); -x_66 = l_Array_isEmpty___redArg(x_57); -if (x_66 == 0) +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreadArgs_spec__0(uint8_t x_1, lean_object* x_2, size_t x_3, size_t x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7) { +_start: { -lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; -x_67 = l_Lean_IR_EmitC_emitFnDeclAux___closed__5; -x_68 = lean_string_append(x_65, x_67); -x_69 = lean_unsigned_to_nat(0u); -x_70 = lean_array_get_size(x_57); -x_71 = l_Lean_IR_EmitC_emitFnDeclAux___closed__6; -x_72 = lean_nat_dec_lt(x_69, x_70); -if (x_72 == 0) +uint8_t x_8; +x_8 = lean_usize_dec_lt(x_4, x_3); +if (x_8 == 0) { -lean_dec_ref(x_57); -x_40 = x_58; -x_41 = x_68; -x_42 = x_69; -x_43 = x_71; -goto block_56; +lean_object* x_9; lean_object* x_10; +x_9 = lean_box(x_5); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_7); +return x_10; } else { -uint8_t x_73; -x_73 = lean_nat_dec_le(x_70, x_70); -if (x_73 == 0) -{ -if (x_72 == 0) +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = lean_array_uget(x_2, x_4); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +if (x_1 == 0) { -lean_dec_ref(x_57); -x_40 = x_58; -x_41 = x_68; -x_42 = x_69; -x_43 = x_71; -goto block_56; +lean_object* x_23; +lean_dec(x_12); +x_23 = lean_box(0); +x_14 = x_23; +goto block_22; } else { -size_t x_74; size_t x_75; lean_object* x_76; -x_74 = 0; -x_75 = lean_usize_of_nat(x_70); -x_76 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__2(x_57, x_74, x_75, x_71); -lean_dec_ref(x_57); -x_40 = x_58; -x_41 = x_68; -x_42 = x_69; -x_43 = x_76; -goto block_56; -} +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_24 = l_Lean_IR_EmitC_argToCString___closed__0; +x_25 = l_Nat_reprFast(x_12); +x_26 = lean_string_append(x_24, x_25); +lean_dec_ref(x_25); +x_27 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_27, 0, x_26); +x_14 = x_27; +goto block_22; } -else +block_22: { -size_t x_77; size_t x_78; lean_object* x_79; -x_77 = 0; -x_78 = lean_usize_of_nat(x_70); -x_79 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__2(x_57, x_77, x_78, x_71); -lean_dec_ref(x_57); -x_40 = x_58; -x_41 = x_68; -x_42 = x_69; -x_43 = x_79; -goto block_56; -} -} +lean_object* x_15; +x_15 = l_Lean_IR_EmitC_emitSpreadArg(x_13, x_14, x_5, x_6, x_7); +lean_dec(x_13); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; uint8_t x_20; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec_ref(x_15); +x_18 = 1; +x_19 = lean_usize_add(x_4, x_18); +x_20 = lean_unbox(x_16); +lean_dec(x_16); +x_4 = x_19; +x_5 = x_20; +x_7 = x_17; +goto _start; } else { -lean_dec_ref(x_58); -lean_dec_ref(x_57); -lean_dec_ref(x_39); -x_6 = x_65; -goto block_13; +return x_15; } } -block_83: -{ -lean_object* x_81; lean_object* x_82; -x_81 = l_Lean_IR_EmitC_emitFnDeclAux___closed__7; -x_82 = lean_string_append(x_5, x_81); -x_58 = x_4; -x_59 = x_82; -goto block_80; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___boxed(lean_object* x_1, lean_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_IR_EmitC_emitSpreadArgs(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_8; -x_8 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec_ref(x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec_ref(x_1); +uint8_t x_5; size_t x_6; size_t x_7; lean_object* x_8; +x_5 = 1; +x_6 = lean_array_size(x_1); +x_7 = 0; +x_8 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreadArgs_spec__0(x_2, x_1, x_6, x_7, x_5, x_3, x_4); +if (lean_obj_tag(x_8) == 0) +{ +uint8_t x_9; +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_8, 0); +lean_dec(x_10); +x_11 = lean_box(0); +lean_ctor_set(x_8, 0, x_11); return x_8; } -} -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +else { -size_t x_5; size_t x_6; lean_object* x_7; -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__1(x_1, x_5, x_6, x_4); -lean_dec_ref(x_1); -return x_7; +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_8, 1); +lean_inc(x_12); +lean_dec(x_8); +x_13 = lean_box(0); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_12); +return x_14; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +else { -size_t x_5; size_t x_6; lean_object* x_7; -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__2(x_1, x_5, x_6, x_4); -lean_dec_ref(x_1); -return x_7; +uint8_t x_15; +x_15 = !lean_is_exclusive(x_8); +if (x_15 == 0) +{ +return x_8; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_8, 0); +x_17 = lean_ctor_get(x_8, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_8); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___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_IR_EmitC_emitSpreadArgs___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_6; -x_6 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_3); -lean_dec(x_2); +uint8_t x_5; lean_object* x_6; +x_5 = lean_unbox(x_2); +x_6 = l_Lean_IR_EmitC_emitSpreadArgs(x_1, x_5, x_3, x_4); +lean_dec_ref(x_3); lean_dec_ref(x_1); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFnDeclAux___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreadArgs_spec__0___boxed(lean_object* x_1, lean_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_6; lean_object* x_7; -x_6 = lean_unbox(x_3); -x_7 = l_Lean_IR_EmitC_emitFnDeclAux(x_1, x_2, x_6, x_4, x_5); +uint8_t x_8; size_t x_9; size_t x_10; uint8_t x_11; lean_object* x_12; +x_8 = lean_unbox(x_1); +x_9 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_10 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_11 = lean_unbox(x_5); +x_12 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreadArgs_spec__0(x_8, x_2, x_9, x_10, x_11, x_6, x_7); +lean_dec_ref(x_6); lean_dec_ref(x_2); -lean_dec_ref(x_1); -return x_7; +return x_12; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFnDecl(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__0() { _start: { -lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_IR_Decl_name(x_1); -lean_inc_ref(x_3); -x_6 = l_Lean_IR_EmitC_toCName(x_5, x_3, x_4); -if (lean_obj_tag(x_6) == 0) -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_7 = lean_ctor_get(x_6, 0); -lean_inc(x_7); -x_8 = lean_ctor_get(x_6, 1); -lean_inc(x_8); -lean_dec_ref(x_6); -x_9 = l_Lean_IR_EmitC_emitFnDeclAux(x_1, x_7, x_2, x_3, x_8); -lean_dec(x_7); -return x_9; +lean_object* x_1; +x_1 = lean_mk_string_unchecked(".i", 2, 2); +return x_1; } -else -{ -uint8_t x_10; -lean_dec_ref(x_3); -x_10 = !lean_is_exclusive(x_6); -if (x_10 == 0) -{ -return x_6; } -else +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__1() { +_start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_11 = lean_ctor_get(x_6, 0); -x_12 = lean_ctor_get(x_6, 1); -lean_inc(x_12); -lean_inc(x_11); -lean_dec(x_6); -x_13 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_13, 0, x_11); -lean_ctor_set(x_13, 1, x_12); -return x_13; -} -} +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" = ", 3, 3); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFnDecl___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Lean_IR_EmitC_emitSpreadValue___closed__0() { _start: { -uint8_t x_5; lean_object* x_6; -x_5 = lean_unbox(x_2); -x_6 = l_Lean_IR_EmitC_emitFnDecl(x_1, x_5, x_3, x_4); -lean_dec_ref(x_1); -return x_6; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("{", 1, 1); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitExternDeclAux(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Lean_IR_EmitC_emitSpreadValue___closed__1() { _start: { -lean_object* x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; -x_5 = lean_ctor_get(x_3, 0); -x_6 = l_Lean_IR_Decl_name(x_1); -lean_inc_ref(x_5); -x_7 = l_Lean_isExternC(x_5, x_6); -x_8 = l_Lean_IR_EmitC_emitFnDeclAux(x_1, x_2, x_7, x_3, x_4); -return x_8; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("}", 1, 1); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitExternDeclAux___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSpreadValue(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_IR_EmitC_emitExternDeclAux(x_1, x_2, x_3, x_4); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -return x_5; -} -} -LEAN_EXPORT lean_object* l_List_foldl___at___00Lean_IR_EmitC_emitFnDecls_spec__0(lean_object* x_1, lean_object* x_2) { -_start: +if (lean_obj_tag(x_1) == 10) { -if (lean_obj_tag(x_2) == 0) +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +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_unsigned_to_nat(0u); +x_14 = lean_nat_dec_eq(x_11, x_13); +if (x_14 == 0) { -return x_1; +x_5 = x_4; +goto block_9; } else { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_3 = lean_ctor_get(x_2, 0); -x_4 = lean_ctor_get(x_2, 1); -x_5 = l_Lean_IR_Decl_name(x_3); -x_6 = l_Lean_NameSet_insert(x_1, x_5); -x_1 = x_6; -x_2 = x_4; -goto _start; -} -} +uint8_t x_15; +x_15 = lean_nat_dec_eq(x_12, x_13); +if (x_15 == 0) +{ +x_5 = x_4; +goto block_9; } -LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_foldrM___at___00Lean_IR_EmitC_emitFnDecls_spec__2(lean_object* x_1, lean_object* x_2) { -_start: +else { -if (lean_obj_tag(x_2) == 0) +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_array_get_size(x_10); +x_17 = l_Lean_IR_EmitC_emitSpreadValue___closed__0; +x_18 = lean_string_append(x_4, x_17); +x_19 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg(x_16, x_10, x_2, x_13, x_15, x_3, x_18); +if (lean_obj_tag(x_19) == 0) { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_3 = lean_ctor_get(x_2, 1); -x_4 = lean_ctor_get(x_2, 3); -x_5 = lean_ctor_get(x_2, 4); -x_6 = l_Std_DTreeMap_Internal_Impl_foldrM___at___00Lean_IR_EmitC_emitFnDecls_spec__2(x_1, x_5); -lean_inc(x_3); -x_7 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_7, 0, x_3); -lean_ctor_set(x_7, 1, x_6); -x_1 = x_7; -x_2 = x_4; -goto _start; +uint8_t x_20; +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_21 = lean_ctor_get(x_19, 1); +x_22 = lean_ctor_get(x_19, 0); +lean_dec(x_22); +x_23 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; +x_24 = lean_string_append(x_21, x_23); +x_25 = lean_box(0); +lean_ctor_set(x_19, 1, x_24); +lean_ctor_set(x_19, 0, x_25); +return x_19; } else { -return x_1; -} +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_26 = lean_ctor_get(x_19, 1); +lean_inc(x_26); +lean_dec(x_19); +x_27 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; +x_28 = lean_string_append(x_26, x_27); +x_29 = lean_box(0); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_28); +return x_30; } } -LEAN_EXPORT lean_object* l_List_foldl___at___00Lean_IR_EmitC_emitFnDecls_spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -if (lean_obj_tag(x_3) == 0) +uint8_t x_31; +x_31 = !lean_is_exclusive(x_19); +if (x_31 == 0) { -lean_dec_ref(x_1); -return x_2; +return x_19; } else { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_4 = lean_ctor_get(x_3, 0); -lean_inc(x_4); -x_5 = lean_ctor_get(x_3, 1); -lean_inc(x_5); -lean_dec_ref(x_3); -x_6 = l_Lean_IR_Decl_name(x_4); -x_7 = l_Lean_NameSet_insert(x_2, x_6); -lean_inc_ref(x_1); -x_8 = l_Lean_IR_collectUsedDecls(x_1, x_4, x_7); -x_2 = x_8; -x_3 = x_5; -goto _start; +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_19, 0); +x_33 = lean_ctor_get(x_19, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_19); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; } } } -static lean_object* _init_l_Lean_IR_EmitC_emitFnDecls___closed__0() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_NameSet_empty; -return x_1; } } -static lean_object* _init_l_List_forM___at___00Lean_IR_EmitC_emitFnDecls_spec__3___closed__0() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("c", 1, 1); -return x_1; -} +x_5 = x_4; +goto block_9; } -static lean_object* _init_l_List_forM___at___00Lean_IR_EmitC_emitFnDecls_spec__3___closed__1() { -_start: +block_9: { -lean_object* x_1; lean_object* x_2; -x_1 = l_List_forM___at___00Lean_IR_EmitC_emitFnDecls_spec__3___closed__0; -x_2 = l_Lean_Name_mkStr1(x_1); -return x_2; +lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_6 = lean_box(0); +x_7 = lean_string_append(x_5, x_2); +lean_dec_ref(x_2); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_6); +lean_ctor_set(x_8, 1, x_7); +return x_8; } } -LEAN_EXPORT lean_object* l_List_forM___at___00Lean_IR_EmitC_emitFnDecls_spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -if (lean_obj_tag(x_3) == 0) +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_9 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__0; +x_10 = lean_string_append(x_8, x_9); +x_11 = l_Nat_reprFast(x_1); +x_12 = lean_string_append(x_10, x_11); +x_13 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__1; +x_14 = lean_string_append(x_12, x_13); +x_15 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadArg_spec__0___redArg___closed__0; +x_16 = lean_string_append(x_2, x_15); +x_17 = lean_string_append(x_16, x_11); +lean_dec_ref(x_11); +x_18 = l_Lean_IR_EmitC_emitSpreadValue(x_3, x_17, x_7, x_14); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_6; lean_object* x_7; -lean_dec_ref(x_4); -lean_dec_ref(x_1); -x_6 = lean_box(0); -x_7 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_7, 0, x_6); -lean_ctor_set(x_7, 1, x_5); -return x_7; +uint8_t x_19; +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_18, 0); +lean_dec(x_20); +x_21 = lean_box(x_4); +x_22 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_18, 0, x_22); +return x_18; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_14; -x_8 = lean_ctor_get(x_3, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_3, 1); -lean_inc(x_9); -lean_dec_ref(x_3); -lean_inc_ref(x_4); -lean_inc(x_8); -x_14 = l_Lean_IR_EmitC_getDecl(x_8, x_4, x_5); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); -lean_inc(x_16); -lean_dec_ref(x_14); -x_17 = l_List_forM___at___00Lean_IR_EmitC_emitFnDecls_spec__3___closed__1; -x_18 = l_Lean_IR_Decl_name(x_15); -lean_inc_ref(x_1); -x_19 = l_Lean_getExternNameFor(x_1, x_17, x_18); -if (lean_obj_tag(x_19) == 0) -{ -uint8_t x_20; -x_20 = l_Lean_NameSet_contains(x_2, x_8); -lean_dec(x_8); -if (x_20 == 0) -{ -uint8_t x_21; lean_object* x_22; -x_21 = 1; -lean_inc_ref(x_4); -x_22 = l_Lean_IR_EmitC_emitFnDecl(x_15, x_21, x_4, x_16); -lean_dec(x_15); -x_10 = x_22; -goto block_13; -} -else -{ -uint8_t x_23; lean_object* x_24; -x_23 = 0; -lean_inc_ref(x_4); -x_24 = l_Lean_IR_EmitC_emitFnDecl(x_15, x_23, x_4, x_16); -lean_dec(x_15); -x_10 = x_24; -goto block_13; -} -} -else -{ -lean_object* x_25; lean_object* x_26; -lean_dec(x_8); -x_25 = lean_ctor_get(x_19, 0); -lean_inc(x_25); -lean_dec_ref(x_19); -lean_inc_ref(x_4); -x_26 = l_Lean_IR_EmitC_emitExternDeclAux(x_15, x_25, x_4, x_16); -lean_dec(x_25); -lean_dec(x_15); -x_10 = x_26; -goto block_13; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_23 = lean_ctor_get(x_18, 1); +lean_inc(x_23); +lean_dec(x_18); +x_24 = lean_box(x_4); +x_25 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_25, 0, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_23); +return x_26; } } else { uint8_t x_27; -lean_dec(x_9); -lean_dec(x_8); -lean_dec_ref(x_4); -lean_dec_ref(x_1); -x_27 = !lean_is_exclusive(x_14); +x_27 = !lean_is_exclusive(x_18); if (x_27 == 0) { -return x_14; +return x_18; } else { lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_14, 0); -x_29 = lean_ctor_get(x_14, 1); +x_28 = lean_ctor_get(x_18, 0); +x_29 = lean_ctor_get(x_18, 1); lean_inc(x_29); lean_inc(x_28); -lean_dec(x_14); +lean_dec(x_18); x_30 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_30, 0, x_28); lean_ctor_set(x_30, 1, x_29); return x_30; } } -block_13: +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7) { +_start: { -if (lean_obj_tag(x_10) == 0) +uint8_t x_8; lean_object* x_9; lean_object* x_14; uint8_t x_30; +x_30 = lean_nat_dec_lt(x_4, x_1); +if (x_30 == 0) { -lean_object* x_11; -x_11 = lean_ctor_get(x_10, 1); -lean_inc(x_11); -lean_dec_ref(x_10); -x_3 = x_9; -x_5 = x_11; -goto _start; +lean_object* x_31; lean_object* x_32; +lean_dec(x_4); +lean_dec_ref(x_3); +x_31 = lean_box(x_5); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_7); +return x_32; } else { -lean_dec(x_9); -lean_dec_ref(x_4); -lean_dec_ref(x_1); -return x_10; -} -} -} +lean_object* x_33; +x_33 = lean_array_fget_borrowed(x_2, x_4); +switch (lean_obj_tag(x_33)) { +case 6: +{ +x_8 = x_5; +x_9 = x_7; +goto block_13; } +case 13: +{ +x_8 = x_5; +x_9 = x_7; +goto block_13; } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFnDecls(lean_object* x_1, lean_object* x_2) { -_start: +default: { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_3 = lean_ctor_get(x_1, 0); -lean_inc_ref(x_3); +uint8_t x_34; +x_34 = 0; +if (x_5 == 0) +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_35 = l_Lean_IR_EmitC_emitSpreadArg___closed__1; +x_36 = lean_box(0); +x_37 = lean_string_append(x_7, x_35); lean_inc_ref(x_3); -x_4 = l_Lean_IR_getDecls(x_3); -x_5 = l_Lean_IR_EmitC_emitFnDecls___closed__0; -x_6 = l_List_foldl___at___00Lean_IR_EmitC_emitFnDecls_spec__0(x_5, x_4); +lean_inc(x_4); +x_38 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0(x_4, x_3, x_33, x_34, x_5, x_36, x_6, x_37); +x_14 = x_38; +goto block_29; +} +else +{ +lean_object* x_39; lean_object* x_40; +x_39 = lean_box(0); lean_inc_ref(x_3); -x_7 = l_List_foldl___at___00Lean_IR_EmitC_emitFnDecls_spec__1(x_3, x_5, x_4); -x_8 = lean_box(0); -x_9 = l_Std_DTreeMap_Internal_Impl_foldrM___at___00Lean_IR_EmitC_emitFnDecls_spec__2(x_8, x_7); -lean_dec(x_7); -x_10 = l_List_forM___at___00Lean_IR_EmitC_emitFnDecls_spec__3(x_3, x_6, x_9, x_1, x_2); -lean_dec(x_6); -return x_10; +lean_inc(x_4); +x_40 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0(x_4, x_3, x_33, x_34, x_5, x_39, x_6, x_7); +x_14 = x_40; +goto block_29; } } -LEAN_EXPORT lean_object* l_List_foldl___at___00Lean_IR_EmitC_emitFnDecls_spec__0___boxed(lean_object* x_1, lean_object* x_2) { -_start: +} +} +block_13: { -lean_object* x_3; -x_3 = l_List_foldl___at___00Lean_IR_EmitC_emitFnDecls_spec__0(x_1, x_2); -lean_dec(x_2); -return x_3; +lean_object* x_10; lean_object* x_11; +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_add(x_4, x_10); +lean_dec(x_4); +x_4 = x_11; +x_5 = x_8; +x_7 = x_9; +goto _start; } +block_29: +{ +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +if (lean_obj_tag(x_15) == 0) +{ +uint8_t x_16; +lean_dec(x_4); +lean_dec_ref(x_3); +x_16 = !lean_is_exclusive(x_14); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_14, 0); +lean_dec(x_17); +x_18 = lean_ctor_get(x_15, 0); +lean_inc(x_18); +lean_dec_ref(x_15); +lean_ctor_set(x_14, 0, x_18); +return x_14; } -LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_foldrM___at___00Lean_IR_EmitC_emitFnDecls_spec__2___boxed(lean_object* x_1, lean_object* x_2) { -_start: +else { -lean_object* x_3; -x_3 = l_Std_DTreeMap_Internal_Impl_foldrM___at___00Lean_IR_EmitC_emitFnDecls_spec__2(x_1, x_2); -lean_dec(x_2); -return x_3; +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_14, 1); +lean_inc(x_19); +lean_dec(x_14); +x_20 = lean_ctor_get(x_15, 0); +lean_inc(x_20); +lean_dec_ref(x_15); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +return x_21; } } -LEAN_EXPORT lean_object* l_List_forM___at___00Lean_IR_EmitC_emitFnDecls_spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +else { -lean_object* x_6; -x_6 = l_List_forM___at___00Lean_IR_EmitC_emitFnDecls_spec__3(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_2); -return x_6; +lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_dec_ref(x_14); +x_23 = lean_ctor_get(x_15, 0); +lean_inc(x_23); +lean_dec_ref(x_15); +x_24 = lean_unbox(x_23); +lean_dec(x_23); +x_8 = x_24; +x_9 = x_22; +goto block_13; } } -LEAN_EXPORT lean_object* l_List_forM___at___00Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0_spec__0___redArg(lean_object* x_1, lean_object* x_2) { -_start: +else { -if (lean_obj_tag(x_1) == 0) +uint8_t x_25; +lean_dec(x_4); +lean_dec_ref(x_3); +x_25 = !lean_is_exclusive(x_14); +if (x_25 == 0) { -lean_object* x_3; lean_object* x_4; -x_3 = lean_box(0); -x_4 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_4, 0, x_3); -lean_ctor_set(x_4, 1, x_2); -return x_4; +return x_14; } else { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_5 = lean_ctor_get(x_1, 0); -x_6 = lean_ctor_get(x_1, 1); -x_7 = lean_string_append(x_2, x_5); -x_8 = l_Lean_IR_EmitC_emitLn___closed__0; -x_9 = lean_string_append(x_7, x_8); -x_1 = x_6; -x_2 = x_9; -goto _start; +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_14, 0); +x_27 = lean_ctor_get(x_14, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_14); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_4; -x_4 = l_List_forM___at___00Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0_spec__0___redArg(x_1, x_3); -return x_4; +lean_object* x_11; +x_11 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg(x_1, x_2, x_3, x_6, x_7, x_9, x_10); +return x_11; } } -LEAN_EXPORT lean_object* l_List_forM___at___00Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_4; -x_4 = l_List_forM___at___00Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0_spec__0___redArg(x_1, x_3); -return x_4; +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_7); +x_12 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0(x_1, x_2, x_3, x_4, x_5, x_6, x_11, x_8, x_9, x_10); +lean_dec_ref(x_9); +lean_dec_ref(x_2); +lean_dec(x_1); +return x_12; } } -static lean_object* _init_l_panic___at___00Lean_IR_EmitC_emitMainFn_spec__2___closed__0() { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_1; -x_1 = l_Lean_instInhabitedConstantInfo_default; -return x_1; +uint8_t x_9; uint8_t x_10; lean_object* x_11; +x_9 = lean_unbox(x_4); +x_10 = lean_unbox(x_5); +x_11 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0(x_1, x_2, x_3, x_9, x_10, x_6, x_7, x_8); +lean_dec_ref(x_7); +lean_dec(x_3); +return x_11; } } -LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_EmitC_emitMainFn_spec__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSpreadValue___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_2; lean_object* x_3; -x_2 = l_panic___at___00Lean_IR_EmitC_emitMainFn_spec__2___closed__0; -x_3 = lean_panic_fn(x_2, x_1); -return x_3; +lean_object* x_5; +x_5 = l_Lean_IR_EmitC_emitSpreadValue(x_1, x_2, x_3, x_4); +lean_dec_ref(x_3); +lean_dec(x_1); +return x_5; } } -LEAN_EXPORT uint8_t l_Option_instBEq_beq___at___00Lean_IR_EmitC_emitMainFn_spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -if (lean_obj_tag(x_1) == 0) -{ -if (lean_obj_tag(x_2) == 0) -{ -uint8_t x_3; -x_3 = 1; -return x_3; -} -else -{ -uint8_t x_4; -x_4 = 0; -return x_4; -} -} -else -{ -if (lean_obj_tag(x_2) == 0) -{ -uint8_t x_5; -x_5 = 0; -return x_5; -} -else -{ -lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_6 = lean_ctor_get(x_1, 0); -x_7 = lean_ctor_get(x_2, 0); -x_8 = lean_name_eq(x_6, x_7); -return x_8; -} -} -} -} -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__0() { -_start: +uint8_t x_8; lean_object* x_9; +x_8 = lean_unbox(x_5); +x_9 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg(x_1, x_2, x_3, x_4, x_8, x_6, x_7); +lean_dec_ref(x_6); +lean_dec_ref(x_2); +lean_dec(x_1); +return x_9; +} +} +static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__0() { +_start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked(" lean_dec_ref(res);", 20, 20); +x_1 = lean_mk_string_unchecked(" = (", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__1() { +static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked(" return ret;", 13, 13); +x_1 = lean_mk_string_unchecked(")", 1, 1); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__2() { +static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("} else {", 8, 8); +x_1 = lean_mk_string_unchecked(";", 1, 1); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__3() { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked(" lean_io_result_show_error(res);", 33, 33); -return x_1; +lean_object* x_8; lean_object* x_9; uint8_t x_14; +x_14 = lean_usize_dec_lt(x_4, x_3); +if (x_14 == 0) +{ +lean_object* x_15; +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_5); +lean_ctor_set(x_15, 1, x_7); +return x_15; } +else +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_array_uget(x_2, x_4); +x_17 = lean_ctor_get(x_16, 1); +lean_inc(x_17); +if (lean_obj_tag(x_17) == 10) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +lean_dec(x_16); +x_19 = lean_ctor_get(x_17, 2); +x_20 = lean_ctor_get(x_17, 3); +x_21 = lean_unsigned_to_nat(0u); +x_22 = lean_nat_dec_eq(x_19, x_21); +if (x_22 == 0) +{ +lean_dec(x_18); +lean_dec_ref(x_17); +x_8 = x_1; +x_9 = x_7; +goto block_13; } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__4() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked(" return 1;", 11, 11); -return x_1; +uint8_t x_23; +x_23 = lean_nat_dec_eq(x_20, x_21); +if (x_23 == 0) +{ +lean_dec(x_18); +lean_dec_ref(x_17); +x_8 = x_1; +x_9 = x_7; +goto block_13; } +else +{ +lean_object* x_24; +x_24 = l_Lean_IR_EmitC_toCType(x_17, x_6, x_7); +if (lean_obj_tag(x_24) == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec_ref(x_24); +x_27 = lean_string_append(x_26, x_25); +lean_dec(x_25); +x_28 = l_Lean_IR_EmitC_emitSpreadArg___closed__0; +x_29 = lean_string_append(x_27, x_28); +x_30 = l_Lean_IR_EmitC_argToCString___closed__0; +x_31 = l_Nat_reprFast(x_18); +x_32 = lean_string_append(x_30, x_31); +lean_dec_ref(x_31); +x_33 = lean_string_append(x_29, x_32); +x_34 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__0; +x_35 = lean_string_append(x_33, x_34); +x_36 = l_Lean_IR_EmitC_toCType(x_17, x_6, x_35); +if (lean_obj_tag(x_36) == 0) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +lean_dec_ref(x_36); +x_39 = lean_string_append(x_38, x_37); +lean_dec(x_37); +x_40 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__1; +x_41 = lean_string_append(x_39, x_40); +x_42 = l_Lean_IR_EmitC_emitSpreadValue(x_17, x_32, x_6, x_41); +lean_dec_ref(x_17); +if (lean_obj_tag(x_42) == 0) +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_43 = lean_ctor_get(x_42, 1); +lean_inc(x_43); +lean_dec_ref(x_42); +x_44 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_45 = lean_string_append(x_43, x_44); +x_46 = l_Lean_IR_EmitC_emitLn___closed__0; +x_47 = lean_string_append(x_45, x_46); +x_8 = x_1; +x_9 = x_47; +goto block_13; } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__5() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_finalize_task_manager();", 29, 29); -return x_1; +return x_42; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__6() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked(" int ret = ", 12, 12); -return x_1; -} +uint8_t x_48; +lean_dec_ref(x_32); +lean_dec_ref(x_17); +x_48 = !lean_is_exclusive(x_36); +if (x_48 == 0) +{ +return x_36; } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__7() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("UInt32", 6, 6); -return x_1; +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_36, 0); +x_50 = lean_ctor_get(x_36, 1); +lean_inc(x_50); +lean_inc(x_49); +lean_dec(x_36); +x_51 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +return x_51; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_IR_EmitC_emitMainFn___closed__7; -x_2 = l_Lean_Name_mkStr1(x_1); -return x_2; } +else +{ +uint8_t x_52; +lean_dec(x_18); +lean_dec_ref(x_17); +x_52 = !lean_is_exclusive(x_24); +if (x_52 == 0) +{ +return x_24; } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__9() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_IR_EmitC_emitMainFn___closed__8; -x_2 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_24, 0); +x_54 = lean_ctor_get(x_24, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_24); +x_55 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +return x_55; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__10() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("0;", 2, 2); -return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__11() { -_start: +} +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_unbox_uint32(lean_io_result_get_value(res));", 49, 49); -return x_1; +lean_dec(x_17); +lean_dec(x_16); +x_8 = x_1; +x_9 = x_7; +goto block_13; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__12() { -_start: +block_13: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("}", 1, 1); -return x_1; +size_t x_10; size_t x_11; +x_10 = 1; +x_11 = lean_usize_add(x_4, x_10); +x_4 = x_11; +x_5 = x_8; +x_7 = x_9; +goto _start; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__13() { +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSpreads(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Init.Data.Option.BasicAux", 25, 25); -return x_1; -} +lean_object* x_4; size_t x_5; size_t x_6; lean_object* x_7; +x_4 = lean_box(0); +x_5 = lean_array_size(x_1); +x_6 = 0; +x_7 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0(x_4, x_1, x_5, x_6, x_4, x_2, x_3); +if (lean_obj_tag(x_7) == 0) +{ +uint8_t x_8; +x_8 = !lean_is_exclusive(x_7); +if (x_8 == 0) +{ +lean_object* x_9; +x_9 = lean_ctor_get(x_7, 0); +lean_dec(x_9); +lean_ctor_set(x_7, 0, x_4); +return x_7; } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__14() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Option.get!", 11, 11); -return x_1; +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_7, 1); +lean_inc(x_10); +lean_dec(x_7); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_4); +lean_ctor_set(x_11, 1, x_10); +return x_11; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__15() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("value is none", 13, 13); -return x_1; +return x_7; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__16() { +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSpreads___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Lean_IR_EmitC_emitMainFn___closed__15; -x_2 = lean_unsigned_to_nat(14u); -x_3 = lean_unsigned_to_nat(22u); -x_4 = l_Lean_IR_EmitC_emitMainFn___closed__14; -x_5 = l_Lean_IR_EmitC_emitMainFn___closed__13; -x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); -return x_6; +lean_object* x_4; +x_4 = l_Lean_IR_EmitC_emitSpreads(x_1, x_2, x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_4; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__17() { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_set_panic_messages(false);", 31, 31); -return x_1; +size_t x_8; size_t x_9; lean_object* x_10; +x_8 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_9 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_10 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0(x_1, x_2, x_8, x_9, x_5, x_6, x_7); +lean_dec_ref(x_6); +lean_dec_ref(x_2); +return x_10; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__18() { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFullAppArg(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("res = ", 6, 6); -return x_1; -} -} -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__19() { -_start: +lean_object* x_6; lean_object* x_12; +if (lean_obj_tag(x_1) == 10) { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("(1 /* builtin */);", 18, 18); -return x_1; +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_18 = lean_ctor_get(x_1, 1); +x_19 = lean_ctor_get(x_1, 2); +x_20 = lean_ctor_get(x_1, 3); +x_21 = lean_unsigned_to_nat(0u); +x_22 = lean_nat_dec_eq(x_19, x_21); +if (x_22 == 0) +{ +x_12 = x_5; +goto block_17; } +else +{ +uint8_t x_23; +x_23 = lean_nat_dec_eq(x_20, x_21); +if (x_23 == 0) +{ +x_12 = x_5; +goto block_17; } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__20() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_set_panic_messages(true);", 30, 30); -return x_1; +lean_object* x_24; lean_object* x_25; +x_24 = lean_array_get_size(x_18); +x_25 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitFullAppArg_spec__0___redArg(x_24, x_18, x_2, x_21, x_3, x_4, x_5); +return x_25; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__21() { -_start: +} +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_io_mark_end_initialization();", 34, 34); -return x_1; +x_12 = x_5; +goto block_17; } +block_11: +{ +lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; +x_7 = lean_string_append(x_6, x_2); +lean_dec_ref(x_2); +x_8 = 0; +x_9 = lean_box(x_8); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_7); +return x_10; } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__22() { -_start: +block_17: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("if (lean_io_result_is_ok(res)) {", 32, 32); -return x_1; +if (lean_obj_tag(x_1) == 13) +{ +lean_object* x_13; lean_object* x_14; +lean_dec_ref(x_2); +x_13 = lean_box(x_3); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_12); +return x_14; } +else +{ +if (x_3 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = l_Lean_IR_EmitC_emitSpreadArg___closed__1; +x_16 = lean_string_append(x_12, x_15); +x_6 = x_16; +goto block_11; } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__23() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_dec_ref(res);", 18, 18); -return x_1; +x_6 = x_12; +goto block_11; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__24() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_init_task_manager();", 25, 25); -return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__25() { +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitFullAppArg_spec__0___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_IR_EmitC_emitMainFn___closed__24; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; +uint8_t x_8; lean_object* x_9; uint8_t x_14; +x_14 = lean_nat_dec_lt(x_4, x_1); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +lean_dec(x_4); +lean_dec_ref(x_3); +x_15 = lean_box(x_5); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_7); +return x_16; } +else +{ +lean_object* x_17; +x_17 = lean_array_fget_borrowed(x_2, x_4); +switch (lean_obj_tag(x_17)) { +case 6: +{ +x_8 = x_5; +x_9 = x_7; +goto block_13; } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__26() { -_start: +case 13: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_IR_EmitC_emitMainFn___closed__25; -x_2 = l_Lean_IR_EmitC_emitMainFn___closed__23; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; +x_8 = x_5; +x_9 = x_7; +goto block_13; } +default: +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_18 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__0; +lean_inc_ref(x_3); +x_19 = lean_string_append(x_3, x_18); +lean_inc(x_4); +x_20 = l_Nat_reprFast(x_4); +x_21 = lean_string_append(x_19, x_20); +lean_dec_ref(x_20); +x_22 = l_Lean_IR_EmitC_emitFullAppArg(x_17, x_21, x_5, x_6, x_7); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec_ref(x_22); +x_25 = lean_unbox(x_23); +lean_dec(x_23); +x_8 = x_25; +x_9 = x_24; +goto block_13; } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__27() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_IR_EmitC_emitMainFn___closed__26; -x_2 = l_Lean_IR_EmitC_emitMainFn___closed__22; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; +lean_dec(x_4); +lean_dec_ref(x_3); +return x_22; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__28() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_IR_EmitC_emitMainFn___closed__27; -x_2 = l_Lean_IR_EmitC_emitMainFn___closed__21; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__29() { -_start: +block_13: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("res = _lean_main();", 19, 19); -return x_1; +lean_object* x_10; lean_object* x_11; +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_add(x_4, x_10); +lean_dec(x_4); +x_4 = x_11; +x_5 = x_8; +x_7 = x_9; +goto _start; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__30() { +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitFullAppArg_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("in = lean_box(0);", 17, 17); -return x_1; +lean_object* x_11; +x_11 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitFullAppArg_spec__0___redArg(x_1, x_2, x_3, x_6, x_7, x_9, x_10); +return x_11; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__31() { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitFullAppArg_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("int i = argc;", 13, 13); -return x_1; +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_7); +x_12 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitFullAppArg_spec__0(x_1, x_2, x_3, x_4, x_5, x_6, x_11, x_8, x_9, x_10); +lean_dec_ref(x_9); +lean_dec_ref(x_2); +lean_dec(x_1); +return x_12; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__32() { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitFullAppArg_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("while (i > 1) {", 15, 15); -return x_1; +uint8_t x_8; lean_object* x_9; +x_8 = lean_unbox(x_5); +x_9 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitFullAppArg_spec__0___redArg(x_1, x_2, x_3, x_4, x_8, x_6, x_7); +lean_dec_ref(x_6); +lean_dec_ref(x_2); +lean_dec(x_1); +return x_9; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__33() { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFullAppArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked(" lean_object* n;", 16, 16); -return x_1; +uint8_t x_6; lean_object* x_7; +x_6 = lean_unbox(x_3); +x_7 = l_Lean_IR_EmitC_emitFullAppArg(x_1, x_2, x_6, x_4, x_5); +lean_dec_ref(x_4); +lean_dec(x_1); +return x_7; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__34() { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitFullAppArgs_spec__0___redArg___lam__0(uint8_t x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked(" i--;", 5, 5); -return x_1; +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_6 = l_Lean_IR_EmitC_argToCString___closed__1; +x_7 = lean_string_append(x_5, x_6); +x_8 = lean_box(x_1); +x_9 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_9, 0, x_8); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_7); +return x_10; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__35() { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitFullAppArgs_spec__0___redArg(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) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked(" n = lean_alloc_ctor(1,2,0); lean_ctor_set(n, 0, lean_mk_string(argv[i])); lean_ctor_set(n, 1, in);", 99, 99); -return x_1; -} +uint8_t x_9; lean_object* x_10; lean_object* x_15; uint8_t x_31; +x_31 = lean_nat_dec_lt(x_5, x_1); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; +lean_dec(x_5); +lean_dec_ref(x_2); +x_32 = lean_box(x_6); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_8); +return x_33; } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__36() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked(" in = n;", 8, 8); -return x_1; +lean_object* x_34; lean_object* x_35; +lean_inc_ref(x_2); +x_34 = lean_array_get_borrowed(x_2, x_3, x_5); +x_35 = lean_ctor_get(x_34, 1); +if (lean_obj_tag(x_35) == 13) +{ +x_9 = x_6; +x_10 = x_8; +goto block_14; } +else +{ +lean_object* x_36; +x_36 = lean_array_fget_borrowed(x_4, x_5); +if (lean_obj_tag(x_36) == 0) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_37 = lean_ctor_get(x_36, 0); +x_38 = l_Lean_IR_EmitC_argToCString___closed__0; +lean_inc(x_37); +x_39 = l_Nat_reprFast(x_37); +x_40 = lean_string_append(x_38, x_39); +lean_dec_ref(x_39); +x_41 = l_Lean_IR_EmitC_emitFullAppArg(x_35, x_40, x_6, x_7, x_8); +if (lean_obj_tag(x_41) == 0) +{ +lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +lean_dec_ref(x_41); +x_44 = lean_unbox(x_42); +lean_dec(x_42); +x_9 = x_44; +x_10 = x_43; +goto block_14; } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__37() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_IR_EmitC_emitMainFn___closed__12; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; +lean_dec(x_5); +lean_dec_ref(x_2); +return x_41; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__38() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_IR_EmitC_emitMainFn___closed__37; -x_2 = l_Lean_IR_EmitC_emitMainFn___closed__36; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; -} +uint8_t x_45; +x_45 = 0; +if (x_6 == 0) +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_46 = l_Lean_IR_EmitC_emitSpreadArg___closed__1; +x_47 = lean_box(0); +x_48 = lean_string_append(x_8, x_46); +x_49 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitFullAppArgs_spec__0___redArg___lam__0(x_45, x_6, x_47, x_7, x_48); +x_15 = x_49; +goto block_30; } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__39() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_IR_EmitC_emitMainFn___closed__38; -x_2 = l_Lean_IR_EmitC_emitMainFn___closed__35; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; +lean_object* x_50; lean_object* x_51; +x_50 = lean_box(0); +x_51 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitFullAppArgs_spec__0___redArg___lam__0(x_45, x_6, x_50, x_7, x_8); +x_15 = x_51; +goto block_30; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__40() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_IR_EmitC_emitMainFn___closed__39; -x_2 = l_Lean_IR_EmitC_emitMainFn___closed__34; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__41() { -_start: +block_14: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_IR_EmitC_emitMainFn___closed__40; -x_2 = l_Lean_IR_EmitC_emitMainFn___closed__33; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; +lean_object* x_11; lean_object* x_12; +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_add(x_5, x_11); +lean_dec(x_5); +x_5 = x_12; +x_6 = x_9; +x_8 = x_10; +goto _start; } +block_30: +{ +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +if (lean_obj_tag(x_16) == 0) +{ +uint8_t x_17; +lean_dec(x_5); +lean_dec_ref(x_2); +x_17 = !lean_is_exclusive(x_15); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_15, 0); +lean_dec(x_18); +x_19 = lean_ctor_get(x_16, 0); +lean_inc(x_19); +lean_dec_ref(x_16); +lean_ctor_set(x_15, 0, x_19); +return x_15; } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__42() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_IR_EmitC_emitMainFn___closed__41; -x_2 = l_Lean_IR_EmitC_emitMainFn___closed__32; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_15, 1); +lean_inc(x_20); +lean_dec(x_15); +x_21 = lean_ctor_get(x_16, 0); +lean_inc(x_21); +lean_dec_ref(x_16); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_20); +return x_22; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__43() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_IR_EmitC_emitMainFn___closed__42; -x_2 = l_Lean_IR_EmitC_emitMainFn___closed__31; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; +lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_23 = lean_ctor_get(x_15, 1); +lean_inc(x_23); +lean_dec_ref(x_15); +x_24 = lean_ctor_get(x_16, 0); +lean_inc(x_24); +lean_dec_ref(x_16); +x_25 = lean_unbox(x_24); +lean_dec(x_24); +x_9 = x_25; +x_10 = x_23; +goto block_14; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__44() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_IR_EmitC_emitMainFn___closed__43; -x_2 = l_Lean_IR_EmitC_emitMainFn___closed__30; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; -} +uint8_t x_26; +lean_dec(x_5); +lean_dec_ref(x_2); +x_26 = !lean_is_exclusive(x_15); +if (x_26 == 0) +{ +return x_15; } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__45() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("res = _lean_main(in);", 21, 21); -return x_1; +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_15, 0); +x_28 = lean_ctor_get(x_15, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_15); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__46() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("\n #if defined(WIN32) || defined(_WIN32)\n #include \n #endif\n\n int main(int argc, char ** argv) {\n #if defined(WIN32) || defined(_WIN32)\n SetErrorMode(SEM_FAILCRITICALERRORS);\n SetConsoleOutputCP(CP_UTF8);\n #endif\n lean_object* in; lean_object* res;", 267, 267); -return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__47() { +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitFullAppArgs_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, uint8_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("argv = lean_setup_args(argc, argv);", 35, 35); -return x_1; +lean_object* x_12; +x_12 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitFullAppArgs_spec__0___redArg(x_1, x_2, x_3, x_4, x_7, x_8, x_10, x_11); +return x_12; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__48() { +static lean_object* _init_l_Lean_IR_EmitC_emitFullAppArgs___closed__0() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_initialize_runtime_module();", 33, 33); +x_1 = l_Lean_IR_instInhabitedParam_default; return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__49() { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFullAppArgs(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_initialize();", 18, 18); -return x_1; -} +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; +x_5 = l_Lean_IR_EmitC_emitFullAppArgs___closed__0; +x_6 = lean_array_get_size(x_2); +x_7 = lean_unsigned_to_nat(0u); +x_8 = 1; +x_9 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitFullAppArgs_spec__0___redArg(x_6, x_5, x_1, x_2, x_7, x_8, x_3, x_4); +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); +lean_dec(x_11); +x_12 = lean_box(0); +lean_ctor_set(x_9, 0, x_12); +return x_9; } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__50() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("invalid main function, incorrect arity when generating code", 59, 59); -return x_1; +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_9, 1); +lean_inc(x_13); +lean_dec(x_9); +x_14 = lean_box(0); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_13); +return x_15; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__51() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean", 4, 4); -return x_1; -} +uint8_t x_16; +x_16 = !lean_is_exclusive(x_9); +if (x_16 == 0) +{ +return x_9; } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__52() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_IR_EmitC_emitMainFn___closed__51; -x_2 = l_Lean_Name_mkStr1(x_1); -return x_2; +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_9, 0); +x_18 = lean_ctor_get(x_9, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_9); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +return x_19; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__53() { +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitFullAppArgs_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("char ** lean_setup_args(int argc, char ** argv);", 48, 48); -return x_1; +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_8); +x_13 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitFullAppArgs_spec__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_12, x_9, x_10, x_11); +lean_dec_ref(x_10); +lean_dec_ref(x_4); +lean_dec_ref(x_3); +lean_dec(x_1); +return x_13; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__54() { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitFullAppArgs_spec__0___redArg___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("void lean_initialize_runtime_module();", 38, 38); -return x_1; +uint8_t x_6; uint8_t x_7; lean_object* x_8; +x_6 = lean_unbox(x_1); +x_7 = lean_unbox(x_2); +x_8 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitFullAppArgs_spec__0___redArg___lam__0(x_6, x_7, x_3, x_4, x_5); +lean_dec_ref(x_4); +return x_8; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__55() { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFullAppArgs___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("void lean_initialize();", 23, 23); -return x_1; +lean_object* x_5; +x_5 = l_Lean_IR_EmitC_emitFullAppArgs(x_1, x_2, x_3, x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_5; } } -static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__56() { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitFullAppArgs_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("function declaration expected", 29, 29); -return x_1; +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_6); +x_10 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitFullAppArgs_spec__0___redArg(x_1, x_2, x_3, x_4, x_5, x_9, x_7, x_8); +lean_dec_ref(x_7); +lean_dec_ref(x_4); +lean_dec_ref(x_3); +lean_dec(x_1); +return x_10; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitMainFn(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_75; -x_59 = l_Lean_IR_EmitC_toCName___closed__1; -lean_inc_ref(x_1); -x_75 = l_Lean_IR_EmitC_getDecl(x_59, x_1, x_2); -if (lean_obj_tag(x_75) == 0) +lean_object* x_5; uint8_t x_10; +x_10 = lean_usize_dec_eq(x_2, x_3); +if (x_10 == 0) { -lean_object* x_76; -x_76 = lean_ctor_get(x_75, 0); -lean_inc(x_76); -if (lean_obj_tag(x_76) == 0) +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = lean_array_uget(x_1, x_2); +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +x_13 = l_Lean_IR_IRType_isErased(x_12); +lean_dec(x_12); +if (x_13 == 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_123; uint8_t x_124; lean_object* x_125; lean_object* x_126; uint8_t x_141; lean_object* x_158; lean_object* x_159; uint8_t x_160; -x_77 = lean_ctor_get(x_75, 1); -lean_inc(x_77); -if (lean_is_exclusive(x_75)) { - lean_ctor_release(x_75, 0); - lean_ctor_release(x_75, 1); - x_78 = x_75; -} else { - lean_dec_ref(x_75); - x_78 = lean_box(0); +lean_object* x_14; +x_14 = lean_array_push(x_4, x_11); +x_5 = x_14; +goto block_9; } -x_79 = lean_ctor_get(x_76, 1); -lean_inc_ref(x_79); -lean_dec_ref(x_76); -x_158 = lean_array_get_size(x_79); -x_159 = lean_unsigned_to_nat(2u); -x_160 = lean_nat_dec_eq(x_158, x_159); -if (x_160 == 0) +else { -lean_object* x_161; uint8_t x_162; -x_161 = lean_unsigned_to_nat(1u); -x_162 = lean_nat_dec_eq(x_158, x_161); -x_141 = x_162; -goto block_157; +lean_dec(x_11); +x_5 = x_4; +goto block_9; +} } else { -x_141 = x_160; -goto block_157; +return x_4; } -block_122: +block_9: { -lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; uint8_t x_112; -x_83 = lean_ctor_get(x_81, 0); -x_84 = lean_ctor_get(x_81, 1); -x_85 = l_Lean_IR_EmitC_getModInitFn___closed__0; -x_86 = lean_ctor_get(x_85, 0); -lean_inc_ref(x_86); -x_87 = lean_ctor_get(x_86, 2); -lean_inc(x_87); -lean_dec_ref(x_86); -x_88 = l_Lean_IR_EmitC_emitMainFn___closed__17; -x_89 = lean_string_append(x_82, x_88); -x_90 = l_Lean_IR_EmitC_emitLn___closed__0; -x_91 = lean_string_append(x_89, x_90); -x_92 = lean_box(0); -x_93 = lean_box(0); -lean_inc_ref(x_83); -x_94 = l_Lean_PersistentEnvExtension_getState___redArg(x_92, x_85, x_83, x_87, x_93); -lean_dec(x_87); -lean_inc(x_84); -x_95 = l_Lean_mkModuleInitializationFunctionName(x_84, x_94); -lean_dec(x_94); -x_96 = l_Lean_IR_EmitC_emitMainFn___closed__18; -x_97 = lean_string_append(x_96, x_95); -lean_dec_ref(x_95); -x_98 = l_Lean_IR_EmitC_emitMainFn___closed__19; -x_99 = lean_string_append(x_97, x_98); -x_100 = lean_string_append(x_91, x_99); -lean_dec_ref(x_99); -x_101 = lean_string_append(x_100, x_90); -x_102 = l_Lean_IR_EmitC_emitMainFn___closed__20; -x_103 = lean_string_append(x_101, x_102); -x_104 = lean_string_append(x_103, x_90); -x_105 = l_Lean_IR_EmitC_emitMainFn___closed__22; -x_106 = lean_box(0); -x_107 = l_Lean_IR_EmitC_emitMainFn___closed__28; -x_108 = l_List_forM___at___00Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0_spec__0___redArg(x_107, x_104); -x_109 = lean_ctor_get(x_108, 1); -lean_inc(x_109); -lean_dec_ref(x_108); -x_110 = lean_array_get_size(x_79); -lean_dec_ref(x_79); -x_111 = lean_unsigned_to_nat(2u); -x_112 = lean_nat_dec_eq(x_110, x_111); -if (x_112 == 0) +size_t x_6; size_t x_7; +x_6 = 1; +x_7 = lean_usize_add(x_2, x_6); +x_2 = x_7; +x_4 = x_5; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__1(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +_start: { -lean_object* x_113; lean_object* x_114; lean_object* x_115; -x_113 = l_Lean_IR_EmitC_emitMainFn___closed__29; -x_114 = lean_string_append(x_109, x_113); -x_115 = lean_string_append(x_114, x_90); -x_60 = x_105; -x_61 = x_106; -x_62 = x_80; -x_63 = x_90; -x_64 = x_81; -x_65 = x_115; -goto block_74; +lean_object* x_5; uint8_t x_10; +x_10 = lean_usize_dec_eq(x_2, x_3); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = lean_array_uget(x_1, x_2); +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +x_13 = l_Lean_IR_IRType_isVoid(x_12); +lean_dec(x_12); +if (x_13 == 0) +{ +lean_object* x_14; +x_14 = lean_array_push(x_4, x_11); +x_5 = x_14; +goto block_9; } else { -lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; -x_116 = l_Lean_IR_EmitC_emitMainFn___closed__44; -x_117 = l_List_forM___at___00Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0_spec__0___redArg(x_116, x_109); -x_118 = lean_ctor_get(x_117, 1); -lean_inc(x_118); -lean_dec_ref(x_117); -x_119 = l_Lean_IR_EmitC_emitMainFn___closed__45; -x_120 = lean_string_append(x_118, x_119); -x_121 = lean_string_append(x_120, x_90); -x_60 = x_105; -x_61 = x_106; -x_62 = x_80; -x_63 = x_90; -x_64 = x_81; -x_65 = x_121; -goto block_74; +lean_dec(x_11); +x_5 = x_4; +goto block_9; } } -block_140: +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 = l_Lean_IR_EmitC_emitMainFn___closed__46; -x_128 = lean_string_append(x_126, x_127); -x_129 = l_Lean_IR_EmitC_emitLn___closed__0; -x_130 = lean_string_append(x_128, x_129); -x_131 = l_Lean_IR_EmitC_emitMainFn___closed__47; -x_132 = lean_string_append(x_130, x_131); -x_133 = lean_string_append(x_132, x_129); -if (x_124 == 0) +return x_4; +} +block_9: { -lean_object* x_134; lean_object* x_135; lean_object* x_136; -x_134 = l_Lean_IR_EmitC_emitMainFn___closed__48; -x_135 = lean_string_append(x_133, x_134); -x_136 = lean_string_append(x_135, x_129); -x_80 = x_123; -x_81 = x_125; -x_82 = x_136; -goto block_122; +size_t x_6; size_t x_7; +x_6 = 1; +x_7 = lean_usize_add(x_2, x_6); +x_2 = x_7; +x_4 = x_5; +goto _start; } -else +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitFnDeclAux___closed__0() { +_start: { -lean_object* x_137; lean_object* x_138; lean_object* x_139; -x_137 = l_Lean_IR_EmitC_emitMainFn___closed__49; -x_138 = lean_string_append(x_133, x_137); -x_139 = lean_string_append(x_138, x_129); -x_80 = x_123; -x_81 = x_125; -x_82 = x_139; -goto block_122; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_object**", 13, 13); +return x_1; } } -block_157: +static lean_object* _init_l_Lean_IR_EmitC_emitFnDeclAux___closed__1() { +_start: { -if (x_141 == 0) +lean_object* x_1; +x_1 = l_Lean_closureMaxArgs; +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitFnDeclAux___closed__2() { +_start: { -lean_object* x_142; lean_object* x_143; -lean_dec_ref(x_79); -lean_dec_ref(x_1); -x_142 = l_Lean_IR_EmitC_emitMainFn___closed__50; -if (lean_is_scalar(x_78)) { - x_143 = lean_alloc_ctor(1, 2, 0); -} else { - x_143 = x_78; - lean_ctor_set_tag(x_143, 1); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("(", 1, 1); +return x_1; } -lean_ctor_set(x_143, 0, x_142); -lean_ctor_set(x_143, 1, x_77); -return x_143; } -else +static lean_object* _init_l_Lean_IR_EmitC_emitFnDeclAux___closed__3() { +_start: { -lean_object* x_144; lean_object* x_145; uint8_t x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; -lean_dec(x_78); -x_144 = lean_ctor_get(x_1, 0); -lean_inc_ref(x_144); -x_145 = l_Lean_IR_EmitC_emitMainFn___closed__52; -x_146 = l_Lean_IR_usesModuleFrom(x_144, x_145); -x_147 = l_Lean_IR_EmitC_emitMainFn___closed__53; -x_148 = lean_string_append(x_77, x_147); -x_149 = l_Lean_IR_EmitC_emitLn___closed__0; -x_150 = lean_string_append(x_148, x_149); -if (x_146 == 0) +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitFnDeclAux___closed__4() { +_start: { -lean_object* x_151; lean_object* x_152; lean_object* x_153; -x_151 = l_Lean_IR_EmitC_emitMainFn___closed__54; -x_152 = lean_string_append(x_150, x_151); -x_153 = lean_string_append(x_152, x_149); -x_123 = x_144; -x_124 = x_146; -x_125 = x_1; -x_126 = x_153; -goto block_140; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("LEAN_EXPORT ", 12, 12); +return x_1; } -else +} +static lean_object* _init_l_Lean_IR_EmitC_emitFnDeclAux___closed__5() { +_start: { -lean_object* x_154; lean_object* x_155; lean_object* x_156; -x_154 = l_Lean_IR_EmitC_emitMainFn___closed__55; -x_155 = lean_string_append(x_150, x_154); -x_156 = lean_string_append(x_155, x_149); -x_123 = x_144; -x_124 = x_146; -x_125 = x_1; -x_126 = x_156; -goto block_140; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("static ", 7, 7); +return x_1; } } +static lean_object* _init_l_Lean_IR_EmitC_emitFnDeclAux___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("extern ", 7, 7); +return x_1; +} } +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFnDeclAux(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; lean_object* x_14; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_83; +x_36 = lean_ctor_get(x_4, 0); +lean_inc_ref(x_36); +x_54 = l_Lean_IR_Decl_params(x_1); +x_83 = l_Array_isEmpty___redArg(x_54); +if (x_83 == 0) +{ +if (x_3 == 0) +{ +goto block_82; } else { -uint8_t x_163; -lean_dec(x_76); -lean_dec_ref(x_1); -x_163 = !lean_is_exclusive(x_75); -if (x_163 == 0) +if (x_83 == 0) { -lean_object* x_164; lean_object* x_165; -x_164 = lean_ctor_get(x_75, 0); -lean_dec(x_164); -x_165 = l_Lean_IR_EmitC_emitMainFn___closed__56; -lean_ctor_set_tag(x_75, 1); -lean_ctor_set(x_75, 0, x_165); -return x_75; +x_55 = x_4; +x_56 = x_5; +goto block_79; } else { -lean_object* x_166; lean_object* x_167; lean_object* x_168; -x_166 = lean_ctor_get(x_75, 1); -lean_inc(x_166); -lean_dec(x_75); -x_167 = l_Lean_IR_EmitC_emitMainFn___closed__56; -x_168 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_168, 0, x_167); -lean_ctor_set(x_168, 1, x_166); -return x_168; +goto block_82; } } } else { -uint8_t x_169; -lean_dec_ref(x_1); -x_169 = !lean_is_exclusive(x_75); -if (x_169 == 0) +if (x_3 == 0) { -return x_75; +lean_object* x_84; uint8_t x_85; +x_84 = l_Lean_IR_Decl_name(x_1); +lean_inc_ref(x_36); +x_85 = l_Lean_isClosedTermName(x_36, x_84); +lean_dec(x_84); +if (x_85 == 0) +{ +lean_object* x_86; lean_object* x_87; +x_86 = l_Lean_IR_EmitC_emitFnDeclAux___closed__4; +x_87 = lean_string_append(x_5, x_86); +x_55 = x_4; +x_56 = x_87; +goto block_79; } else { -lean_object* x_170; lean_object* x_171; lean_object* x_172; -x_170 = lean_ctor_get(x_75, 0); -x_171 = lean_ctor_get(x_75, 1); -lean_inc(x_171); -lean_inc(x_170); -lean_dec(x_75); -x_172 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_172, 0, x_170); -lean_ctor_set(x_172, 1, x_171); -return x_172; +lean_object* x_88; lean_object* x_89; +x_88 = l_Lean_IR_EmitC_emitFnDeclAux___closed__5; +x_89 = lean_string_append(x_5, x_88); +x_55 = x_4; +x_56 = x_89; +goto block_79; } } -block_40: +else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; -lean_dec_ref(x_3); -x_12 = lean_string_append(x_10, x_11); -lean_dec_ref(x_11); -x_13 = l_Lean_IR_EmitC_emitMainFn___closed__0; -x_14 = l_Lean_IR_EmitC_emitMainFn___closed__1; -x_15 = l_Lean_IR_EmitC_emitMainFn___closed__2; -x_16 = l_Lean_IR_EmitC_emitMainFn___closed__3; -x_17 = l_Lean_IR_EmitC_emitMainFn___closed__4; -lean_inc_ref(x_8); -x_18 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_18, 0, x_8); -lean_ctor_set(x_18, 1, x_5); -x_19 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_19, 1, x_18); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_13); -lean_ctor_set(x_20, 1, x_19); -x_21 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_16); -lean_ctor_set(x_21, 1, x_20); -x_22 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_22, 0, x_15); -lean_ctor_set(x_22, 1, x_21); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_14); -lean_ctor_set(x_23, 1, x_22); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_13); -lean_ctor_set(x_24, 1, x_23); -x_25 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_25, 0, x_12); -lean_ctor_set(x_25, 1, x_24); -x_26 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_26, 0, x_4); -lean_ctor_set(x_26, 1, x_25); -x_27 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_27, 0, x_6); -lean_ctor_set(x_27, 1, x_26); -x_28 = l_List_forM___at___00Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0_spec__0___redArg(x_27, x_7); -lean_dec_ref(x_27); -x_29 = !lean_is_exclusive(x_28); -if (x_29 == 0) +lean_object* x_90; lean_object* x_91; +x_90 = l_Lean_IR_EmitC_emitFnDeclAux___closed__6; +x_91 = lean_string_append(x_5, x_90); +x_55 = x_4; +x_56 = x_91; +goto block_79; +} +} +block_13: { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_30 = lean_ctor_get(x_28, 1); -x_31 = lean_ctor_get(x_28, 0); -lean_dec(x_31); -x_32 = lean_string_append(x_30, x_8); -lean_dec_ref(x_8); -x_33 = lean_box(0); -x_34 = lean_string_append(x_32, x_9); -lean_dec_ref(x_9); -lean_ctor_set(x_28, 1, x_34); -lean_ctor_set(x_28, 0, x_33); -return x_28; +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_7 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_8 = lean_string_append(x_6, x_7); +x_9 = l_Lean_IR_EmitC_emitLn___closed__0; +x_10 = lean_box(0); +x_11 = lean_string_append(x_8, x_9); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set(x_12, 1, x_11); +return x_12; +} +block_17: +{ +lean_object* x_15; lean_object* x_16; +x_15 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__1; +x_16 = lean_string_append(x_14, x_15); +x_6 = x_16; +goto block_13; +} +block_26: +{ +if (x_21 == 0) +{ +lean_object* x_22; +x_22 = l_Lean_IR_EmitC_emitSpreadArgs(x_19, x_21, x_18, x_20); +lean_dec_ref(x_18); +lean_dec_ref(x_19); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec_ref(x_22); +x_14 = x_23; +goto block_17; } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_35 = lean_ctor_get(x_28, 1); -lean_inc(x_35); -lean_dec(x_28); -x_36 = lean_string_append(x_35, x_8); -lean_dec_ref(x_8); -x_37 = lean_box(0); -x_38 = lean_string_append(x_36, x_9); -lean_dec_ref(x_9); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_37); -lean_ctor_set(x_39, 1, x_38); -return x_39; +return x_22; } } -block_58: +else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; -x_48 = l_Lean_ConstantInfo_type(x_47); -lean_dec_ref(x_47); -x_49 = l_Lean_Expr_getForallBody(x_48); -lean_dec_ref(x_48); -x_50 = l_Lean_Expr_appArg_x21(x_49); -lean_dec_ref(x_49); -x_51 = l_Lean_IR_EmitC_emitMainFn___closed__5; -x_52 = l_Lean_IR_EmitC_emitMainFn___closed__6; -x_53 = l_Lean_Expr_constName_x3f(x_50); -lean_dec_ref(x_50); -x_54 = l_Lean_IR_EmitC_emitMainFn___closed__9; -x_55 = l_Option_instBEq_beq___at___00Lean_IR_EmitC_emitMainFn_spec__1(x_53, x_54); -lean_dec(x_53); -if (x_55 == 0) +lean_object* x_24; lean_object* x_25; +lean_dec_ref(x_19); +lean_dec_ref(x_18); +x_24 = l_Lean_IR_EmitC_emitFnDeclAux___closed__0; +x_25 = lean_string_append(x_20, x_24); +x_14 = x_25; +goto block_17; +} +} +block_35: { -lean_object* x_56; -x_56 = l_Lean_IR_EmitC_emitMainFn___closed__10; -x_3 = x_42; -x_4 = x_41; -x_5 = x_43; -x_6 = x_51; -x_7 = x_44; -x_8 = x_45; -x_9 = x_46; -x_10 = x_52; -x_11 = x_56; -goto block_40; +lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_31 = l_Lean_IR_EmitC_emitFnDeclAux___closed__1; +x_32 = lean_array_get_size(x_30); +x_33 = lean_nat_dec_lt(x_31, x_32); +if (x_33 == 0) +{ +lean_dec(x_28); +x_18 = x_27; +x_19 = x_30; +x_20 = x_29; +x_21 = x_33; +goto block_26; } else { -lean_object* x_57; -x_57 = l_Lean_IR_EmitC_emitMainFn___closed__11; -x_3 = x_42; -x_4 = x_41; -x_5 = x_43; -x_6 = x_51; -x_7 = x_44; -x_8 = x_45; -x_9 = x_46; -x_10 = x_52; -x_11 = x_57; -goto block_40; +uint8_t x_34; +x_34 = l_Lean_IR_ExplicitBoxing_isBoxedName(x_28); +lean_dec(x_28); +x_18 = x_27; +x_19 = x_30; +x_20 = x_29; +x_21 = x_34; +goto block_26; } } -block_74: +block_53: { -lean_object* x_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; lean_object* x_70; -x_66 = l_Lean_IR_EmitC_emitMainFn___closed__12; -x_67 = lean_string_append(x_65, x_66); -x_68 = lean_string_append(x_67, x_63); -x_69 = 0; -x_70 = l_Lean_Environment_find_x3f(x_62, x_59, x_69); -if (lean_obj_tag(x_70) == 0) +lean_object* x_41; uint8_t x_42; +x_41 = l_Lean_IR_Decl_name(x_1); +lean_inc(x_41); +x_42 = l_Lean_isExternC(x_36, x_41); +if (x_42 == 0) { -lean_object* x_71; lean_object* x_72; -x_71 = l_Lean_IR_EmitC_emitMainFn___closed__16; -x_72 = l_panic___at___00Lean_IR_EmitC_emitMainFn_spec__2(x_71); -x_41 = x_60; -x_42 = x_64; -x_43 = x_61; -x_44 = x_68; -x_45 = x_66; -x_46 = x_63; -x_47 = x_72; -goto block_58; +x_27 = x_37; +x_28 = x_41; +x_29 = x_38; +x_30 = x_40; +goto block_35; } else { -lean_object* x_73; -x_73 = lean_ctor_get(x_70, 0); -lean_inc(x_73); -lean_dec_ref(x_70); -x_41 = x_60; -x_42 = x_64; -x_43 = x_61; -x_44 = x_68; -x_45 = x_66; -x_46 = x_63; -x_47 = x_73; -goto block_58; +lean_object* x_43; lean_object* x_44; uint8_t x_45; +x_43 = lean_array_get_size(x_40); +x_44 = lean_mk_empty_array_with_capacity(x_39); +x_45 = lean_nat_dec_lt(x_39, x_43); +if (x_45 == 0) +{ +lean_dec_ref(x_40); +x_27 = x_37; +x_28 = x_41; +x_29 = x_38; +x_30 = x_44; +goto block_35; } +else +{ +uint8_t x_46; +x_46 = lean_nat_dec_le(x_43, x_43); +if (x_46 == 0) +{ +if (x_45 == 0) +{ +lean_dec_ref(x_40); +x_27 = x_37; +x_28 = x_41; +x_29 = x_38; +x_30 = x_44; +goto block_35; } +else +{ +size_t x_47; size_t x_48; lean_object* x_49; +x_47 = 0; +x_48 = lean_usize_of_nat(x_43); +x_49 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0(x_40, x_47, x_48, x_44); +lean_dec_ref(x_40); +x_27 = x_37; +x_28 = x_41; +x_29 = x_38; +x_30 = x_49; +goto block_35; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -lean_object* x_4; -x_4 = l_Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0(x_1, x_2, x_3); -lean_dec_ref(x_2); -lean_dec(x_1); -return x_4; +size_t x_50; size_t x_51; lean_object* x_52; +x_50 = 0; +x_51 = lean_usize_of_nat(x_43); +x_52 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0(x_40, x_50, x_51, x_44); +lean_dec_ref(x_40); +x_27 = x_37; +x_28 = x_41; +x_29 = x_38; +x_30 = x_52; +goto block_35; } } -LEAN_EXPORT lean_object* l_List_forM___at___00Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +} +} +block_79: { -lean_object* x_4; -x_4 = l_List_forM___at___00Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0_spec__0(x_1, x_2, x_3); -lean_dec_ref(x_2); -lean_dec(x_1); -return x_4; +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; uint8_t x_65; +x_57 = l_Lean_IR_Decl_resultType(x_1); +x_58 = l_Lean_IR_EmitC_toCType(x_57, x_55, x_56); +lean_dec(x_57); +x_59 = lean_ctor_get(x_58, 0); +lean_inc(x_59); +x_60 = lean_ctor_get(x_58, 1); +lean_inc(x_60); +lean_dec_ref(x_58); +x_61 = l_Lean_IR_EmitC_emitSpreadArg___closed__0; +x_62 = lean_string_append(x_59, x_61); +x_63 = lean_string_append(x_62, x_2); +x_64 = lean_string_append(x_60, x_63); +lean_dec_ref(x_63); +x_65 = l_Array_isEmpty___redArg(x_54); +if (x_65 == 0) +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; +x_66 = l_Lean_IR_EmitC_emitFnDeclAux___closed__2; +x_67 = lean_string_append(x_64, x_66); +x_68 = lean_unsigned_to_nat(0u); +x_69 = lean_array_get_size(x_54); +x_70 = l_Lean_IR_EmitC_emitFnDeclAux___closed__3; +x_71 = lean_nat_dec_lt(x_68, x_69); +if (x_71 == 0) +{ +lean_dec_ref(x_54); +x_37 = x_55; +x_38 = x_67; +x_39 = x_68; +x_40 = x_70; +goto block_53; } +else +{ +uint8_t x_72; +x_72 = lean_nat_dec_le(x_69, x_69); +if (x_72 == 0) +{ +if (x_71 == 0) +{ +lean_dec_ref(x_54); +x_37 = x_55; +x_38 = x_67; +x_39 = x_68; +x_40 = x_70; +goto block_53; } -LEAN_EXPORT lean_object* l_List_forM___at___00Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2) { -_start: +else { -lean_object* x_3; -x_3 = l_List_forM___at___00Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0_spec__0___redArg(x_1, x_2); -lean_dec(x_1); -return x_3; +size_t x_73; size_t x_74; lean_object* x_75; +x_73 = 0; +x_74 = lean_usize_of_nat(x_69); +x_75 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__1(x_54, x_73, x_74, x_70); +lean_dec_ref(x_54); +x_37 = x_55; +x_38 = x_67; +x_39 = x_68; +x_40 = x_75; +goto block_53; } } -LEAN_EXPORT lean_object* l_Option_instBEq_beq___at___00Lean_IR_EmitC_emitMainFn_spec__1___boxed(lean_object* x_1, lean_object* x_2) { -_start: +else { -uint8_t x_3; lean_object* x_4; -x_3 = l_Option_instBEq_beq___at___00Lean_IR_EmitC_emitMainFn_spec__1(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(x_3); -return x_4; +size_t x_76; size_t x_77; lean_object* x_78; +x_76 = 0; +x_77 = lean_usize_of_nat(x_69); +x_78 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__1(x_54, x_76, x_77, x_70); +lean_dec_ref(x_54); +x_37 = x_55; +x_38 = x_67; +x_39 = x_68; +x_40 = x_78; +goto block_53; } } -LEAN_EXPORT uint8_t l_Lean_IR_EmitC_hasMainFn___lam__0(lean_object* x_1) { -_start: +} +else { -lean_object* x_2; lean_object* x_3; uint8_t x_4; -x_2 = l_Lean_IR_Decl_name(x_1); -x_3 = l_Lean_IR_EmitC_toCName___closed__1; -x_4 = lean_name_eq(x_2, x_3); -lean_dec(x_2); -return x_4; +lean_dec_ref(x_55); +lean_dec_ref(x_54); +lean_dec_ref(x_36); +x_6 = x_64; +goto block_13; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_hasMainFn___lam__0___boxed(lean_object* x_1) { -_start: +block_82: { -uint8_t x_2; lean_object* x_3; -x_2 = l_Lean_IR_EmitC_hasMainFn___lam__0(x_1); -lean_dec_ref(x_1); -x_3 = lean_box(x_2); -return x_3; +lean_object* x_80; lean_object* x_81; +x_80 = l_Lean_IR_EmitC_emitFnDeclAux___closed__4; +x_81 = lean_string_append(x_5, x_80); +x_55 = x_4; +x_56 = x_81; +goto block_79; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_hasMainFn(lean_object* x_1, lean_object* x_2) { +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; -x_3 = lean_ctor_get(x_1, 0); -lean_inc_ref(x_3); +size_t x_5; size_t x_6; lean_object* x_7; +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0(x_1, x_5, x_6, x_4); lean_dec_ref(x_1); -x_4 = lean_alloc_closure((void*)(l_Lean_IR_EmitC_hasMainFn___lam__0___boxed), 1, 0); -x_5 = l_Lean_IR_getDecls(x_3); -x_6 = l_List_any___redArg(x_5, x_4); -x_7 = lean_box(x_6); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_7); -lean_ctor_set(x_8, 1, x_2); -return x_8; +return x_7; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitMainFnIfNeeded(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_3; lean_object* x_4; uint8_t x_5; -lean_inc_ref(x_1); -x_3 = l_Lean_IR_EmitC_hasMainFn(x_1, x_2); -x_4 = lean_ctor_get(x_3, 0); -lean_inc(x_4); -x_5 = lean_unbox(x_4); -lean_dec(x_4); -if (x_5 == 0) +size_t x_5; size_t x_6; lean_object* x_7; +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__1(x_1, x_5, x_6, x_4); +lean_dec_ref(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFnDeclAux___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -uint8_t x_6; +uint8_t x_6; lean_object* x_7; +x_6 = lean_unbox(x_3); +x_7 = l_Lean_IR_EmitC_emitFnDeclAux(x_1, x_2, x_6, x_4, x_5); +lean_dec_ref(x_2); lean_dec_ref(x_1); -x_6 = !lean_is_exclusive(x_3); -if (x_6 == 0) +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFnDecl(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_7; lean_object* x_8; -x_7 = lean_ctor_get(x_3, 0); +lean_object* x_5; lean_object* x_6; +x_5 = l_Lean_IR_Decl_name(x_1); +lean_inc_ref(x_3); +x_6 = l_Lean_IR_EmitC_toCName(x_5, x_3, x_4); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +lean_dec_ref(x_6); +x_9 = l_Lean_IR_EmitC_emitFnDeclAux(x_1, x_7, x_2, x_3, x_8); lean_dec(x_7); -x_8 = lean_box(0); -lean_ctor_set(x_3, 0, x_8); -return x_3; +return x_9; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_9 = lean_ctor_get(x_3, 1); -lean_inc(x_9); -lean_dec(x_3); -x_10 = lean_box(0); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_10); -lean_ctor_set(x_11, 1, x_9); -return x_11; -} +uint8_t x_10; +lean_dec_ref(x_3); +x_10 = !lean_is_exclusive(x_6); +if (x_10 == 0) +{ +return x_6; } else { -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_3, 1); +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_6, 0); +x_12 = lean_ctor_get(x_6, 1); lean_inc(x_12); -lean_dec_ref(x_3); -x_13 = l_Lean_IR_EmitC_emitMainFn(x_1, x_12); +lean_inc(x_11); +lean_dec(x_6); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_11); +lean_ctor_set(x_13, 1, x_12); return x_13; } } } -static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__0() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("import ", 7, 7); -return x_1; -} } -static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__1() { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFnDecl___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("all ", 4, 4); -return x_1; +uint8_t x_5; lean_object* x_6; +x_5 = lean_unbox(x_2); +x_6 = l_Lean_IR_EmitC_emitFnDecl(x_1, x_5, x_3, x_4); +lean_dec_ref(x_1); +return x_6; } } -static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__2() { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitExternDeclAux(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("meta ", 5, 5); -return x_1; +lean_object* x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; +x_5 = lean_ctor_get(x_3, 0); +x_6 = l_Lean_IR_Decl_name(x_1); +lean_inc_ref(x_5); +x_7 = l_Lean_isExternC(x_5, x_6); +x_8 = l_Lean_IR_EmitC_emitFnDeclAux(x_1, x_2, x_7, x_3, x_4); +return x_8; } } -static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__3() { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitExternDeclAux___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("public ", 7, 7); -return x_1; +lean_object* x_5; +x_5 = l_Lean_IR_EmitC_emitExternDeclAux(x_1, x_2, x_3, x_4); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_5; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_List_foldl___at___00Lean_IR_EmitC_emitFnDecls_spec__0(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_6; -x_6 = lean_usize_dec_eq(x_2, x_3); -if (x_6 == 0) -{ -lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_24; lean_object* x_25; lean_object* x_33; -x_7 = lean_array_uget(x_1, x_2); -x_8 = lean_ctor_get_uint8(x_7, sizeof(void*)*1 + 1); -x_9 = l_Lean_IR_EmitC_emitFnDeclAux___closed__4; -if (x_8 == 0) +if (lean_obj_tag(x_2) == 0) { -lean_object* x_38; -x_38 = l_panic___at___00Lean_IR_EmitC_toCType_spec__0___closed__0; -x_33 = x_38; -goto block_37; +return x_1; } else { -lean_object* x_39; -x_39 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__3; -x_33 = x_39; -goto block_37; -} -block_23: -{ -lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; size_t x_20; size_t x_21; -x_12 = lean_ctor_get(x_7, 0); -lean_inc(x_12); -lean_dec(x_7); -x_13 = lean_string_append(x_10, x_11); -lean_dec_ref(x_11); -x_14 = 1; -x_15 = l_Lean_Name_toStringWithToken___at___00Lean_Name_toString_spec__0(x_12, x_14); -x_16 = lean_string_append(x_13, x_15); -lean_dec_ref(x_15); -x_17 = lean_string_append(x_9, x_16); -lean_dec_ref(x_16); -x_18 = lean_box(0); -x_19 = lean_string_append(x_5, x_17); -lean_dec_ref(x_17); -x_20 = 1; -x_21 = lean_usize_add(x_2, x_20); -x_2 = x_21; -x_4 = x_18; -x_5 = x_19; +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_3 = lean_ctor_get(x_2, 0); +x_4 = lean_ctor_get(x_2, 1); +x_5 = l_Lean_IR_Decl_name(x_3); +x_6 = l_Lean_NameSet_insert(x_1, x_5); +x_1 = x_6; +x_2 = x_4; goto _start; } -block_32: +} +} +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_foldrM___at___00Lean_IR_EmitC_emitFnDecls_spec__2(lean_object* x_1, lean_object* x_2) { +_start: { -uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_26 = lean_ctor_get_uint8(x_7, sizeof(void*)*1); -x_27 = lean_string_append(x_24, x_25); -lean_dec_ref(x_25); -x_28 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__0; -x_29 = lean_string_append(x_27, x_28); -if (x_26 == 0) +if (lean_obj_tag(x_2) == 0) { -lean_object* x_30; -x_30 = l_panic___at___00Lean_IR_EmitC_toCType_spec__0___closed__0; -x_10 = x_29; -x_11 = x_30; -goto block_23; +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_3 = lean_ctor_get(x_2, 1); +x_4 = lean_ctor_get(x_2, 3); +x_5 = lean_ctor_get(x_2, 4); +x_6 = l_Std_DTreeMap_Internal_Impl_foldrM___at___00Lean_IR_EmitC_emitFnDecls_spec__2(x_1, x_5); +lean_inc(x_3); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_3); +lean_ctor_set(x_7, 1, x_6); +x_1 = x_7; +x_2 = x_4; +goto _start; } else { -lean_object* x_31; -x_31 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__1; -x_10 = x_29; -x_11 = x_31; -goto block_23; +return x_1; } } -block_37: +} +LEAN_EXPORT lean_object* l_List_foldl___at___00Lean_IR_EmitC_emitFnDecls_spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -uint8_t x_34; -x_34 = lean_ctor_get_uint8(x_7, sizeof(void*)*1 + 2); -if (x_34 == 0) +if (lean_obj_tag(x_3) == 0) { -lean_object* x_35; -x_35 = l_panic___at___00Lean_IR_EmitC_toCType_spec__0___closed__0; -x_24 = x_33; -x_25 = x_35; -goto block_32; +lean_dec_ref(x_1); +return x_2; } else { -lean_object* x_36; -x_36 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__2; -x_24 = x_33; -x_25 = x_36; -goto block_32; +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_3, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_3, 1); +lean_inc(x_5); +lean_dec_ref(x_3); +x_6 = l_Lean_IR_Decl_name(x_4); +x_7 = l_Lean_NameSet_insert(x_2, x_6); +lean_inc_ref(x_1); +x_8 = l_Lean_IR_collectUsedDecls(x_1, x_4, x_7); +x_2 = x_8; +x_3 = x_5; +goto _start; } } } -else -{ -lean_object* x_40; -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_4); -lean_ctor_set(x_40, 1, x_5); -return x_40; -} -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; -x_7 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg(x_1, x_2, x_3, x_4, x_6); -return x_7; -} -} -static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__0() { -_start: +static lean_object* _init_l_Lean_IR_EmitC_emitFnDecls___closed__0() { +_start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("#include ", 22, 22); +x_1 = l_Lean_NameSet_empty; return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__1() { +static lean_object* _init_l_List_forM___at___00Lean_IR_EmitC_emitFnDecls_spec__3___closed__0() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("#if defined(__clang__)", 22, 22); +x_1 = lean_mk_string_unchecked("c", 1, 1); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__2() { +static lean_object* _init_l_List_forM___at___00Lean_IR_EmitC_emitFnDecls_spec__3___closed__1() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("#pragma clang diagnostic ignored \"-Wunused-parameter\"", 53, 53); -return x_1; +lean_object* x_1; lean_object* x_2; +x_1 = l_List_forM___at___00Lean_IR_EmitC_emitFnDecls_spec__3___closed__0; +x_2 = l_Lean_Name_mkStr1(x_1); +return x_2; } } -static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__3() { +LEAN_EXPORT lean_object* l_List_forM___at___00Lean_IR_EmitC_emitFnDecls_spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("#pragma clang diagnostic ignored \"-Wunused-label\"", 49, 49); -return x_1; -} -} -static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__4() { -_start: +if (lean_obj_tag(x_3) == 0) { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("#elif defined(__GNUC__) && !defined(__CLANG__)", 46, 46); -return x_1; -} +lean_object* x_6; lean_object* x_7; +lean_dec_ref(x_4); +lean_dec_ref(x_1); +x_6 = lean_box(0); +x_7 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_5); +return x_7; } -static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__5() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("#pragma GCC diagnostic ignored \"-Wunused-parameter\"", 51, 51); -return x_1; -} +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_14; +x_8 = lean_ctor_get(x_3, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_3, 1); +lean_inc(x_9); +lean_dec_ref(x_3); +lean_inc_ref(x_4); +lean_inc(x_8); +x_14 = l_Lean_IR_EmitC_getDecl(x_8, x_4, x_5); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec_ref(x_14); +x_17 = l_List_forM___at___00Lean_IR_EmitC_emitFnDecls_spec__3___closed__1; +x_18 = l_Lean_IR_Decl_name(x_15); +lean_inc_ref(x_1); +x_19 = l_Lean_getExternNameFor(x_1, x_17, x_18); +if (lean_obj_tag(x_19) == 0) +{ +uint8_t x_20; +x_20 = l_Lean_NameSet_contains(x_2, x_8); +lean_dec(x_8); +if (x_20 == 0) +{ +uint8_t x_21; lean_object* x_22; +x_21 = 1; +lean_inc_ref(x_4); +x_22 = l_Lean_IR_EmitC_emitFnDecl(x_15, x_21, x_4, x_16); +lean_dec(x_15); +x_10 = x_22; +goto block_13; } -static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__6() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("#pragma GCC diagnostic ignored \"-Wunused-label\"", 47, 47); -return x_1; +uint8_t x_23; lean_object* x_24; +x_23 = 0; +lean_inc_ref(x_4); +x_24 = l_Lean_IR_EmitC_emitFnDecl(x_15, x_23, x_4, x_16); +lean_dec(x_15); +x_10 = x_24; +goto block_13; } } -static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__7() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("#pragma GCC diagnostic ignored \"-Wunused-but-set-variable\"", 58, 58); -return x_1; +lean_object* x_25; lean_object* x_26; +lean_dec(x_8); +x_25 = lean_ctor_get(x_19, 0); +lean_inc(x_25); +lean_dec_ref(x_19); +lean_inc_ref(x_4); +x_26 = l_Lean_IR_EmitC_emitExternDeclAux(x_15, x_25, x_4, x_16); +lean_dec(x_25); +lean_dec(x_15); +x_10 = x_26; +goto block_13; } } -static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__8() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("#endif", 6, 6); -return x_1; -} +uint8_t x_27; +lean_dec(x_9); +lean_dec(x_8); +lean_dec_ref(x_4); +lean_dec_ref(x_1); +x_27 = !lean_is_exclusive(x_14); +if (x_27 == 0) +{ +return x_14; } -static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__9() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("#ifdef __cplusplus", 18, 18); -return x_1; +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_14, 0); +x_29 = lean_ctor_get(x_14, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_14); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; } } -static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__10() { -_start: +block_13: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("extern \"C\" {", 12, 12); -return x_1; -} +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec_ref(x_10); +x_3 = x_9; +x_5 = x_11; +goto _start; } -static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__11() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_IR_EmitC_emitFileHeader___closed__8; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; +lean_dec(x_9); +lean_dec_ref(x_4); +lean_dec_ref(x_1); +return x_10; } } -static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__12() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_IR_EmitC_emitFileHeader___closed__11; -x_2 = l_Lean_IR_EmitC_emitFileHeader___closed__10; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; } } -static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__13() { +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFnDecls(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_IR_EmitC_emitFileHeader___closed__12; -x_2 = l_Lean_IR_EmitC_emitFileHeader___closed__9; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_3 = lean_ctor_get(x_1, 0); +lean_inc_ref(x_3); +lean_inc_ref(x_3); +x_4 = l_Lean_IR_getDecls(x_3); +x_5 = l_Lean_IR_EmitC_emitFnDecls___closed__0; +x_6 = l_List_foldl___at___00Lean_IR_EmitC_emitFnDecls_spec__0(x_5, x_4); +lean_inc_ref(x_3); +x_7 = l_List_foldl___at___00Lean_IR_EmitC_emitFnDecls_spec__1(x_3, x_5, x_4); +x_8 = lean_box(0); +x_9 = l_Std_DTreeMap_Internal_Impl_foldrM___at___00Lean_IR_EmitC_emitFnDecls_spec__2(x_8, x_7); +lean_dec(x_7); +x_10 = l_List_forM___at___00Lean_IR_EmitC_emitFnDecls_spec__3(x_3, x_6, x_9, x_1, x_2); +lean_dec(x_6); +return x_10; } } -static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__14() { +LEAN_EXPORT lean_object* l_List_foldl___at___00Lean_IR_EmitC_emitFnDecls_spec__0___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_IR_EmitC_emitFileHeader___closed__13; -x_2 = l_Lean_IR_EmitC_emitFileHeader___closed__8; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); +lean_object* x_3; +x_3 = l_List_foldl___at___00Lean_IR_EmitC_emitFnDecls_spec__0(x_1, x_2); +lean_dec(x_2); return x_3; } } -static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__15() { +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_foldrM___at___00Lean_IR_EmitC_emitFnDecls_spec__2___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_IR_EmitC_emitFileHeader___closed__14; -x_2 = l_Lean_IR_EmitC_emitFileHeader___closed__7; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); +lean_object* x_3; +x_3 = l_Std_DTreeMap_Internal_Impl_foldrM___at___00Lean_IR_EmitC_emitFnDecls_spec__2(x_1, x_2); +lean_dec(x_2); return x_3; } } -static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__16() { +LEAN_EXPORT lean_object* l_List_forM___at___00Lean_IR_EmitC_emitFnDecls_spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_IR_EmitC_emitFileHeader___closed__15; -x_2 = l_Lean_IR_EmitC_emitFileHeader___closed__6; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; +lean_object* x_6; +x_6 = l_List_forM___at___00Lean_IR_EmitC_emitFnDecls_spec__3(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_2); +return x_6; } } -static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__17() { +LEAN_EXPORT lean_object* l_List_forM___at___00Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0_spec__0___redArg(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_IR_EmitC_emitFileHeader___closed__16; -x_2 = l_Lean_IR_EmitC_emitFileHeader___closed__5; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; -} -} -static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__18() { -_start: +if (lean_obj_tag(x_1) == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_IR_EmitC_emitFileHeader___closed__17; -x_2 = l_Lean_IR_EmitC_emitFileHeader___closed__4; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; -} +lean_object* x_3; lean_object* x_4; +x_3 = lean_box(0); +x_4 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; } -static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__19() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_IR_EmitC_emitFileHeader___closed__18; -x_2 = l_Lean_IR_EmitC_emitFileHeader___closed__3; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; -} +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_5 = lean_ctor_get(x_1, 0); +x_6 = lean_ctor_get(x_1, 1); +x_7 = lean_string_append(x_2, x_5); +x_8 = l_Lean_IR_EmitC_emitLn___closed__0; +x_9 = lean_string_append(x_7, x_8); +x_1 = x_6; +x_2 = x_9; +goto _start; } -static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__20() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_IR_EmitC_emitFileHeader___closed__19; -x_2 = l_Lean_IR_EmitC_emitFileHeader___closed__2; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; } } -static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__21() { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_IR_EmitC_emitFileHeader___closed__20; -x_2 = l_Lean_IR_EmitC_emitFileHeader___closed__1; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; +lean_object* x_4; +x_4 = l_List_forM___at___00Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0_spec__0___redArg(x_1, x_3); +return x_4; } } -static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__22() { +LEAN_EXPORT lean_object* l_List_forM___at___00Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("// Lean compiler output", 23, 23); -return x_1; +lean_object* x_4; +x_4 = l_List_forM___at___00Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0_spec__0___redArg(x_1, x_3); +return x_4; } } -static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__23() { +static lean_object* _init_l_panic___at___00Lean_IR_EmitC_emitMainFn_spec__2___closed__0() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("// Module: ", 11, 11); +x_1 = l_Lean_instInhabitedConstantInfo_default; return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__24() { +LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_EmitC_emitMainFn_spec__2(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("// Imports:", 11, 11); -return x_1; +lean_object* x_2; lean_object* x_3; +x_2 = l_panic___at___00Lean_IR_EmitC_emitMainFn_spec__2___closed__0; +x_3 = lean_panic_fn(x_2, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFileHeader(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l_Option_instBEq_beq___at___00Lean_IR_EmitC_emitMainFn_spec__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_3; lean_object* x_12; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; -x_15 = lean_ctor_get(x_1, 0); -lean_inc_ref(x_15); -x_16 = lean_ctor_get(x_1, 1); -lean_inc(x_16); -lean_dec_ref(x_1); -x_17 = l_Lean_IR_EmitC_emitFileHeader___closed__22; -x_18 = lean_string_append(x_2, x_17); -x_19 = l_Lean_IR_EmitC_emitLn___closed__0; -x_20 = lean_string_append(x_18, x_19); -x_21 = l_Lean_IR_EmitC_emitFileHeader___closed__23; -x_22 = 1; -x_23 = l_Lean_Name_toStringWithToken___at___00Lean_Name_toString_spec__0(x_16, x_22); -x_24 = lean_string_append(x_21, x_23); -lean_dec_ref(x_23); -x_25 = lean_string_append(x_20, x_24); -lean_dec_ref(x_24); -x_26 = lean_string_append(x_25, x_19); -x_27 = l_Lean_IR_EmitC_emitFileHeader___closed__24; -x_28 = lean_string_append(x_26, x_27); -x_29 = l_Lean_Environment_imports(x_15); -lean_dec_ref(x_15); -x_30 = lean_unsigned_to_nat(0u); -x_31 = lean_array_get_size(x_29); -x_32 = lean_nat_dec_lt(x_30, x_31); -if (x_32 == 0) -{ -lean_dec_ref(x_29); -x_3 = x_28; -goto block_11; -} -else -{ -lean_object* x_33; uint8_t x_34; -x_33 = lean_box(0); -x_34 = lean_nat_dec_le(x_31, x_31); -if (x_34 == 0) +if (lean_obj_tag(x_1) == 0) { -if (x_32 == 0) +if (lean_obj_tag(x_2) == 0) { -lean_dec_ref(x_29); -x_3 = x_28; -goto block_11; +uint8_t x_3; +x_3 = 1; +return x_3; } else { -size_t x_35; size_t x_36; lean_object* x_37; -x_35 = 0; -x_36 = lean_usize_of_nat(x_31); -x_37 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg(x_29, x_35, x_36, x_33, x_28); -lean_dec_ref(x_29); -x_12 = x_37; -goto block_14; +uint8_t x_4; +x_4 = 0; +return x_4; } } else { -size_t x_38; size_t x_39; lean_object* x_40; -x_38 = 0; -x_39 = lean_usize_of_nat(x_31); -x_40 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg(x_29, x_38, x_39, x_33, x_28); -lean_dec_ref(x_29); -x_12 = x_40; -goto block_14; -} -} -block_11: -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_4 = l_Lean_IR_EmitC_emitLn___closed__0; -x_5 = lean_string_append(x_3, x_4); -x_6 = l_Lean_IR_EmitC_emitFileHeader___closed__0; -x_7 = lean_string_append(x_5, x_6); -x_8 = lean_string_append(x_7, x_4); -x_9 = l_Lean_IR_EmitC_emitFileHeader___closed__21; -x_10 = l_List_forM___at___00Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0_spec__0___redArg(x_9, x_8); -return x_10; -} -block_14: -{ -if (lean_obj_tag(x_12) == 0) +if (lean_obj_tag(x_2) == 0) { -lean_object* x_13; -x_13 = lean_ctor_get(x_12, 1); -lean_inc(x_13); -lean_dec_ref(x_12); -x_3 = x_13; -goto block_11; +uint8_t x_5; +x_5 = 0; +return x_5; } else { -return x_12; +lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_6 = lean_ctor_get(x_1, 0); +x_7 = lean_ctor_get(x_2, 0); +x_8 = lean_name_eq(x_6, x_7); +return x_8; } } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__0() { _start: { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_8 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_9 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0(x_1, x_7, x_8, x_4, x_5, x_6); -lean_dec_ref(x_5); -lean_dec_ref(x_1); -return x_9; +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" lean_dec_ref(res);", 20, 20); +return x_1; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__1() { _start: { -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_7 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_8 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg(x_1, x_6, x_7, x_4, x_5); -lean_dec_ref(x_1); -return x_8; +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" return ret;", 13, 13); +return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitFileFooter___redArg___closed__0() { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_IR_EmitC_emitFileHeader___closed__11; -x_2 = l_Lean_IR_EmitC_emitMainFn___closed__12; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("} else {", 8, 8); +return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitFileFooter___redArg___closed__1() { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__3() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_IR_EmitC_emitFileFooter___redArg___closed__0; -x_2 = l_Lean_IR_EmitC_emitFileHeader___closed__9; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" lean_io_result_show_error(res);", 33, 33); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFileFooter___redArg(lean_object* x_1) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__4() { _start: { -lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_IR_EmitC_emitFileFooter___redArg___closed__1; -x_3 = l_List_forM___at___00Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0_spec__0___redArg(x_2, x_1); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" return 1;", 11, 11); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFileFooter(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__5() { _start: { -lean_object* x_3; -x_3 = l_Lean_IR_EmitC_emitFileFooter___redArg(x_2); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_finalize_task_manager();", 29, 29); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFileFooter___boxed(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__6() { _start: { -lean_object* x_3; -x_3 = l_Lean_IR_EmitC_emitFileFooter(x_1, x_2); -lean_dec_ref(x_1); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" int ret = ", 12, 12); +return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_throwUnknownVar___redArg___closed__0() { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__7() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("unknown variable '", 18, 18); +x_1 = lean_mk_string_unchecked("UInt32", 6, 6); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_throwUnknownVar___redArg(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__8() { _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_Lean_IR_EmitC_throwUnknownVar___redArg___closed__0; -x_4 = l_Lean_IR_EmitC_argToCString___closed__0; -x_5 = l_Nat_reprFast(x_1); -x_6 = lean_string_append(x_4, x_5); -lean_dec_ref(x_5); -x_7 = lean_string_append(x_3, x_6); -lean_dec_ref(x_6); -x_8 = l_Lean_IR_EmitC_getDecl___closed__1; -x_9 = lean_string_append(x_7, x_8); -x_10 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_2); -return x_10; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_IR_EmitC_emitMainFn___closed__7; +x_2 = l_Lean_Name_mkStr1(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_throwUnknownVar(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__9() { _start: { -lean_object* x_5; -x_5 = l_Lean_IR_EmitC_throwUnknownVar___redArg(x_2, x_4); -return x_5; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_IR_EmitC_emitMainFn___closed__8; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_throwUnknownVar___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__10() { _start: { -lean_object* x_5; -x_5 = l_Lean_IR_EmitC_throwUnknownVar(x_1, x_2, x_3, x_4); -lean_dec_ref(x_3); -return x_5; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("0;", 2, 2); +return x_1; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0_spec__0___redArg(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__11() { _start: { -if (lean_obj_tag(x_2) == 0) -{ -lean_object* x_3; -x_3 = lean_box(0); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_unbox_uint32(lean_io_result_get_value(res));", 49, 49); +return x_1; } -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 = l_Lean_IR_instBEqJoinPointId_beq(x_4, x_1); -if (x_7 == 0) -{ -x_2 = x_6; -goto _start; } -else +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__12() { +_start: { -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_object* x_1; +x_1 = lean_mk_string_unchecked("Init.Data.Option.BasicAux", 25, 25); +return x_1; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__13() { _start: { -lean_object* x_4; -x_4 = l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0_spec__0___redArg(x_2, x_3); -return x_4; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Option.get!", 11, 11); +return x_1; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0___redArg(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__14() { _start: { -lean_object* x_3; lean_object* x_4; uint64_t x_5; uint64_t x_6; uint64_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; size_t x_12; size_t x_13; size_t x_14; size_t x_15; size_t x_16; lean_object* x_17; lean_object* x_18; -x_3 = lean_ctor_get(x_1, 1); -x_4 = lean_array_get_size(x_3); -x_5 = l_Lean_IR_instHashableJoinPointId_hash(x_2); -x_6 = 32; -x_7 = lean_uint64_shift_right(x_5, x_6); -x_8 = lean_uint64_xor(x_5, x_7); -x_9 = 16; -x_10 = lean_uint64_shift_right(x_8, x_9); -x_11 = lean_uint64_xor(x_8, x_10); -x_12 = lean_uint64_to_usize(x_11); -x_13 = lean_usize_of_nat(x_4); -x_14 = 1; -x_15 = lean_usize_sub(x_13, x_14); -x_16 = lean_usize_land(x_12, x_15); -x_17 = lean_array_uget(x_3, x_16); -x_18 = l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0_spec__0___redArg(x_2, x_17); -lean_dec(x_17); -return x_18; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("value is none", 13, 13); +return x_1; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__15() { _start: { -lean_object* x_4; -x_4 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0___redArg(x_2, x_3); -return x_4; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_IR_EmitC_emitMainFn___closed__14; +x_2 = lean_unsigned_to_nat(14u); +x_3 = lean_unsigned_to_nat(22u); +x_4 = l_Lean_IR_EmitC_emitMainFn___closed__13; +x_5 = l_Lean_IR_EmitC_emitMainFn___closed__12; +x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); +return x_6; } } -static lean_object* _init_l_Lean_IR_EmitC_getJPParams___closed__0() { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__16() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("unknown join point", 18, 18); +x_1 = lean_mk_string_unchecked("lean_set_panic_messages(false);", 31, 31); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_getJPParams(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__17() { _start: { -lean_object* x_4; lean_object* x_5; -x_4 = lean_ctor_get(x_2, 2); -x_5 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0___redArg(x_4, x_1); -if (lean_obj_tag(x_5) == 0) -{ -lean_object* x_6; lean_object* x_7; -x_6 = l_Lean_IR_EmitC_getJPParams___closed__0; -x_7 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_7, 0, x_6); -lean_ctor_set(x_7, 1, x_3); -return x_7; -} -else -{ -lean_object* x_8; lean_object* x_9; -x_8 = lean_ctor_get(x_5, 0); -lean_inc(x_8); -lean_dec_ref(x_5); -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_8); -lean_ctor_set(x_9, 1, x_3); -return x_9; -} +lean_object* x_1; +x_1 = lean_mk_string_unchecked("res = ", 6, 6); +return x_1; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__18() { _start: { -lean_object* x_4; -x_4 = l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0_spec__0(x_1, x_2, x_3); -lean_dec(x_3); -lean_dec(x_2); -return x_4; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("(1 /* builtin */);", 18, 18); +return x_1; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__19() { _start: { -lean_object* x_4; -x_4 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0(x_1, x_2, x_3); -lean_dec(x_3); -lean_dec_ref(x_2); -return x_4; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_set_panic_messages(true);", 30, 30); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_getJPParams___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__20() { _start: { -lean_object* x_4; -x_4 = l_Lean_IR_EmitC_getJPParams(x_1, x_2, x_3); -lean_dec_ref(x_2); -lean_dec(x_1); -return x_4; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_io_mark_end_initialization();", 34, 34); +return x_1; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__21() { _start: { -lean_object* x_3; -x_3 = l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0_spec__0___redArg(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("if (lean_io_result_is_ok(res)) {", 32, 32); +return x_1; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__22() { _start: { -lean_object* x_3; -x_3 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0___redArg(x_1, x_2); -lean_dec(x_2); -lean_dec_ref(x_1); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_dec_ref(res);", 18, 18); +return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_declareVar___redArg___closed__0() { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__23() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("; ", 2, 2); +x_1 = lean_mk_string_unchecked("lean_init_task_manager();", 25, 25); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_declareVar___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__24() { _start: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_4 = l_Lean_IR_EmitC_toCType(x_2); -x_5 = lean_string_append(x_3, x_4); -lean_dec_ref(x_4); -x_6 = l_Lean_IR_EmitC_emitFnDeclAux___closed__4; -x_7 = lean_string_append(x_5, x_6); -x_8 = l_Lean_IR_EmitC_argToCString___closed__0; -x_9 = l_Nat_reprFast(x_1); -x_10 = lean_string_append(x_8, x_9); -lean_dec_ref(x_9); -x_11 = lean_string_append(x_7, x_10); -lean_dec_ref(x_10); -x_12 = l_Lean_IR_EmitC_declareVar___redArg___closed__0; -x_13 = lean_box(0); -x_14 = lean_string_append(x_11, x_12); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_13); -lean_ctor_set(x_15, 1, x_14); -return x_15; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_IR_EmitC_emitMainFn___closed__23; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_declareVar(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__25() { _start: { -lean_object* x_5; -x_5 = l_Lean_IR_EmitC_declareVar___redArg(x_1, x_2, x_4); -return x_5; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_EmitC_emitMainFn___closed__24; +x_2 = l_Lean_IR_EmitC_emitMainFn___closed__22; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_declareVar___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__26() { _start: { -lean_object* x_5; -x_5 = l_Lean_IR_EmitC_declareVar(x_1, x_2, x_3, x_4); -lean_dec_ref(x_3); -lean_dec(x_2); -return x_5; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_EmitC_emitMainFn___closed__25; +x_2 = l_Lean_IR_EmitC_emitMainFn___closed__21; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_declareVar___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__27() { _start: { -lean_object* x_4; -x_4 = l_Lean_IR_EmitC_declareVar___redArg(x_1, x_2, x_3); -lean_dec(x_2); -return x_4; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_EmitC_emitMainFn___closed__26; +x_2 = l_Lean_IR_EmitC_emitMainFn___closed__20; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_declareParams_spec__0___redArg(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__28() { _start: { -uint8_t x_6; -x_6 = lean_usize_dec_eq(x_2, x_3); -if (x_6 == 0) -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_7 = lean_array_uget(x_1, x_2); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_7, 1); -lean_inc(x_9); -lean_dec(x_7); -x_10 = l_Lean_IR_EmitC_declareVar___redArg(x_8, x_9, x_5); -lean_dec(x_9); -if (lean_obj_tag(x_10) == 0) -{ -lean_object* x_11; lean_object* x_12; size_t x_13; size_t x_14; -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); -lean_dec_ref(x_10); -x_13 = 1; -x_14 = lean_usize_add(x_2, x_13); -x_2 = x_14; -x_4 = x_11; -x_5 = x_12; -goto _start; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("res = _lean_main();", 19, 19); +return x_1; } -else +} +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__29() { +_start: { -return x_10; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("in = lean_box(0);", 17, 17); +return x_1; } } -else +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__30() { +_start: { -lean_object* x_16; -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_4); -lean_ctor_set(x_16, 1, x_5); -return x_16; -} +lean_object* x_1; +x_1 = lean_mk_string_unchecked("int i = argc;", 13, 13); +return x_1; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_declareParams_spec__0(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__31() { _start: { -lean_object* x_7; -x_7 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_declareParams_spec__0___redArg(x_1, x_2, x_3, x_4, x_6); -return x_7; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("while (i > 1) {", 15, 15); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_declareParams(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__32() { _start: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_4 = lean_unsigned_to_nat(0u); -x_5 = lean_array_get_size(x_1); -x_6 = lean_box(0); -x_7 = lean_nat_dec_lt(x_4, x_5); -if (x_7 == 0) -{ -lean_object* x_8; -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_6); -lean_ctor_set(x_8, 1, x_3); -return x_8; +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" lean_object* n;", 16, 16); +return x_1; } -else -{ -uint8_t x_9; -x_9 = lean_nat_dec_le(x_5, x_5); -if (x_9 == 0) -{ -if (x_7 == 0) -{ -lean_object* x_10; -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_6); -lean_ctor_set(x_10, 1, x_3); -return x_10; } -else +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__33() { +_start: { -size_t x_11; size_t x_12; lean_object* x_13; -x_11 = 0; -x_12 = lean_usize_of_nat(x_5); -x_13 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_declareParams_spec__0___redArg(x_1, x_11, x_12, x_6, x_3); -return x_13; +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" i--;", 5, 5); +return x_1; } } -else +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__34() { +_start: { -size_t x_14; size_t x_15; lean_object* x_16; -x_14 = 0; -x_15 = lean_usize_of_nat(x_5); -x_16 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_declareParams_spec__0___redArg(x_1, x_14, x_15, x_6, x_3); -return x_16; +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" n = lean_alloc_ctor(1,2,0); lean_ctor_set(n, 0, lean_mk_string(argv[i])); lean_ctor_set(n, 1, in);", 99, 99); +return x_1; } } +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__35() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" in = n;", 8, 8); +return x_1; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_declareParams_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__36() { _start: { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_8 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_9 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_declareParams_spec__0(x_1, x_7, x_8, x_4, x_5, x_6); -lean_dec_ref(x_5); -lean_dec_ref(x_1); -return x_9; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_declareParams_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__37() { _start: { -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_7 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_8 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_declareParams_spec__0___redArg(x_1, x_6, x_7, x_4, x_5); -lean_dec_ref(x_1); -return x_8; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_EmitC_emitMainFn___closed__36; +x_2 = l_Lean_IR_EmitC_emitMainFn___closed__35; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_declareParams___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__38() { _start: { -lean_object* x_4; -x_4 = l_Lean_IR_EmitC_declareParams(x_1, x_2, x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -return x_4; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_EmitC_emitMainFn___closed__37; +x_2 = l_Lean_IR_EmitC_emitMainFn___closed__34; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_declareVars(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__39() { _start: { -switch (lean_obj_tag(x_1)) { -case 0: -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_5 = lean_ctor_get(x_1, 0); -lean_inc(x_5); -x_6 = lean_ctor_get(x_1, 1); -lean_inc(x_6); -x_7 = lean_ctor_get(x_1, 3); -lean_inc(x_7); -x_8 = lean_ctor_get(x_3, 3); -x_9 = l_Lean_IR_isTailCallTo(x_8, x_1); -lean_dec_ref(x_1); -if (x_9 == 0) -{ -lean_object* x_10; -x_10 = l_Lean_IR_EmitC_declareVar___redArg(x_5, x_6, x_4); -lean_dec(x_6); -if (lean_obj_tag(x_10) == 0) -{ -lean_object* x_11; uint8_t x_12; -x_11 = lean_ctor_get(x_10, 1); -lean_inc(x_11); -lean_dec_ref(x_10); -x_12 = 1; -x_1 = x_7; -x_2 = x_12; -x_4 = x_11; -goto _start; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_EmitC_emitMainFn___closed__38; +x_2 = l_Lean_IR_EmitC_emitMainFn___closed__33; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } -else -{ -uint8_t x_14; -lean_dec(x_7); -x_14 = !lean_is_exclusive(x_10); -if (x_14 == 0) -{ -return x_10; } -else +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__40() { +_start: { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_10, 0); -x_16 = lean_ctor_get(x_10, 1); -lean_inc(x_16); -lean_inc(x_15); -lean_dec(x_10); -x_17 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_16); -return x_17; -} +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_EmitC_emitMainFn___closed__39; +x_2 = l_Lean_IR_EmitC_emitMainFn___closed__32; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -else +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__41() { +_start: { -lean_object* x_18; lean_object* x_19; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -x_18 = lean_box(x_2); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_4); -return x_19; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_EmitC_emitMainFn___closed__40; +x_2 = l_Lean_IR_EmitC_emitMainFn___closed__31; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -case 1: +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__42() { +_start: { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_1, 1); -lean_inc_ref(x_20); -x_21 = lean_ctor_get(x_1, 3); -lean_inc(x_21); -lean_dec_ref(x_1); -x_22 = l_Lean_IR_EmitC_declareParams(x_20, x_3, x_4); -if (lean_obj_tag(x_22) == 0) -{ -if (x_2 == 0) -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; -x_23 = lean_ctor_get(x_22, 1); -lean_inc(x_23); -lean_dec_ref(x_22); -x_24 = lean_unsigned_to_nat(0u); -x_25 = lean_array_get_size(x_20); -lean_dec_ref(x_20); -x_26 = lean_nat_dec_lt(x_24, x_25); -x_1 = x_21; -x_2 = x_26; -x_4 = x_23; -goto _start; -} -else -{ -lean_object* x_28; -lean_dec_ref(x_20); -x_28 = lean_ctor_get(x_22, 1); -lean_inc(x_28); -lean_dec_ref(x_22); -x_1 = x_21; -x_4 = x_28; -goto _start; -} -} -else -{ -uint8_t x_30; -lean_dec(x_21); -lean_dec_ref(x_20); -x_30 = !lean_is_exclusive(x_22); -if (x_30 == 0) -{ -return x_22; -} -else -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_22, 0); -x_32 = lean_ctor_get(x_22, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_22); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; -} -} -} -default: -{ -uint8_t x_34; -x_34 = l_Lean_IR_FnBody_isTerminal(x_1); -if (x_34 == 0) -{ -lean_object* x_35; -x_35 = l_Lean_IR_FnBody_body(x_1); -lean_dec(x_1); -x_1 = x_35; -goto _start; -} -else -{ -lean_object* x_37; lean_object* x_38; -lean_dec(x_1); -x_37 = lean_box(x_2); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_38, 1, x_4); -return x_38; -} -} -} +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_EmitC_emitMainFn___closed__41; +x_2 = l_Lean_IR_EmitC_emitMainFn___closed__30; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_declareVars___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__43() { _start: { -uint8_t x_5; lean_object* x_6; -x_5 = lean_unbox(x_2); -x_6 = l_Lean_IR_EmitC_declareVars(x_1, x_5, x_3, x_4); -lean_dec_ref(x_3); -return x_6; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_EmitC_emitMainFn___closed__42; +x_2 = l_Lean_IR_EmitC_emitMainFn___closed__29; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -static lean_object* _init_l_Lean_IR_EmitC_emitTag___redArg___closed__0() { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__44() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_obj_tag(", 13, 13); +x_1 = lean_mk_string_unchecked("res = _lean_main(in);", 21, 21); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitTag___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__45() { _start: { -uint8_t x_4; -x_4 = l_Lean_IR_IRType_isObj(x_2); -if (x_4 == 0) -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_5 = lean_box(0); -x_6 = l_Lean_IR_EmitC_argToCString___closed__0; -x_7 = l_Nat_reprFast(x_1); -x_8 = lean_string_append(x_6, x_7); -lean_dec_ref(x_7); -x_9 = lean_string_append(x_3, x_8); -lean_dec_ref(x_8); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_5); -lean_ctor_set(x_10, 1, x_9); -return x_10; -} -else -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_11 = l_Lean_IR_EmitC_emitTag___redArg___closed__0; -x_12 = lean_string_append(x_3, x_11); -x_13 = l_Lean_IR_EmitC_argToCString___closed__0; -x_14 = l_Nat_reprFast(x_1); -x_15 = lean_string_append(x_13, x_14); -lean_dec_ref(x_14); -x_16 = lean_string_append(x_12, x_15); -lean_dec_ref(x_15); -x_17 = l_Lean_IR_EmitC_emitFnDeclAux___closed__1; -x_18 = lean_box(0); -x_19 = lean_string_append(x_16, x_17); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_18); -lean_ctor_set(x_20, 1, x_19); -return x_20; -} +lean_object* x_1; +x_1 = lean_mk_string_unchecked("\n #if defined(WIN32) || defined(_WIN32)\n #include \n #endif\n\n int main(int argc, char ** argv) {\n #if defined(WIN32) || defined(_WIN32)\n SetErrorMode(SEM_FAILCRITICALERRORS);\n SetConsoleOutputCP(CP_UTF8);\n #endif\n lean_object* in; lean_object* res;", 267, 267); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitTag(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__46() { _start: { -lean_object* x_5; -x_5 = l_Lean_IR_EmitC_emitTag___redArg(x_1, x_2, x_4); -return x_5; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("argv = lean_setup_args(argc, argv);", 35, 35); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitTag___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__47() { _start: { -lean_object* x_5; -x_5 = l_Lean_IR_EmitC_emitTag(x_1, x_2, x_3, x_4); -lean_dec_ref(x_3); -lean_dec(x_2); -return x_5; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_initialize_runtime_module();", 33, 33); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitTag___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__48() { _start: { -lean_object* x_4; -x_4 = l_Lean_IR_EmitC_emitTag___redArg(x_1, x_2, x_3); -lean_dec(x_2); -return x_4; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_initialize();", 18, 18); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_isIf(lean_object* x_1) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__49() { _start: { -lean_object* x_2; lean_object* x_3; uint8_t x_4; -x_2 = lean_array_get_size(x_1); -x_3 = lean_unsigned_to_nat(2u); -x_4 = lean_nat_dec_eq(x_2, x_3); -if (x_4 == 0) -{ -lean_object* x_5; -x_5 = lean_box(0); -return x_5; -} -else -{ -lean_object* x_6; lean_object* x_7; -x_6 = lean_unsigned_to_nat(0u); -x_7 = lean_array_fget(x_1, x_6); -if (lean_obj_tag(x_7) == 0) -{ -uint8_t x_8; -x_8 = !lean_is_exclusive(x_7); -if (x_8 == 0) -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; 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_7, 0); -x_10 = lean_ctor_get(x_7, 1); -x_11 = lean_ctor_get(x_9, 1); -lean_inc(x_11); -lean_dec_ref(x_9); -x_12 = lean_unsigned_to_nat(1u); -x_13 = lean_array_fget_borrowed(x_1, x_12); -x_14 = l_Lean_IR_Alt_body(x_13); -lean_ctor_set(x_7, 1, x_14); -lean_ctor_set(x_7, 0, x_10); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_11); -lean_ctor_set(x_15, 1, x_7); -x_16 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_16, 0, x_15); -return x_16; -} -else -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_17 = lean_ctor_get(x_7, 0); -x_18 = lean_ctor_get(x_7, 1); -lean_inc(x_18); -lean_inc(x_17); -lean_dec(x_7); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); -lean_dec_ref(x_17); -x_20 = lean_unsigned_to_nat(1u); -x_21 = lean_array_fget_borrowed(x_1, x_20); -x_22 = l_Lean_IR_Alt_body(x_21); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_18); -lean_ctor_set(x_23, 1, x_22); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_19); -lean_ctor_set(x_24, 1, x_23); -x_25 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_25, 0, x_24); -return x_25; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid main function, incorrect arity when generating code", 59, 59); +return x_1; } } -else +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__50() { +_start: { -lean_object* x_26; -lean_dec(x_7); -x_26 = lean_box(0); -return x_26; -} -} +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_isIf___boxed(lean_object* x_1) { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__51() { _start: { -lean_object* x_2; -x_2 = l_Lean_IR_EmitC_isIf(x_1); -lean_dec_ref(x_1); +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_IR_EmitC_emitMainFn___closed__50; +x_2 = l_Lean_Name_mkStr1(x_1); return x_2; } } -static lean_object* _init_l_Lean_IR_EmitC_emitInc___redArg___closed__0() { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__52() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked(");", 2, 2); +x_1 = lean_mk_string_unchecked("char ** lean_setup_args(int argc, char ** argv);", 48, 48); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitInc___redArg___closed__1() { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__53() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_inc_ref_n", 14, 14); +x_1 = lean_mk_string_unchecked("void lean_initialize_runtime_module();", 38, 38); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitInc___redArg___closed__2() { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__54() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_inc_ref", 12, 12); +x_1 = lean_mk_string_unchecked("void lean_initialize();", 23, 23); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitInc___redArg___closed__3() { +static lean_object* _init_l_Lean_IR_EmitC_emitMainFn___closed__55() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_inc_n", 10, 10); +x_1 = lean_mk_string_unchecked("function declaration expected", 29, 29); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitInc___redArg___closed__4() { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitMainFn(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_inc", 8, 8); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitInc___redArg(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4) { -_start: +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_75; +x_59 = l_Lean_IR_EmitC_toCName___closed__1; +lean_inc_ref(x_1); +x_75 = l_Lean_IR_EmitC_getDecl(x_59, x_1, x_2); +if (lean_obj_tag(x_75) == 0) { -lean_object* x_5; lean_object* x_13; -if (x_3 == 0) +lean_object* x_76; +x_76 = lean_ctor_get(x_75, 0); +lean_inc(x_76); +if (lean_obj_tag(x_76) == 0) { -lean_object* x_28; uint8_t x_29; -x_28 = lean_unsigned_to_nat(1u); -x_29 = lean_nat_dec_eq(x_2, x_28); -if (x_29 == 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; uint8_t x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; uint8_t x_141; lean_object* x_158; lean_object* x_159; uint8_t x_160; +x_77 = lean_ctor_get(x_75, 1); +lean_inc(x_77); +if (lean_is_exclusive(x_75)) { + lean_ctor_release(x_75, 0); + lean_ctor_release(x_75, 1); + x_78 = x_75; +} else { + lean_dec_ref(x_75); + x_78 = lean_box(0); +} +x_79 = lean_ctor_get(x_76, 1); +lean_inc_ref(x_79); +lean_dec_ref(x_76); +x_158 = lean_array_get_size(x_79); +x_159 = lean_unsigned_to_nat(2u); +x_160 = lean_nat_dec_eq(x_158, x_159); +if (x_160 == 0) { -lean_object* x_30; -x_30 = l_Lean_IR_EmitC_emitInc___redArg___closed__1; -x_13 = x_30; -goto block_27; +lean_object* x_161; uint8_t x_162; +x_161 = lean_unsigned_to_nat(1u); +x_162 = lean_nat_dec_eq(x_158, x_161); +x_141 = x_162; +goto block_157; } else { -lean_object* x_31; -x_31 = l_Lean_IR_EmitC_emitInc___redArg___closed__2; -x_13 = x_31; -goto block_27; -} +x_141 = x_160; +goto block_157; } -else +block_122: { -lean_object* x_32; uint8_t x_33; -x_32 = lean_unsigned_to_nat(1u); -x_33 = lean_nat_dec_eq(x_2, x_32); -if (x_33 == 0) +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; uint8_t x_112; +x_83 = lean_ctor_get(x_81, 0); +x_84 = lean_ctor_get(x_81, 1); +x_85 = l_Lean_IR_EmitC_getModInitFn___closed__0; +x_86 = lean_ctor_get(x_85, 0); +lean_inc_ref(x_86); +x_87 = lean_ctor_get(x_86, 2); +lean_inc(x_87); +lean_dec_ref(x_86); +x_88 = l_Lean_IR_EmitC_emitMainFn___closed__16; +x_89 = lean_string_append(x_82, x_88); +x_90 = l_Lean_IR_EmitC_emitLn___closed__0; +x_91 = lean_string_append(x_89, x_90); +x_92 = lean_box(0); +x_93 = lean_box(0); +lean_inc_ref(x_83); +x_94 = l_Lean_PersistentEnvExtension_getState___redArg(x_92, x_85, x_83, x_87, x_93); +lean_dec(x_87); +lean_inc(x_84); +x_95 = l_Lean_mkModuleInitializationFunctionName(x_84, x_94); +lean_dec(x_94); +x_96 = l_Lean_IR_EmitC_emitMainFn___closed__17; +x_97 = lean_string_append(x_96, x_95); +lean_dec_ref(x_95); +x_98 = l_Lean_IR_EmitC_emitMainFn___closed__18; +x_99 = lean_string_append(x_97, x_98); +x_100 = lean_string_append(x_91, x_99); +lean_dec_ref(x_99); +x_101 = lean_string_append(x_100, x_90); +x_102 = l_Lean_IR_EmitC_emitMainFn___closed__19; +x_103 = lean_string_append(x_101, x_102); +x_104 = lean_string_append(x_103, x_90); +x_105 = l_Lean_IR_EmitC_emitMainFn___closed__21; +x_106 = lean_box(0); +x_107 = l_Lean_IR_EmitC_emitMainFn___closed__27; +x_108 = l_List_forM___at___00Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0_spec__0___redArg(x_107, x_104); +x_109 = lean_ctor_get(x_108, 1); +lean_inc(x_109); +lean_dec_ref(x_108); +x_110 = lean_array_get_size(x_79); +lean_dec_ref(x_79); +x_111 = lean_unsigned_to_nat(2u); +x_112 = lean_nat_dec_eq(x_110, x_111); +if (x_112 == 0) { -lean_object* x_34; -x_34 = l_Lean_IR_EmitC_emitInc___redArg___closed__3; -x_13 = x_34; -goto block_27; +lean_object* x_113; lean_object* x_114; lean_object* x_115; +x_113 = l_Lean_IR_EmitC_emitMainFn___closed__28; +x_114 = lean_string_append(x_109, x_113); +x_115 = lean_string_append(x_114, x_90); +x_60 = x_106; +x_61 = x_105; +x_62 = x_80; +x_63 = x_90; +x_64 = x_81; +x_65 = x_115; +goto block_74; } else { -lean_object* x_35; -x_35 = l_Lean_IR_EmitC_emitInc___redArg___closed__4; -x_13 = x_35; -goto block_27; -} +lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; +x_116 = l_Lean_IR_EmitC_emitMainFn___closed__43; +x_117 = l_List_forM___at___00Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0_spec__0___redArg(x_116, x_109); +x_118 = lean_ctor_get(x_117, 1); +lean_inc(x_118); +lean_dec_ref(x_117); +x_119 = l_Lean_IR_EmitC_emitMainFn___closed__44; +x_120 = lean_string_append(x_118, x_119); +x_121 = lean_string_append(x_120, x_90); +x_60 = x_106; +x_61 = x_105; +x_62 = x_80; +x_63 = x_90; +x_64 = x_81; +x_65 = x_121; +goto block_74; } -block_12: -{ -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 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_7 = lean_string_append(x_5, x_6); -x_8 = l_Lean_IR_EmitC_emitLn___closed__0; -x_9 = lean_box(0); -x_10 = lean_string_append(x_7, x_8); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_9); -lean_ctor_set(x_11, 1, x_10); -return x_11; } -block_27: +block_140: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_14 = lean_string_append(x_4, x_13); -lean_dec_ref(x_13); -x_15 = l_Lean_IR_EmitC_emitFnDeclAux___closed__5; -x_16 = lean_string_append(x_14, x_15); -x_17 = l_Lean_IR_EmitC_argToCString___closed__0; -x_18 = l_Nat_reprFast(x_1); -x_19 = lean_string_append(x_17, x_18); -lean_dec_ref(x_18); -x_20 = lean_string_append(x_16, x_19); -lean_dec_ref(x_19); -x_21 = lean_unsigned_to_nat(1u); -x_22 = lean_nat_dec_eq(x_2, x_21); -if (x_22 == 0) +lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; +x_127 = l_Lean_IR_EmitC_emitMainFn___closed__45; +x_128 = lean_string_append(x_126, x_127); +x_129 = l_Lean_IR_EmitC_emitLn___closed__0; +x_130 = lean_string_append(x_128, x_129); +x_131 = l_Lean_IR_EmitC_emitMainFn___closed__46; +x_132 = lean_string_append(x_130, x_131); +x_133 = lean_string_append(x_132, x_129); +if (x_123 == 0) { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_23 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___closed__0; -x_24 = lean_string_append(x_20, x_23); -x_25 = l_Nat_reprFast(x_2); -x_26 = lean_string_append(x_24, x_25); -lean_dec_ref(x_25); -x_5 = x_26; -goto block_12; +lean_object* x_134; lean_object* x_135; lean_object* x_136; +x_134 = l_Lean_IR_EmitC_emitMainFn___closed__47; +x_135 = lean_string_append(x_133, x_134); +x_136 = lean_string_append(x_135, x_129); +x_80 = x_124; +x_81 = x_125; +x_82 = x_136; +goto block_122; } else { -lean_dec(x_2); -x_5 = x_20; -goto block_12; -} -} +lean_object* x_137; lean_object* x_138; lean_object* x_139; +x_137 = l_Lean_IR_EmitC_emitMainFn___closed__48; +x_138 = lean_string_append(x_133, x_137); +x_139 = lean_string_append(x_138, x_129); +x_80 = x_124; +x_81 = x_125; +x_82 = x_139; +goto block_122; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitInc(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5) { -_start: +block_157: { -lean_object* x_6; -x_6 = l_Lean_IR_EmitC_emitInc___redArg(x_1, x_2, x_3, x_5); -return x_6; -} -} -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitInc___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +if (x_141 == 0) { -uint8_t x_6; lean_object* x_7; -x_6 = lean_unbox(x_3); -x_7 = l_Lean_IR_EmitC_emitInc(x_1, x_2, x_6, x_4, x_5); -lean_dec_ref(x_4); -return x_7; +lean_object* x_142; lean_object* x_143; +lean_dec_ref(x_79); +lean_dec_ref(x_1); +x_142 = l_Lean_IR_EmitC_emitMainFn___closed__49; +if (lean_is_scalar(x_78)) { + x_143 = lean_alloc_ctor(1, 2, 0); +} else { + x_143 = x_78; + lean_ctor_set_tag(x_143, 1); } +lean_ctor_set(x_143, 0, x_142); +lean_ctor_set(x_143, 1, x_77); +return x_143; } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitInc___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +else { -uint8_t x_5; lean_object* x_6; -x_5 = lean_unbox(x_3); -x_6 = l_Lean_IR_EmitC_emitInc___redArg(x_1, x_2, x_5, x_4); -return x_6; -} +lean_object* x_144; lean_object* x_145; uint8_t x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; +lean_dec(x_78); +x_144 = lean_ctor_get(x_1, 0); +lean_inc_ref(x_144); +x_145 = l_Lean_IR_EmitC_emitMainFn___closed__51; +x_146 = l_Lean_IR_usesModuleFrom(x_144, x_145); +x_147 = l_Lean_IR_EmitC_emitMainFn___closed__52; +x_148 = lean_string_append(x_77, x_147); +x_149 = l_Lean_IR_EmitC_emitLn___closed__0; +x_150 = lean_string_append(x_148, x_149); +if (x_146 == 0) +{ +lean_object* x_151; lean_object* x_152; lean_object* x_153; +x_151 = l_Lean_IR_EmitC_emitMainFn___closed__53; +x_152 = lean_string_append(x_150, x_151); +x_153 = lean_string_append(x_152, x_149); +x_123 = x_146; +x_124 = x_144; +x_125 = x_1; +x_126 = x_153; +goto block_140; } -static lean_object* _init_l_Lean_IR_EmitC_emitDec___redArg___closed__0() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_dec_ref", 12, 12); -return x_1; +lean_object* x_154; lean_object* x_155; lean_object* x_156; +x_154 = l_Lean_IR_EmitC_emitMainFn___closed__54; +x_155 = lean_string_append(x_150, x_154); +x_156 = lean_string_append(x_155, x_149); +x_123 = x_146; +x_124 = x_144; +x_125 = x_1; +x_126 = x_156; +goto block_140; } } -static lean_object* _init_l_Lean_IR_EmitC_emitDec___redArg___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_dec", 8, 8); -return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitDec___redArg(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4) { -_start: +else { -lean_object* x_5; lean_object* x_13; -if (x_3 == 0) +uint8_t x_163; +lean_dec(x_76); +lean_dec_ref(x_1); +x_163 = !lean_is_exclusive(x_75); +if (x_163 == 0) { -lean_object* x_28; -x_28 = l_Lean_IR_EmitC_emitDec___redArg___closed__0; -x_13 = x_28; -goto block_27; +lean_object* x_164; lean_object* x_165; +x_164 = lean_ctor_get(x_75, 0); +lean_dec(x_164); +x_165 = l_Lean_IR_EmitC_emitMainFn___closed__55; +lean_ctor_set_tag(x_75, 1); +lean_ctor_set(x_75, 0, x_165); +return x_75; } else { -lean_object* x_29; -x_29 = l_Lean_IR_EmitC_emitDec___redArg___closed__1; -x_13 = x_29; -goto block_27; +lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_166 = lean_ctor_get(x_75, 1); +lean_inc(x_166); +lean_dec(x_75); +x_167 = l_Lean_IR_EmitC_emitMainFn___closed__55; +x_168 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_168, 0, x_167); +lean_ctor_set(x_168, 1, x_166); +return x_168; } -block_12: -{ -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 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_7 = lean_string_append(x_5, x_6); -x_8 = l_Lean_IR_EmitC_emitLn___closed__0; -x_9 = lean_box(0); -x_10 = lean_string_append(x_7, x_8); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_9); -lean_ctor_set(x_11, 1, x_10); -return x_11; } -block_27: +} +else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_14 = lean_string_append(x_4, x_13); -lean_dec_ref(x_13); -x_15 = l_Lean_IR_EmitC_emitFnDeclAux___closed__5; -x_16 = lean_string_append(x_14, x_15); -x_17 = l_Lean_IR_EmitC_argToCString___closed__0; -x_18 = l_Nat_reprFast(x_1); -x_19 = lean_string_append(x_17, x_18); -lean_dec_ref(x_18); -x_20 = lean_string_append(x_16, x_19); -lean_dec_ref(x_19); -x_21 = lean_unsigned_to_nat(1u); -x_22 = lean_nat_dec_eq(x_2, x_21); -if (x_22 == 0) +uint8_t x_169; +lean_dec_ref(x_1); +x_169 = !lean_is_exclusive(x_75); +if (x_169 == 0) { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_23 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___closed__0; -x_24 = lean_string_append(x_20, x_23); -x_25 = l_Nat_reprFast(x_2); -x_26 = lean_string_append(x_24, x_25); -lean_dec_ref(x_25); -x_5 = x_26; -goto block_12; +return x_75; } else { -lean_dec(x_2); -x_5 = x_20; -goto block_12; +lean_object* x_170; lean_object* x_171; lean_object* x_172; +x_170 = lean_ctor_get(x_75, 0); +x_171 = lean_ctor_get(x_75, 1); +lean_inc(x_171); +lean_inc(x_170); +lean_dec(x_75); +x_172 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_172, 0, x_170); +lean_ctor_set(x_172, 1, x_171); +return x_172; +} } +block_40: +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; +lean_dec_ref(x_6); +x_12 = lean_string_append(x_7, x_11); +lean_dec_ref(x_11); +x_13 = l_Lean_IR_EmitC_emitMainFn___closed__0; +x_14 = l_Lean_IR_EmitC_emitMainFn___closed__1; +x_15 = l_Lean_IR_EmitC_emitMainFn___closed__2; +x_16 = l_Lean_IR_EmitC_emitMainFn___closed__3; +x_17 = l_Lean_IR_EmitC_emitMainFn___closed__4; +lean_inc_ref(x_10); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_10); +lean_ctor_set(x_18, 1, x_4); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_13); +lean_ctor_set(x_20, 1, x_19); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_16); +lean_ctor_set(x_21, 1, x_20); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_15); +lean_ctor_set(x_22, 1, x_21); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_14); +lean_ctor_set(x_23, 1, x_22); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_13); +lean_ctor_set(x_24, 1, x_23); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_12); +lean_ctor_set(x_25, 1, x_24); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_5); +lean_ctor_set(x_26, 1, x_25); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_9); +lean_ctor_set(x_27, 1, x_26); +x_28 = l_List_forM___at___00Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0_spec__0___redArg(x_27, x_3); +lean_dec_ref(x_27); +x_29 = !lean_is_exclusive(x_28); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_30 = lean_ctor_get(x_28, 1); +x_31 = lean_ctor_get(x_28, 0); +lean_dec(x_31); +x_32 = lean_string_append(x_30, x_10); +lean_dec_ref(x_10); +x_33 = lean_box(0); +x_34 = lean_string_append(x_32, x_8); +lean_dec_ref(x_8); +lean_ctor_set(x_28, 1, x_34); +lean_ctor_set(x_28, 0, x_33); +return x_28; } +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_35 = lean_ctor_get(x_28, 1); +lean_inc(x_35); +lean_dec(x_28); +x_36 = lean_string_append(x_35, x_10); +lean_dec_ref(x_10); +x_37 = lean_box(0); +x_38 = lean_string_append(x_36, x_8); +lean_dec_ref(x_8); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitDec(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5) { -_start: +block_58: { -lean_object* x_6; -x_6 = l_Lean_IR_EmitC_emitDec___redArg(x_1, x_2, x_3, x_5); -return x_6; +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_48 = l_Lean_ConstantInfo_type(x_47); +lean_dec_ref(x_47); +x_49 = l_Lean_Expr_getForallBody(x_48); +lean_dec_ref(x_48); +x_50 = l_Lean_Expr_appArg_x21(x_49); +lean_dec_ref(x_49); +x_51 = l_Lean_IR_EmitC_emitMainFn___closed__5; +x_52 = l_Lean_IR_EmitC_emitMainFn___closed__6; +x_53 = l_Lean_Expr_constName_x3f(x_50); +lean_dec_ref(x_50); +x_54 = l_Lean_IR_EmitC_emitMainFn___closed__9; +x_55 = l_Option_instBEq_beq___at___00Lean_IR_EmitC_emitMainFn_spec__1(x_53, x_54); +lean_dec(x_53); +if (x_55 == 0) +{ +lean_object* x_56; +x_56 = l_Lean_IR_EmitC_emitMainFn___closed__10; +x_3 = x_41; +x_4 = x_42; +x_5 = x_43; +x_6 = x_44; +x_7 = x_52; +x_8 = x_45; +x_9 = x_51; +x_10 = x_46; +x_11 = x_56; +goto block_40; +} +else +{ +lean_object* x_57; +x_57 = l_Lean_IR_EmitC_emitMainFn___closed__11; +x_3 = x_41; +x_4 = x_42; +x_5 = x_43; +x_6 = x_44; +x_7 = x_52; +x_8 = x_45; +x_9 = x_51; +x_10 = x_46; +x_11 = x_57; +goto block_40; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitDec___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +block_74: { -uint8_t x_6; lean_object* x_7; -x_6 = lean_unbox(x_3); -x_7 = l_Lean_IR_EmitC_emitDec(x_1, x_2, x_6, x_4, x_5); -lean_dec_ref(x_4); -return x_7; +lean_object* x_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; lean_object* x_70; +x_66 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; +x_67 = lean_string_append(x_65, x_66); +x_68 = lean_string_append(x_67, x_63); +x_69 = 0; +x_70 = l_Lean_Environment_find_x3f(x_62, x_59, x_69); +if (lean_obj_tag(x_70) == 0) +{ +lean_object* x_71; lean_object* x_72; +x_71 = l_Lean_IR_EmitC_emitMainFn___closed__15; +x_72 = l_panic___at___00Lean_IR_EmitC_emitMainFn_spec__2(x_71); +x_41 = x_68; +x_42 = x_60; +x_43 = x_61; +x_44 = x_64; +x_45 = x_63; +x_46 = x_66; +x_47 = x_72; +goto block_58; +} +else +{ +lean_object* x_73; +x_73 = lean_ctor_get(x_70, 0); +lean_inc(x_73); +lean_dec_ref(x_70); +x_41 = x_68; +x_42 = x_60; +x_43 = x_61; +x_44 = x_64; +x_45 = x_63; +x_46 = x_66; +x_47 = x_73; +goto block_58; +} } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitDec___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -uint8_t x_5; lean_object* x_6; -x_5 = lean_unbox(x_3); -x_6 = l_Lean_IR_EmitC_emitDec___redArg(x_1, x_2, x_5, x_4); -return x_6; +lean_object* x_4; +x_4 = l_Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0(x_1, x_2, x_3); +lean_dec_ref(x_2); +lean_dec(x_1); +return x_4; } } -static lean_object* _init_l_Lean_IR_EmitC_emitDel___redArg___closed__0() { +LEAN_EXPORT lean_object* l_List_forM___at___00Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_free_object(", 17, 17); -return x_1; +lean_object* x_4; +x_4 = l_List_forM___at___00Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0_spec__0(x_1, x_2, x_3); +lean_dec_ref(x_2); +lean_dec(x_1); +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitDel___redArg(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_List_forM___at___00Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_3 = l_Lean_IR_EmitC_emitDel___redArg___closed__0; -x_4 = lean_string_append(x_2, x_3); -x_5 = l_Lean_IR_EmitC_argToCString___closed__0; -x_6 = l_Nat_reprFast(x_1); -x_7 = lean_string_append(x_5, x_6); -lean_dec_ref(x_6); -x_8 = lean_string_append(x_4, x_7); -lean_dec_ref(x_7); -x_9 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_10 = lean_string_append(x_8, x_9); -x_11 = l_Lean_IR_EmitC_emitLn___closed__0; -x_12 = lean_box(0); -x_13 = lean_string_append(x_10, x_11); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_12); -lean_ctor_set(x_14, 1, x_13); -return x_14; +lean_object* x_3; +x_3 = l_List_forM___at___00Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0_spec__0___redArg(x_1, x_2); +lean_dec(x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitDel(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Option_instBEq_beq___at___00Lean_IR_EmitC_emitMainFn_spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_4; -x_4 = l_Lean_IR_EmitC_emitDel___redArg(x_1, x_3); +uint8_t x_3; lean_object* x_4; +x_3 = l_Option_instBEq_beq___at___00Lean_IR_EmitC_emitMainFn_spec__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitDel___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT uint8_t l_Lean_IR_EmitC_hasMainFn___lam__0(lean_object* x_1) { _start: { -lean_object* x_4; -x_4 = l_Lean_IR_EmitC_emitDel(x_1, x_2, x_3); -lean_dec_ref(x_2); +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = l_Lean_IR_Decl_name(x_1); +x_3 = l_Lean_IR_EmitC_toCName___closed__1; +x_4 = lean_name_eq(x_2, x_3); +lean_dec(x_2); return x_4; } } -static lean_object* _init_l_Lean_IR_EmitC_emitSetTag___redArg___closed__0() { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_hasMainFn___lam__0___boxed(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_ctor_set_tag(", 18, 18); -return x_1; +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_IR_EmitC_hasMainFn___lam__0(x_1); +lean_dec_ref(x_1); +x_3 = lean_box(x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSetTag___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_hasMainFn(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_4 = l_Lean_IR_EmitC_emitSetTag___redArg___closed__0; -x_5 = lean_string_append(x_3, x_4); -x_6 = l_Lean_IR_EmitC_argToCString___closed__0; -x_7 = l_Nat_reprFast(x_1); -x_8 = lean_string_append(x_6, x_7); -lean_dec_ref(x_7); -x_9 = lean_string_append(x_5, x_8); -lean_dec_ref(x_8); -x_10 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___closed__0; -x_11 = lean_string_append(x_9, x_10); -x_12 = l_Nat_reprFast(x_2); -x_13 = lean_string_append(x_11, x_12); -lean_dec_ref(x_12); -x_14 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_15 = lean_string_append(x_13, x_14); -x_16 = l_Lean_IR_EmitC_emitLn___closed__0; -x_17 = lean_box(0); -x_18 = lean_string_append(x_15, x_16); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_19, 1, x_18); -return x_19; +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; +x_3 = lean_ctor_get(x_1, 0); +lean_inc_ref(x_3); +lean_dec_ref(x_1); +x_4 = lean_alloc_closure((void*)(l_Lean_IR_EmitC_hasMainFn___lam__0___boxed), 1, 0); +x_5 = l_Lean_IR_getDecls(x_3); +x_6 = l_List_any___redArg(x_5, x_4); +x_7 = lean_box(x_6); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_2); +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSetTag(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitMainFnIfNeeded(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_5; -x_5 = l_Lean_IR_EmitC_emitSetTag___redArg(x_1, x_2, x_4); -return x_5; +lean_object* x_3; lean_object* x_4; uint8_t x_5; +lean_inc_ref(x_1); +x_3 = l_Lean_IR_EmitC_hasMainFn(x_1, x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +x_5 = lean_unbox(x_4); +lean_dec(x_4); +if (x_5 == 0) +{ +uint8_t x_6; +lean_dec_ref(x_1); +x_6 = !lean_is_exclusive(x_3); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_3, 0); +lean_dec(x_7); +x_8 = lean_box(0); +lean_ctor_set(x_3, 0, x_8); +return x_3; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = lean_ctor_get(x_3, 1); +lean_inc(x_9); +lean_dec(x_3); +x_10 = lean_box(0); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_9); +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSetTag___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +else { -lean_object* x_5; -x_5 = l_Lean_IR_EmitC_emitSetTag(x_1, x_2, x_3, x_4); +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_3, 1); +lean_inc(x_12); lean_dec_ref(x_3); -return x_5; +x_13 = l_Lean_IR_EmitC_emitMainFn(x_1, x_12); +return x_13; } } -static lean_object* _init_l_Lean_IR_EmitC_emitSet___redArg___closed__0() { +} +static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__0() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_ctor_set(", 14, 14); +x_1 = lean_mk_string_unchecked("import ", 7, 7); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSet___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_5 = l_Lean_IR_EmitC_emitSet___redArg___closed__0; -x_6 = lean_string_append(x_4, x_5); -x_7 = l_Lean_IR_EmitC_argToCString___closed__0; -x_8 = l_Nat_reprFast(x_1); -x_9 = lean_string_append(x_7, x_8); -lean_dec_ref(x_8); -x_10 = lean_string_append(x_6, x_9); -lean_dec_ref(x_9); -x_11 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___closed__0; -x_12 = lean_string_append(x_10, x_11); -x_13 = l_Nat_reprFast(x_2); -x_14 = lean_string_append(x_12, x_13); -lean_dec_ref(x_13); -x_15 = lean_string_append(x_14, x_11); -x_16 = l_Lean_IR_EmitC_emitArg___redArg(x_3, x_15); -x_17 = !lean_is_exclusive(x_16); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_18 = lean_ctor_get(x_16, 1); -x_19 = lean_ctor_get(x_16, 0); -lean_dec(x_19); -x_20 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_21 = lean_string_append(x_18, x_20); -x_22 = l_Lean_IR_EmitC_emitLn___closed__0; -x_23 = lean_box(0); -x_24 = lean_string_append(x_21, x_22); -lean_ctor_set(x_16, 1, x_24); -lean_ctor_set(x_16, 0, x_23); -return x_16; -} -else -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_25 = lean_ctor_get(x_16, 1); -lean_inc(x_25); -lean_dec(x_16); -x_26 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_27 = lean_string_append(x_25, x_26); -x_28 = l_Lean_IR_EmitC_emitLn___closed__0; -x_29 = lean_box(0); -x_30 = lean_string_append(x_27, x_28); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_30); -return x_31; -} -} -} -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSet(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__1() { _start: { -lean_object* x_6; -x_6 = l_Lean_IR_EmitC_emitSet___redArg(x_1, x_2, x_3, x_5); -return x_6; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSet___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__2() { _start: { -lean_object* x_6; -x_6 = l_Lean_IR_EmitC_emitSet(x_1, x_2, x_3, x_4, x_5); -lean_dec_ref(x_4); -return x_6; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("all ", 4, 4); +return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitOffset___redArg___closed__0() { +static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("sizeof(void*)*", 14, 14); +x_1 = lean_mk_string_unchecked("meta ", 5, 5); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitOffset___redArg___closed__1() { +static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked(" + ", 3, 3); +x_1 = lean_mk_string_unchecked("public ", 7, 7); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitOffset___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_4; uint8_t x_5; -x_4 = lean_unsigned_to_nat(0u); -x_5 = lean_nat_dec_lt(x_4, x_1); -if (x_5 == 0) -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; -lean_dec(x_1); -x_6 = lean_box(0); -x_7 = l_Nat_reprFast(x_2); -x_8 = lean_string_append(x_3, x_7); -lean_dec_ref(x_7); -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_6); -lean_ctor_set(x_9, 1, x_8); -return x_9; -} -else +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_2, x_3); +if (x_6 == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_10 = l_Lean_IR_EmitC_emitOffset___redArg___closed__0; -x_11 = lean_string_append(x_3, x_10); -x_12 = l_Nat_reprFast(x_1); -x_13 = lean_string_append(x_11, x_12); -lean_dec_ref(x_12); -x_14 = lean_nat_dec_lt(x_4, x_2); -if (x_14 == 0) +lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_24; lean_object* x_25; lean_object* x_33; +x_7 = lean_array_uget(x_1, x_2); +x_8 = lean_ctor_get_uint8(x_7, sizeof(void*)*1 + 1); +x_9 = l_Lean_IR_EmitC_emitSpreadArg___closed__0; +if (x_8 == 0) { -lean_object* x_15; lean_object* x_16; -lean_dec(x_2); -x_15 = lean_box(0); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_13); -return x_16; +lean_object* x_38; +x_38 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__1; +x_33 = x_38; +goto block_37; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_17 = l_Lean_IR_EmitC_emitOffset___redArg___closed__1; -x_18 = lean_string_append(x_13, x_17); -x_19 = lean_box(0); -x_20 = l_Nat_reprFast(x_2); -x_21 = lean_string_append(x_18, x_20); -lean_dec_ref(x_20); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_19); -lean_ctor_set(x_22, 1, x_21); -return x_22; -} -} -} +lean_object* x_39; +x_39 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__4; +x_33 = x_39; +goto block_37; } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitOffset(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +block_23: { -lean_object* x_5; -x_5 = l_Lean_IR_EmitC_emitOffset___redArg(x_1, x_2, x_4); -return x_5; -} +lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; size_t x_20; size_t x_21; +x_12 = lean_ctor_get(x_7, 0); +lean_inc(x_12); +lean_dec(x_7); +x_13 = lean_string_append(x_10, x_11); +lean_dec_ref(x_11); +x_14 = 1; +x_15 = l_Lean_Name_toStringWithToken___at___00Lean_Name_toString_spec__0(x_12, x_14); +x_16 = lean_string_append(x_13, x_15); +lean_dec_ref(x_15); +x_17 = lean_string_append(x_9, x_16); +lean_dec_ref(x_16); +x_18 = lean_box(0); +x_19 = lean_string_append(x_5, x_17); +lean_dec_ref(x_17); +x_20 = 1; +x_21 = lean_usize_add(x_2, x_20); +x_2 = x_21; +x_4 = x_18; +x_5 = x_19; +goto _start; } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitOffset___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +block_32: { -lean_object* x_5; -x_5 = l_Lean_IR_EmitC_emitOffset(x_1, x_2, x_3, x_4); -lean_dec_ref(x_3); -return x_5; -} +uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_26 = lean_ctor_get_uint8(x_7, sizeof(void*)*1); +x_27 = lean_string_append(x_24, x_25); +lean_dec_ref(x_25); +x_28 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__0; +x_29 = lean_string_append(x_27, x_28); +if (x_26 == 0) +{ +lean_object* x_30; +x_30 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__1; +x_10 = x_29; +x_11 = x_30; +goto block_23; } -static lean_object* _init_l_Lean_IR_EmitC_emitUSet___redArg___closed__0() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_ctor_set_usize(", 20, 20); -return x_1; +lean_object* x_31; +x_31 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__2; +x_10 = x_29; +x_11 = x_31; +goto block_23; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUSet___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +block_37: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; 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_5 = l_Lean_IR_EmitC_emitUSet___redArg___closed__0; -x_6 = lean_string_append(x_4, x_5); -x_7 = l_Lean_IR_EmitC_argToCString___closed__0; -x_8 = l_Nat_reprFast(x_1); -x_9 = lean_string_append(x_7, x_8); -lean_dec_ref(x_8); -x_10 = lean_string_append(x_6, x_9); -lean_dec_ref(x_9); -x_11 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___closed__0; -x_12 = lean_string_append(x_10, x_11); -x_13 = l_Nat_reprFast(x_2); -x_14 = lean_string_append(x_12, x_13); -lean_dec_ref(x_13); -x_15 = lean_string_append(x_14, x_11); -x_16 = l_Nat_reprFast(x_3); -x_17 = lean_string_append(x_7, x_16); -lean_dec_ref(x_16); -x_18 = lean_string_append(x_15, x_17); -lean_dec_ref(x_17); -x_19 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_20 = lean_string_append(x_18, x_19); -x_21 = l_Lean_IR_EmitC_emitLn___closed__0; -x_22 = lean_box(0); -x_23 = lean_string_append(x_20, x_21); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); -return x_24; +uint8_t x_34; +x_34 = lean_ctor_get_uint8(x_7, sizeof(void*)*1 + 2); +if (x_34 == 0) +{ +lean_object* x_35; +x_35 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__1; +x_24 = x_33; +x_25 = x_35; +goto block_32; +} +else +{ +lean_object* x_36; +x_36 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__3; +x_24 = x_33; +x_25 = x_36; +goto block_32; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUSet(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +} +else { -lean_object* x_6; -x_6 = l_Lean_IR_EmitC_emitUSet___redArg(x_1, x_2, x_3, x_5); -return x_6; +lean_object* x_40; +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_4); +lean_ctor_set(x_40, 1, x_5); +return x_40; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUSet___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_6; -x_6 = l_Lean_IR_EmitC_emitUSet(x_1, x_2, x_3, x_4, x_5); -lean_dec_ref(x_4); -return x_6; +lean_object* x_7; +x_7 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg(x_1, x_2, x_3, x_4, x_6); +return x_7; } } -static lean_object* _init_l_Lean_IR_EmitC_emitSSet___redArg___closed__0() { +static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__0() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_ctor_set_float", 19, 19); +x_1 = lean_mk_string_unchecked("#include ", 22, 22); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitSSet___redArg___closed__1() { +static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_ctor_set_float32", 21, 21); +x_1 = lean_mk_string_unchecked("#if defined(__clang__)", 22, 22); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitSSet___redArg___closed__2() { +static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__2() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_ctor_set_uint8", 19, 19); +x_1 = lean_mk_string_unchecked("#pragma clang diagnostic ignored \"-Wunused-parameter\"", 53, 53); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitSSet___redArg___closed__3() { +static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_ctor_set_uint16", 20, 20); +x_1 = lean_mk_string_unchecked("#pragma clang diagnostic ignored \"-Wunused-label\"", 49, 49); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitSSet___redArg___closed__4() { +static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_ctor_set_uint32", 20, 20); +x_1 = lean_mk_string_unchecked("#elif defined(__GNUC__) && !defined(__CLANG__)", 46, 46); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitSSet___redArg___closed__5() { +static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__5() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_ctor_set_uint64", 20, 20); +x_1 = lean_mk_string_unchecked("#pragma GCC diagnostic ignored \"-Wunused-parameter\"", 51, 51); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitSSet___redArg___closed__6() { +static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__6() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("invalid instruction", 19, 19); +x_1 = lean_mk_string_unchecked("#pragma GCC diagnostic ignored \"-Wunused-label\"", 47, 47); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSSet___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__7() { _start: { -lean_object* x_7; -switch (lean_obj_tag(x_5)) { -case 0: -{ -lean_object* x_41; lean_object* x_42; -x_41 = l_Lean_IR_EmitC_emitSSet___redArg___closed__0; -x_42 = lean_string_append(x_6, x_41); -x_7 = x_42; -goto block_40; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("#pragma GCC diagnostic ignored \"-Wunused-but-set-variable\"", 58, 58); +return x_1; } -case 9: -{ -lean_object* x_43; lean_object* x_44; -x_43 = l_Lean_IR_EmitC_emitSSet___redArg___closed__1; -x_44 = lean_string_append(x_6, x_43); -x_7 = x_44; -goto block_40; } -case 1: +static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__8() { +_start: { -lean_object* x_45; lean_object* x_46; -x_45 = l_Lean_IR_EmitC_emitSSet___redArg___closed__2; -x_46 = lean_string_append(x_6, x_45); -x_7 = x_46; -goto block_40; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("#endif", 6, 6); +return x_1; } -case 2: -{ -lean_object* x_47; lean_object* x_48; -x_47 = l_Lean_IR_EmitC_emitSSet___redArg___closed__3; -x_48 = lean_string_append(x_6, x_47); -x_7 = x_48; -goto block_40; } -case 3: +static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__9() { +_start: { -lean_object* x_49; lean_object* x_50; -x_49 = l_Lean_IR_EmitC_emitSSet___redArg___closed__4; -x_50 = lean_string_append(x_6, x_49); -x_7 = x_50; -goto block_40; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("#ifdef __cplusplus", 18, 18); +return x_1; } -case 4: -{ -lean_object* x_51; lean_object* x_52; -x_51 = l_Lean_IR_EmitC_emitSSet___redArg___closed__5; -x_52 = lean_string_append(x_6, x_51); -x_7 = x_52; -goto block_40; } -default: +static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__10() { +_start: { -lean_object* x_53; lean_object* x_54; -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_53 = l_Lean_IR_EmitC_emitSSet___redArg___closed__6; -x_54 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_6); -return x_54; -} +lean_object* x_1; +x_1 = lean_mk_string_unchecked("extern \"C\" {", 12, 12); +return x_1; } -block_40: -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_8 = l_Lean_IR_EmitC_emitFnDeclAux___closed__5; -x_9 = lean_string_append(x_7, x_8); -x_10 = l_Lean_IR_EmitC_argToCString___closed__0; -x_11 = l_Nat_reprFast(x_1); -x_12 = lean_string_append(x_10, x_11); -lean_dec_ref(x_11); -x_13 = lean_string_append(x_9, x_12); -lean_dec_ref(x_12); -x_14 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___closed__0; -x_15 = lean_string_append(x_13, x_14); -x_16 = l_Lean_IR_EmitC_emitOffset___redArg(x_2, x_3, x_15); -x_17 = !lean_is_exclusive(x_16); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_18 = lean_ctor_get(x_16, 1); -x_19 = lean_ctor_get(x_16, 0); -lean_dec(x_19); -x_20 = lean_string_append(x_18, x_14); -x_21 = l_Nat_reprFast(x_4); -x_22 = lean_string_append(x_10, x_21); -lean_dec_ref(x_21); -x_23 = lean_string_append(x_20, x_22); -lean_dec_ref(x_22); -x_24 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_25 = lean_string_append(x_23, x_24); -x_26 = l_Lean_IR_EmitC_emitLn___closed__0; -x_27 = lean_box(0); -x_28 = lean_string_append(x_25, x_26); -lean_ctor_set(x_16, 1, x_28); -lean_ctor_set(x_16, 0, x_27); -return x_16; } -else +static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__11() { +_start: { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_29 = lean_ctor_get(x_16, 1); -lean_inc(x_29); -lean_dec(x_16); -x_30 = lean_string_append(x_29, x_14); -x_31 = l_Nat_reprFast(x_4); -x_32 = lean_string_append(x_10, x_31); -lean_dec_ref(x_31); -x_33 = lean_string_append(x_30, x_32); -lean_dec_ref(x_32); -x_34 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_35 = lean_string_append(x_33, x_34); -x_36 = l_Lean_IR_EmitC_emitLn___closed__0; -x_37 = lean_box(0); -x_38 = lean_string_append(x_35, x_36); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_37); -lean_ctor_set(x_39, 1, x_38); -return x_39; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_IR_EmitC_emitFileHeader___closed__8; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } +static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_EmitC_emitFileHeader___closed__11; +x_2 = l_Lean_IR_EmitC_emitFileHeader___closed__10; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSSet(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__13() { _start: { -lean_object* x_8; -x_8 = l_Lean_IR_EmitC_emitSSet___redArg(x_1, x_2, x_3, x_4, x_5, x_7); -return x_8; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_EmitC_emitFileHeader___closed__12; +x_2 = l_Lean_IR_EmitC_emitFileHeader___closed__9; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSSet___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__14() { _start: { -lean_object* x_8; -x_8 = l_Lean_IR_EmitC_emitSSet(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec_ref(x_6); -lean_dec(x_5); -return x_8; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_EmitC_emitFileHeader___closed__13; +x_2 = l_Lean_IR_EmitC_emitFileHeader___closed__8; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSSet___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__15() { _start: { -lean_object* x_7; -x_7 = l_Lean_IR_EmitC_emitSSet___redArg(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_5); -return x_7; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_EmitC_emitFileHeader___closed__14; +x_2 = l_Lean_IR_EmitC_emitFileHeader___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -static lean_object* _init_l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0___redArg___closed__0() { +static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__16() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked(" = ", 3, 3); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_EmitC_emitFileHeader___closed__15; +x_2 = l_Lean_IR_EmitC_emitFileHeader___closed__6; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__17() { _start: { -lean_object* x_6; uint8_t x_7; -x_6 = lean_unsigned_to_nat(0u); -x_7 = lean_nat_dec_eq(x_4, x_6); -if (x_7 == 1) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_EmitC_emitFileHeader___closed__16; +x_2 = l_Lean_IR_EmitC_emitFileHeader___closed__5; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__18() { +_start: { -lean_object* x_8; lean_object* x_9; -lean_dec(x_4); -x_8 = lean_box(0); -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_8); -lean_ctor_set(x_9, 1, x_5); -return x_9; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_EmitC_emitFileHeader___closed__17; +x_2 = l_Lean_IR_EmitC_emitFileHeader___closed__4; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } -else +} +static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__19() { +_start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_10 = lean_unsigned_to_nat(1u); -x_11 = lean_nat_sub(x_4, x_10); -lean_dec(x_4); -x_12 = lean_nat_sub(x_3, x_11); -x_13 = lean_nat_sub(x_12, x_10); -lean_dec(x_12); -x_14 = lean_array_fget_borrowed(x_1, x_13); -x_15 = lean_ctor_get(x_14, 0); -x_16 = lean_array_fget_borrowed(x_2, x_13); -lean_dec(x_13); -x_17 = l_Lean_IR_EmitC_argToCString___closed__0; -lean_inc(x_15); -x_18 = l_Nat_reprFast(x_15); -x_19 = lean_string_append(x_17, x_18); -lean_dec_ref(x_18); -x_20 = lean_string_append(x_5, x_19); -lean_dec_ref(x_19); -x_21 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0___redArg___closed__0; -x_22 = lean_string_append(x_20, x_21); -lean_inc(x_16); -x_23 = l_Lean_IR_EmitC_emitArg___redArg(x_16, x_22); -x_24 = lean_ctor_get(x_23, 1); -lean_inc(x_24); -lean_dec_ref(x_23); -x_25 = l_Lean_IR_EmitC_emitFnDeclAux___closed__0; -x_26 = lean_string_append(x_24, x_25); -x_27 = l_Lean_IR_EmitC_emitLn___closed__0; -x_28 = lean_string_append(x_26, x_27); -x_4 = x_11; -x_5 = x_28; -goto _start; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_EmitC_emitFileHeader___closed__18; +x_2 = l_Lean_IR_EmitC_emitFileHeader___closed__3; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } +static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_EmitC_emitFileHeader___closed__19; +x_2 = l_Lean_IR_EmitC_emitFileHeader___closed__2; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__21() { _start: { -lean_object* x_8; -x_8 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0___redArg(x_1, x_2, x_3, x_4, x_7); -return x_8; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_EmitC_emitFileHeader___closed__20; +x_2 = l_Lean_IR_EmitC_emitFileHeader___closed__1; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -static lean_object* _init_l_Lean_IR_EmitC_emitJmp___closed__0() { +static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__22() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("invalid goto", 12, 12); +x_1 = lean_mk_string_unchecked("// Lean compiler output", 23, 23); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitJmp___closed__1() { +static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__23() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("goto ", 5, 5); +x_1 = lean_mk_string_unchecked("// Module: ", 11, 11); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitJmp___closed__2() { +static lean_object* _init_l_Lean_IR_EmitC_emitFileHeader___closed__24() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("block_", 6, 6); +x_1 = lean_mk_string_unchecked("// Imports:", 11, 11); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitJmp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFileHeader(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_5; -x_5 = l_Lean_IR_EmitC_getJPParams(x_1, x_3, x_4); -if (lean_obj_tag(x_5) == 0) -{ -uint8_t x_6; -x_6 = !lean_is_exclusive(x_5); -if (x_6 == 0) -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_7 = lean_ctor_get(x_5, 0); -x_8 = lean_ctor_get(x_5, 1); -x_9 = lean_array_get_size(x_2); -x_10 = lean_array_get_size(x_7); -x_11 = lean_nat_dec_eq(x_9, x_10); -if (x_11 == 0) +lean_object* x_3; lean_object* x_12; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_15 = lean_ctor_get(x_1, 0); +lean_inc_ref(x_15); +x_16 = lean_ctor_get(x_1, 1); +lean_inc(x_16); +lean_dec_ref(x_1); +x_17 = l_Lean_IR_EmitC_emitFileHeader___closed__22; +x_18 = lean_string_append(x_2, x_17); +x_19 = l_Lean_IR_EmitC_emitLn___closed__0; +x_20 = lean_string_append(x_18, x_19); +x_21 = l_Lean_IR_EmitC_emitFileHeader___closed__23; +x_22 = 1; +x_23 = l_Lean_Name_toStringWithToken___at___00Lean_Name_toString_spec__0(x_16, x_22); +x_24 = lean_string_append(x_21, x_23); +lean_dec_ref(x_23); +x_25 = lean_string_append(x_20, x_24); +lean_dec_ref(x_24); +x_26 = lean_string_append(x_25, x_19); +x_27 = l_Lean_IR_EmitC_emitFileHeader___closed__24; +x_28 = lean_string_append(x_26, x_27); +x_29 = l_Lean_Environment_imports(x_15); +lean_dec_ref(x_15); +x_30 = lean_unsigned_to_nat(0u); +x_31 = lean_array_get_size(x_29); +x_32 = lean_nat_dec_lt(x_30, x_31); +if (x_32 == 0) { -lean_object* x_12; -lean_dec(x_7); -lean_dec(x_1); -x_12 = l_Lean_IR_EmitC_emitJmp___closed__0; -lean_ctor_set_tag(x_5, 1); -lean_ctor_set(x_5, 0, x_12); -return x_5; +lean_dec_ref(x_29); +x_3 = x_28; +goto block_11; } else { -lean_object* x_13; uint8_t x_14; -lean_free_object(x_5); -x_13 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0___redArg(x_7, x_2, x_9, x_9, x_8); -lean_dec(x_7); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) +lean_object* x_33; uint8_t x_34; +x_33 = lean_box(0); +x_34 = lean_nat_dec_le(x_31, x_31); +if (x_34 == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_15 = lean_ctor_get(x_13, 1); -x_16 = lean_ctor_get(x_13, 0); -lean_dec(x_16); -x_17 = l_Lean_IR_EmitC_emitJmp___closed__1; -x_18 = lean_string_append(x_15, x_17); -x_19 = l_Lean_IR_EmitC_emitJmp___closed__2; -x_20 = l_Nat_reprFast(x_1); -x_21 = lean_string_append(x_19, x_20); -lean_dec_ref(x_20); -x_22 = lean_string_append(x_18, x_21); -lean_dec_ref(x_21); -x_23 = l_Lean_IR_EmitC_emitFnDeclAux___closed__0; -x_24 = lean_string_append(x_22, x_23); -x_25 = l_Lean_IR_EmitC_emitLn___closed__0; -x_26 = lean_box(0); -x_27 = lean_string_append(x_24, x_25); -lean_ctor_set(x_13, 1, x_27); -lean_ctor_set(x_13, 0, x_26); -return x_13; +if (x_32 == 0) +{ +lean_dec_ref(x_29); +x_3 = x_28; +goto block_11; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_28 = lean_ctor_get(x_13, 1); -lean_inc(x_28); -lean_dec(x_13); -x_29 = l_Lean_IR_EmitC_emitJmp___closed__1; -x_30 = lean_string_append(x_28, x_29); -x_31 = l_Lean_IR_EmitC_emitJmp___closed__2; -x_32 = l_Nat_reprFast(x_1); -x_33 = lean_string_append(x_31, x_32); -lean_dec_ref(x_32); -x_34 = lean_string_append(x_30, x_33); -lean_dec_ref(x_33); -x_35 = l_Lean_IR_EmitC_emitFnDeclAux___closed__0; -x_36 = lean_string_append(x_34, x_35); -x_37 = l_Lean_IR_EmitC_emitLn___closed__0; -x_38 = lean_box(0); -x_39 = lean_string_append(x_36, x_37); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; -} +size_t x_35; size_t x_36; lean_object* x_37; +x_35 = 0; +x_36 = lean_usize_of_nat(x_31); +x_37 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg(x_29, x_35, x_36, x_33, x_28); +lean_dec_ref(x_29); +x_12 = x_37; +goto block_14; } } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; -x_41 = lean_ctor_get(x_5, 0); -x_42 = lean_ctor_get(x_5, 1); -lean_inc(x_42); -lean_inc(x_41); -lean_dec(x_5); -x_43 = lean_array_get_size(x_2); -x_44 = lean_array_get_size(x_41); -x_45 = lean_nat_dec_eq(x_43, x_44); -if (x_45 == 0) -{ -lean_object* x_46; lean_object* x_47; -lean_dec(x_41); -lean_dec(x_1); -x_46 = l_Lean_IR_EmitC_emitJmp___closed__0; -x_47 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_42); -return x_47; +size_t x_38; size_t x_39; lean_object* x_40; +x_38 = 0; +x_39 = lean_usize_of_nat(x_31); +x_40 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg(x_29, x_38, x_39, x_33, x_28); +lean_dec_ref(x_29); +x_12 = x_40; +goto block_14; } -else +} +block_11: { -lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_48 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0___redArg(x_41, x_2, x_43, x_43, x_42); -lean_dec(x_41); -x_49 = lean_ctor_get(x_48, 1); -lean_inc(x_49); -if (lean_is_exclusive(x_48)) { - lean_ctor_release(x_48, 0); - lean_ctor_release(x_48, 1); - x_50 = x_48; -} else { - lean_dec_ref(x_48); - x_50 = lean_box(0); -} -x_51 = l_Lean_IR_EmitC_emitJmp___closed__1; -x_52 = lean_string_append(x_49, x_51); -x_53 = l_Lean_IR_EmitC_emitJmp___closed__2; -x_54 = l_Nat_reprFast(x_1); -x_55 = lean_string_append(x_53, x_54); -lean_dec_ref(x_54); -x_56 = lean_string_append(x_52, x_55); -lean_dec_ref(x_55); -x_57 = l_Lean_IR_EmitC_emitFnDeclAux___closed__0; -x_58 = lean_string_append(x_56, x_57); -x_59 = l_Lean_IR_EmitC_emitLn___closed__0; -x_60 = lean_box(0); -x_61 = lean_string_append(x_58, x_59); -if (lean_is_scalar(x_50)) { - x_62 = lean_alloc_ctor(0, 2, 0); -} else { - x_62 = x_50; -} -lean_ctor_set(x_62, 0, x_60); -lean_ctor_set(x_62, 1, x_61); -return x_62; -} -} +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_4 = l_Lean_IR_EmitC_emitLn___closed__0; +x_5 = lean_string_append(x_3, x_4); +x_6 = l_Lean_IR_EmitC_emitFileHeader___closed__0; +x_7 = lean_string_append(x_5, x_6); +x_8 = lean_string_append(x_7, x_4); +x_9 = l_Lean_IR_EmitC_emitFileHeader___closed__21; +x_10 = l_List_forM___at___00Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0_spec__0___redArg(x_9, x_8); +return x_10; } -else +block_14: { -uint8_t x_63; -lean_dec(x_1); -x_63 = !lean_is_exclusive(x_5); -if (x_63 == 0) +if (lean_obj_tag(x_12) == 0) { -return x_5; +lean_object* x_13; +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +lean_dec_ref(x_12); +x_3 = x_13; +goto block_11; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_5, 0); -x_65 = lean_ctor_get(x_5, 1); -lean_inc(x_65); -lean_inc(x_64); -lean_dec(x_5); -x_66 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_66, 0, x_64); -lean_ctor_set(x_66, 1, x_65); -return x_66; +return x_12; } } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_8; -x_8 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec_ref(x_6); +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_8 = lean_unbox_usize(x_3); lean_dec(x_3); -lean_dec_ref(x_2); +x_9 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0(x_1, x_7, x_8, x_4, x_5, x_6); +lean_dec_ref(x_5); lean_dec_ref(x_1); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitJmp___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_IR_EmitC_emitJmp(x_1, x_2, x_3, x_4); -lean_dec_ref(x_3); -lean_dec_ref(x_2); -return x_5; +return x_9; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_6; -x_6 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0___redArg(x_1, x_2, x_3, x_4, x_5); +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_7 = lean_unbox_usize(x_3); lean_dec(x_3); -lean_dec_ref(x_2); +x_8 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg(x_1, x_6, x_7, x_4, x_5); lean_dec_ref(x_1); -return x_6; +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLhs___redArg(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_IR_EmitC_emitFileFooter___redArg___closed__0() { _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_Lean_IR_EmitC_argToCString___closed__0; -x_4 = l_Nat_reprFast(x_1); -x_5 = lean_string_append(x_3, x_4); -lean_dec_ref(x_4); -x_6 = lean_string_append(x_2, x_5); -lean_dec_ref(x_5); -x_7 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0___redArg___closed__0; -x_8 = lean_box(0); -x_9 = lean_string_append(x_6, x_7); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_8); -lean_ctor_set(x_10, 1, x_9); -return x_10; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_EmitC_emitFileHeader___closed__11; +x_2 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLhs(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_IR_EmitC_emitFileFooter___redArg___closed__1() { _start: { -lean_object* x_4; -x_4 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_3); -return x_4; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_EmitC_emitFileFooter___redArg___closed__0; +x_2 = l_Lean_IR_EmitC_emitFileHeader___closed__9; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLhs___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFileFooter___redArg(lean_object* x_1) { _start: { -lean_object* x_4; -x_4 = l_Lean_IR_EmitC_emitLhs(x_1, x_2, x_3); -lean_dec_ref(x_2); -return x_4; +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_IR_EmitC_emitFileFooter___redArg___closed__1; +x_3 = l_List_forM___at___00Lean_IR_EmitC_emitLns___at___00Lean_IR_EmitC_emitMainFn_spec__0_spec__0___redArg(x_2, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitArgs_spec__0___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFileFooter(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_5; uint8_t x_6; -x_5 = lean_unsigned_to_nat(0u); -x_6 = lean_nat_dec_eq(x_3, x_5); -if (x_6 == 1) -{ -lean_object* x_7; lean_object* x_8; -lean_dec(x_3); -x_7 = lean_box(0); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_7); -lean_ctor_set(x_8, 1, x_4); -return x_8; -} -else -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_19; -x_9 = lean_unsigned_to_nat(1u); -x_10 = lean_nat_sub(x_3, x_9); -lean_dec(x_3); -x_11 = lean_nat_sub(x_2, x_10); -x_12 = lean_nat_sub(x_11, x_9); -lean_dec(x_11); -x_19 = lean_nat_dec_lt(x_5, x_12); -if (x_19 == 0) -{ -x_13 = x_4; -goto block_18; -} -else -{ -lean_object* x_20; lean_object* x_21; -x_20 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___closed__0; -x_21 = lean_string_append(x_4, x_20); -x_13 = x_21; -goto block_18; -} -block_18: -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_array_fget_borrowed(x_1, x_12); -lean_dec(x_12); -lean_inc(x_14); -x_15 = l_Lean_IR_EmitC_emitArg___redArg(x_14, x_13); -x_16 = lean_ctor_get(x_15, 1); -lean_inc(x_16); -lean_dec_ref(x_15); -x_3 = x_10; -x_4 = x_16; -goto _start; -} -} +lean_object* x_3; +x_3 = l_Lean_IR_EmitC_emitFileFooter___redArg(x_2); +return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitArgs_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFileFooter___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_7; -x_7 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitArgs_spec__0___redArg(x_1, x_2, x_3, x_6); -return x_7; +lean_object* x_3; +x_3 = l_Lean_IR_EmitC_emitFileFooter(x_1, x_2); +lean_dec_ref(x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitArgs(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_IR_EmitC_throwUnknownVar___redArg___closed__0() { _start: { -lean_object* x_4; lean_object* x_5; -x_4 = lean_array_get_size(x_1); -x_5 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitArgs_spec__0___redArg(x_1, x_4, x_4, x_3); -return x_5; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("unknown variable '", 18, 18); +return x_1; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitArgs_spec__0___boxed(lean_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_IR_EmitC_throwUnknownVar___redArg(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_7; -x_7 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitArgs_spec__0(x_1, x_2, x_3, x_4, x_5, x_6); +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_3 = l_Lean_IR_EmitC_throwUnknownVar___redArg___closed__0; +x_4 = l_Lean_IR_EmitC_argToCString___closed__0; +x_5 = l_Nat_reprFast(x_1); +x_6 = lean_string_append(x_4, x_5); lean_dec_ref(x_5); -lean_dec(x_2); -lean_dec_ref(x_1); -return x_7; +x_7 = lean_string_append(x_3, x_6); +lean_dec_ref(x_6); +x_8 = l_Lean_IR_EmitC_getDecl___closed__1; +x_9 = lean_string_append(x_7, x_8); +x_10 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_2); +return x_10; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitArgs___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_throwUnknownVar(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_4; -x_4 = l_Lean_IR_EmitC_emitArgs(x_1, x_2, x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -return x_4; +lean_object* x_5; +x_5 = l_Lean_IR_EmitC_throwUnknownVar___redArg(x_2, x_4); +return x_5; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitArgs_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_throwUnknownVar___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitArgs_spec__0___redArg(x_1, x_2, x_3, x_4); -lean_dec(x_2); -lean_dec_ref(x_1); +x_5 = l_Lean_IR_EmitC_throwUnknownVar(x_1, x_2, x_3, x_4); +lean_dec_ref(x_3); return x_5; } } -static lean_object* _init_l_Lean_IR_EmitC_emitCtorScalarSize___redArg___closed__0() { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0_spec__0___redArg(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("sizeof(size_t)*", 15, 15); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCtorScalarSize___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +if (lean_obj_tag(x_2) == 0) { -lean_object* x_4; uint8_t x_5; -x_4 = lean_unsigned_to_nat(0u); -x_5 = lean_nat_dec_eq(x_1, x_4); -if (x_5 == 0) +lean_object* x_3; +x_3 = lean_box(0); +return x_3; +} +else { -uint8_t x_6; -x_6 = lean_nat_dec_eq(x_2, x_4); -if (x_6 == 0) +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 1); +x_6 = lean_ctor_get(x_2, 2); +x_7 = l_Lean_IR_instBEqJoinPointId_beq(x_4, x_1); +if (x_7 == 0) { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_7 = l_Lean_IR_EmitC_emitCtorScalarSize___redArg___closed__0; -x_8 = lean_string_append(x_3, x_7); -x_9 = l_Nat_reprFast(x_1); -x_10 = lean_string_append(x_8, x_9); -lean_dec_ref(x_9); -x_11 = l_Lean_IR_EmitC_emitOffset___redArg___closed__1; -x_12 = lean_string_append(x_10, x_11); -x_13 = lean_box(0); -x_14 = l_Nat_reprFast(x_2); -x_15 = lean_string_append(x_12, x_14); -lean_dec_ref(x_14); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_13); -lean_ctor_set(x_16, 1, x_15); -return x_16; +x_2 = x_6; +goto _start; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -lean_dec(x_2); -x_17 = l_Lean_IR_EmitC_emitCtorScalarSize___redArg___closed__0; -x_18 = lean_string_append(x_3, x_17); -x_19 = lean_box(0); -x_20 = l_Nat_reprFast(x_1); -x_21 = lean_string_append(x_18, x_20); -lean_dec_ref(x_20); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_19); -lean_ctor_set(x_22, 1, x_21); -return x_22; +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; } } -else -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -lean_dec(x_1); -x_23 = lean_box(0); -x_24 = l_Nat_reprFast(x_2); -x_25 = lean_string_append(x_3, x_24); -lean_dec_ref(x_24); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_23); -lean_ctor_set(x_26, 1, x_25); -return x_26; } } +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0_spec__0___redArg(x_2, x_3); +return x_4; +} } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCtorScalarSize(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0___redArg(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_5; -x_5 = l_Lean_IR_EmitC_emitCtorScalarSize___redArg(x_1, x_2, x_4); -return x_5; +lean_object* x_3; lean_object* x_4; uint64_t x_5; uint64_t x_6; uint64_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; size_t x_12; size_t x_13; size_t x_14; size_t x_15; size_t x_16; lean_object* x_17; lean_object* x_18; +x_3 = lean_ctor_get(x_1, 1); +x_4 = lean_array_get_size(x_3); +x_5 = l_Lean_IR_instHashableJoinPointId_hash(x_2); +x_6 = 32; +x_7 = lean_uint64_shift_right(x_5, x_6); +x_8 = lean_uint64_xor(x_5, x_7); +x_9 = 16; +x_10 = lean_uint64_shift_right(x_8, x_9); +x_11 = lean_uint64_xor(x_8, x_10); +x_12 = lean_uint64_to_usize(x_11); +x_13 = lean_usize_of_nat(x_4); +x_14 = 1; +x_15 = lean_usize_sub(x_13, x_14); +x_16 = lean_usize_land(x_12, x_15); +x_17 = lean_array_uget(x_3, x_16); +x_18 = l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0_spec__0___redArg(x_2, x_17); +lean_dec(x_17); +return x_18; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCtorScalarSize___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_5; -x_5 = l_Lean_IR_EmitC_emitCtorScalarSize(x_1, x_2, x_3, x_4); -lean_dec_ref(x_3); -return x_5; +lean_object* x_4; +x_4 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0___redArg(x_2, x_3); +return x_4; } } -static lean_object* _init_l_Lean_IR_EmitC_emitAllocCtor___redArg___closed__0() { +static lean_object* _init_l_Lean_IR_EmitC_getJPParams___closed__0() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_alloc_ctor(", 16, 16); +x_1 = lean_mk_string_unchecked("unknown join point", 18, 18); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitAllocCtor___redArg(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_getJPParams(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_3 = lean_ctor_get(x_1, 1); -lean_inc(x_3); -x_4 = lean_ctor_get(x_1, 2); -lean_inc(x_4); -x_5 = lean_ctor_get(x_1, 3); -lean_inc(x_5); -x_6 = lean_ctor_get(x_1, 4); -lean_inc(x_6); -lean_dec_ref(x_1); -x_7 = l_Lean_IR_EmitC_emitAllocCtor___redArg___closed__0; -x_8 = lean_string_append(x_2, x_7); -x_9 = l_Nat_reprFast(x_3); -x_10 = lean_string_append(x_8, x_9); -lean_dec_ref(x_9); -x_11 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___closed__0; -x_12 = lean_string_append(x_10, x_11); -x_13 = l_Nat_reprFast(x_4); -x_14 = lean_string_append(x_12, x_13); -lean_dec_ref(x_13); -x_15 = lean_string_append(x_14, x_11); -x_16 = l_Lean_IR_EmitC_emitCtorScalarSize___redArg(x_5, x_6, x_15); -x_17 = !lean_is_exclusive(x_16); -if (x_17 == 0) +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_2, 2); +x_5 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0___redArg(x_4, x_1); +if (lean_obj_tag(x_5) == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_18 = lean_ctor_get(x_16, 1); -x_19 = lean_ctor_get(x_16, 0); -lean_dec(x_19); -x_20 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_21 = lean_string_append(x_18, x_20); -x_22 = l_Lean_IR_EmitC_emitLn___closed__0; -x_23 = lean_box(0); -x_24 = lean_string_append(x_21, x_22); -lean_ctor_set(x_16, 1, x_24); -lean_ctor_set(x_16, 0, x_23); -return x_16; +lean_object* x_6; lean_object* x_7; +x_6 = l_Lean_IR_EmitC_getJPParams___closed__0; +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_3); +return x_7; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_25 = lean_ctor_get(x_16, 1); -lean_inc(x_25); -lean_dec(x_16); -x_26 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_27 = lean_string_append(x_25, x_26); -x_28 = l_Lean_IR_EmitC_emitLn___closed__0; -x_29 = lean_box(0); -x_30 = lean_string_append(x_27, x_28); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_30); -return x_31; +lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_5, 0); +lean_inc(x_8); +lean_dec_ref(x_5); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_3); +return x_9; } } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitAllocCtor(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_IR_EmitC_emitAllocCtor___redArg(x_1, x_3); +x_4 = l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0_spec__0(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitAllocCtor___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_IR_EmitC_emitAllocCtor(x_1, x_2, x_3); +x_4 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0(x_1, x_2, x_3); +lean_dec(x_3); lean_dec_ref(x_2); return x_4; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitCtorSetArgs_spec__0___redArg(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_IR_EmitC_getJPParams___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_6; uint8_t x_7; -x_6 = lean_unsigned_to_nat(0u); -x_7 = lean_nat_dec_eq(x_4, x_6); -if (x_7 == 1) -{ -lean_object* x_8; lean_object* x_9; -lean_dec(x_4); +lean_object* x_4; +x_4 = l_Lean_IR_EmitC_getJPParams(x_1, x_2, x_3); +lean_dec_ref(x_2); lean_dec(x_1); -x_8 = lean_box(0); -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_8); -lean_ctor_set(x_9, 1, x_5); -return x_9; -} -else -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_10 = lean_unsigned_to_nat(1u); -x_11 = lean_nat_sub(x_4, x_10); -lean_dec(x_4); -x_12 = lean_nat_sub(x_3, x_11); -x_13 = lean_nat_sub(x_12, x_10); -lean_dec(x_12); -x_14 = l_Lean_IR_EmitC_emitSet___redArg___closed__0; -x_15 = lean_string_append(x_5, x_14); -x_16 = l_Lean_IR_EmitC_argToCString___closed__0; -lean_inc(x_1); -x_17 = l_Nat_reprFast(x_1); -x_18 = lean_string_append(x_16, x_17); -lean_dec_ref(x_17); -x_19 = lean_string_append(x_15, x_18); -lean_dec_ref(x_18); -x_20 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___closed__0; -x_21 = lean_string_append(x_19, x_20); -lean_inc(x_13); -x_22 = l_Nat_reprFast(x_13); -x_23 = lean_string_append(x_21, x_22); -lean_dec_ref(x_22); -x_24 = lean_string_append(x_23, x_20); -x_25 = lean_array_fget_borrowed(x_2, x_13); -lean_dec(x_13); -lean_inc(x_25); -x_26 = l_Lean_IR_EmitC_emitArg___redArg(x_25, x_24); -x_27 = lean_ctor_get(x_26, 1); -lean_inc(x_27); -lean_dec_ref(x_26); -x_28 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_29 = lean_string_append(x_27, x_28); -x_30 = l_Lean_IR_EmitC_emitLn___closed__0; -x_31 = lean_string_append(x_29, x_30); -x_4 = x_11; -x_5 = x_31; -goto _start; -} +return x_4; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitCtorSetArgs_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_8; -x_8 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitCtorSetArgs_spec__0___redArg(x_1, x_2, x_3, x_4, x_7); -return x_8; +lean_object* x_3; +x_3 = l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0_spec__0___redArg(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCtorSetArgs(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_5; lean_object* x_6; -x_5 = lean_array_get_size(x_2); -x_6 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitCtorSetArgs_spec__0___redArg(x_1, x_2, x_5, x_5, x_4); -return x_6; +lean_object* x_3; +x_3 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0___redArg(x_1, x_2); +lean_dec(x_2); +lean_dec_ref(x_1); +return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitCtorSetArgs_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +static lean_object* _init_l_Lean_IR_EmitC_declareVar___closed__0() { _start: { -lean_object* x_8; -x_8 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitCtorSetArgs_spec__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec_ref(x_6); -lean_dec(x_3); -lean_dec_ref(x_2); -return x_8; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("; ", 2, 2); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCtorSetArgs___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_declareVar(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_IR_EmitC_emitCtorSetArgs(x_1, x_2, x_3, x_4); -lean_dec_ref(x_3); -lean_dec_ref(x_2); +lean_object* x_5; uint8_t x_6; +x_5 = l_Lean_IR_EmitC_toCType(x_2, x_3, x_4); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_ctor_get(x_5, 1); +x_9 = lean_string_append(x_8, x_7); +lean_dec(x_7); +x_10 = l_Lean_IR_EmitC_emitSpreadArg___closed__0; +x_11 = lean_string_append(x_9, x_10); +x_12 = l_Lean_IR_EmitC_argToCString___closed__0; +x_13 = l_Nat_reprFast(x_1); +x_14 = lean_string_append(x_12, x_13); +lean_dec_ref(x_13); +x_15 = lean_string_append(x_11, x_14); +lean_dec_ref(x_14); +x_16 = l_Lean_IR_EmitC_declareVar___closed__0; +x_17 = lean_box(0); +x_18 = lean_string_append(x_15, x_16); +lean_ctor_set(x_5, 1, x_18); +lean_ctor_set(x_5, 0, x_17); return x_5; } -} -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitCtorSetArgs_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +else { -lean_object* x_6; -x_6 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitCtorSetArgs_spec__0___redArg(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_3); -lean_dec_ref(x_2); -return x_6; +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_19 = lean_ctor_get(x_5, 0); +x_20 = lean_ctor_get(x_5, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_5); +x_21 = lean_string_append(x_20, x_19); +lean_dec(x_19); +x_22 = l_Lean_IR_EmitC_emitSpreadArg___closed__0; +x_23 = lean_string_append(x_21, x_22); +x_24 = l_Lean_IR_EmitC_argToCString___closed__0; +x_25 = l_Nat_reprFast(x_1); +x_26 = lean_string_append(x_24, x_25); +lean_dec_ref(x_25); +x_27 = lean_string_append(x_23, x_26); +lean_dec_ref(x_26); +x_28 = l_Lean_IR_EmitC_declareVar___closed__0; +x_29 = lean_box(0); +x_30 = lean_string_append(x_27, x_28); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; } } -static lean_object* _init_l_Lean_IR_EmitC_emitCtor___closed__0() { +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_declareVar___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_box(", 9, 9); -return x_1; +lean_object* x_5; +x_5 = l_Lean_IR_EmitC_declareVar(x_1, x_2, x_3, x_4); +lean_dec_ref(x_3); +lean_dec(x_2); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCtor(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_declareParams_spec__0(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_31; uint8_t x_32; -lean_inc(x_1); -x_6 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_5); -x_7 = lean_ctor_get(x_6, 1); -lean_inc(x_7); -if (lean_is_exclusive(x_6)) { - lean_ctor_release(x_6, 0); - lean_ctor_release(x_6, 1); - x_8 = x_6; -} else { - lean_dec_ref(x_6); - x_8 = lean_box(0); -} -x_13 = lean_ctor_get(x_2, 1); -x_14 = lean_ctor_get(x_2, 2); -x_15 = lean_ctor_get(x_2, 3); -x_16 = lean_ctor_get(x_2, 4); -x_31 = lean_unsigned_to_nat(0u); -x_32 = lean_nat_dec_eq(x_14, x_31); -if (x_32 == 0) +uint8_t x_7; +x_7 = lean_usize_dec_eq(x_2, x_3); +if (x_7 == 0) { -x_17 = x_32; -goto block_30; +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_array_uget(x_1, x_2); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +lean_dec(x_8); +x_11 = l_Lean_IR_EmitC_declareVar(x_9, x_10, x_5, x_6); +lean_dec(x_10); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; lean_object* x_13; size_t x_14; size_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_ref(x_11); +x_14 = 1; +x_15 = lean_usize_add(x_2, x_14); +x_2 = x_15; +x_4 = x_12; +x_6 = x_13; +goto _start; } else { -uint8_t x_33; -x_33 = lean_nat_dec_eq(x_15, x_31); -x_17 = x_33; -goto block_30; +return x_11; } -block_12: +} +else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_9 = l_Lean_IR_EmitC_emitAllocCtor___redArg(x_2, x_7); -x_10 = lean_ctor_get(x_9, 1); -lean_inc(x_10); -lean_dec_ref(x_9); -x_11 = l_Lean_IR_EmitC_emitCtorSetArgs(x_1, x_3, x_4, x_10); -return x_11; +lean_object* x_17; +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_4); +lean_ctor_set(x_17, 1, x_6); +return x_17; } -block_30: +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_declareParams(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -if (x_17 == 0) +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_array_get_size(x_1); +x_6 = lean_box(0); +x_7 = lean_nat_dec_lt(x_4, x_5); +if (x_7 == 0) { -lean_dec(x_8); -goto block_12; +lean_object* x_8; +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_6); +lean_ctor_set(x_8, 1, x_3); +return x_8; } else { -lean_object* x_18; uint8_t x_19; -x_18 = lean_unsigned_to_nat(0u); -x_19 = lean_nat_dec_eq(x_16, x_18); -if (x_19 == 0) +uint8_t x_9; +x_9 = lean_nat_dec_le(x_5, x_5); +if (x_9 == 0) { -lean_dec(x_8); -goto block_12; +if (x_7 == 0) +{ +lean_object* x_10; +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_6); +lean_ctor_set(x_10, 1, x_3); +return x_10; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -lean_inc(x_13); -lean_dec_ref(x_2); -lean_dec(x_1); -x_20 = l_Lean_IR_EmitC_emitCtor___closed__0; -x_21 = lean_string_append(x_7, x_20); -x_22 = l_Nat_reprFast(x_13); -x_23 = lean_string_append(x_21, x_22); -lean_dec_ref(x_22); -x_24 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_25 = lean_string_append(x_23, x_24); -x_26 = l_Lean_IR_EmitC_emitLn___closed__0; -x_27 = lean_box(0); -x_28 = lean_string_append(x_25, x_26); -if (lean_is_scalar(x_8)) { - x_29 = lean_alloc_ctor(0, 2, 0); -} else { - x_29 = x_8; +size_t x_11; size_t x_12; lean_object* x_13; +x_11 = 0; +x_12 = lean_usize_of_nat(x_5); +x_13 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_declareParams_spec__0(x_1, x_11, x_12, x_6, x_2, x_3); +return x_13; } -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -return x_29; } +else +{ +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = 0; +x_15 = lean_usize_of_nat(x_5); +x_16 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_declareParams_spec__0(x_1, x_14, x_15, x_6, x_2, x_3); +return x_16; } } } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCtor___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_declareParams_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_6; -x_6 = l_Lean_IR_EmitC_emitCtor(x_1, x_2, x_3, x_4, x_5); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -return x_6; +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_8 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_9 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_declareParams_spec__0(x_1, x_7, x_8, x_4, x_5, x_6); +lean_dec_ref(x_5); +lean_dec_ref(x_1); +return x_9; } } -static lean_object* _init_l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitReset_spec__0___redArg___closed__0() { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_declareParams___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked(" lean_ctor_release(", 19, 19); -return x_1; +lean_object* x_4; +x_4 = l_Lean_IR_EmitC_declareParams(x_1, x_2, x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_4; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitReset_spec__0___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_declareVars(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_5; uint8_t x_6; -x_5 = lean_unsigned_to_nat(0u); -x_6 = lean_nat_dec_eq(x_3, x_5); -if (x_6 == 1) +switch (lean_obj_tag(x_1)) { +case 0: { -lean_object* x_7; lean_object* x_8; -lean_dec(x_3); -lean_dec(x_1); -x_7 = lean_box(0); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_7); -lean_ctor_set(x_8, 1, x_4); -return x_8; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +x_7 = lean_ctor_get(x_1, 3); +lean_inc(x_7); +x_8 = lean_ctor_get(x_3, 3); +x_9 = l_Lean_IR_isTailCallTo(x_8, x_1); +lean_dec_ref(x_1); +if (x_9 == 0) +{ +lean_object* x_10; +x_10 = l_Lean_IR_EmitC_declareVar(x_5, x_6, x_3, x_4); +lean_dec(x_6); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; uint8_t x_12; +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec_ref(x_10); +x_12 = 1; +x_1 = x_7; +x_2 = x_12; +x_4 = x_11; +goto _start; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_9 = lean_unsigned_to_nat(1u); -x_10 = lean_nat_sub(x_3, x_9); -lean_dec(x_3); -x_11 = lean_nat_sub(x_2, x_10); -x_12 = lean_nat_sub(x_11, x_9); -lean_dec(x_11); -x_13 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitReset_spec__0___redArg___closed__0; -x_14 = lean_string_append(x_4, x_13); -x_15 = l_Lean_IR_EmitC_argToCString___closed__0; -lean_inc(x_1); -x_16 = l_Nat_reprFast(x_1); -x_17 = lean_string_append(x_15, x_16); -lean_dec_ref(x_16); -x_18 = lean_string_append(x_14, x_17); -lean_dec_ref(x_17); -x_19 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___closed__0; -x_20 = lean_string_append(x_18, x_19); -x_21 = l_Nat_reprFast(x_12); -x_22 = lean_string_append(x_20, x_21); -lean_dec_ref(x_21); -x_23 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_24 = lean_string_append(x_22, x_23); -x_25 = l_Lean_IR_EmitC_emitLn___closed__0; -x_26 = lean_string_append(x_24, x_25); -x_3 = x_10; -x_4 = x_26; -goto _start; -} -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitReset_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +uint8_t x_14; +lean_dec(x_7); +x_14 = !lean_is_exclusive(x_10); +if (x_14 == 0) { -lean_object* x_7; -x_7 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitReset_spec__0___redArg(x_1, x_2, x_3, x_6); -return x_7; -} +return x_10; } -static lean_object* _init_l_Lean_IR_EmitC_emitReset___closed__0() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("if (lean_is_exclusive(", 22, 22); -return x_1; +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_10, 0); +x_16 = lean_ctor_get(x_10, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_10); +x_17 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +return x_17; } } -static lean_object* _init_l_Lean_IR_EmitC_emitReset___closed__1() { -_start: +} +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked(")) {", 4, 4); -return x_1; +lean_object* x_18; lean_object* x_19; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_18 = lean_box(x_2); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_4); +return x_19; } } -static lean_object* _init_l_Lean_IR_EmitC_emitReset___closed__2() { -_start: +case 1: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked(" lean_dec_ref(", 14, 14); -return x_1; -} +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_20); +x_21 = lean_ctor_get(x_1, 3); +lean_inc(x_21); +lean_dec_ref(x_1); +x_22 = l_Lean_IR_EmitC_declareParams(x_20, x_3, x_4); +if (lean_obj_tag(x_22) == 0) +{ +if (x_2 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec_ref(x_22); +x_24 = lean_unsigned_to_nat(0u); +x_25 = lean_array_get_size(x_20); +lean_dec_ref(x_20); +x_26 = lean_nat_dec_lt(x_24, x_25); +x_1 = x_21; +x_2 = x_26; +x_4 = x_23; +goto _start; } -static lean_object* _init_l_Lean_IR_EmitC_emitReset___closed__3() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_box(0);", 12, 12); -return x_1; +lean_object* x_28; +lean_dec_ref(x_20); +x_28 = lean_ctor_get(x_22, 1); +lean_inc(x_28); +lean_dec_ref(x_22); +x_1 = x_21; +x_4 = x_28; +goto _start; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitReset(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +else { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_6 = l_Lean_IR_EmitC_emitReset___closed__0; -x_7 = lean_string_append(x_5, x_6); -x_8 = l_Lean_IR_EmitC_argToCString___closed__0; -lean_inc(x_3); -x_9 = l_Nat_reprFast(x_3); -x_10 = lean_string_append(x_8, x_9); -lean_dec_ref(x_9); -x_11 = lean_string_append(x_7, x_10); -x_12 = l_Lean_IR_EmitC_emitReset___closed__1; -x_13 = lean_string_append(x_11, x_12); -x_14 = l_Lean_IR_EmitC_emitLn___closed__0; -x_15 = lean_string_append(x_13, x_14); -lean_inc(x_2); -x_16 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitReset_spec__0___redArg(x_3, x_2, x_2, x_15); -lean_dec(x_2); -x_17 = lean_ctor_get(x_16, 1); -lean_inc(x_17); -lean_dec_ref(x_16); -x_18 = l_Lean_IR_EmitC_emitFnDeclAux___closed__4; -x_19 = lean_string_append(x_17, x_18); -lean_inc(x_1); -x_20 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_19); -x_21 = lean_ctor_get(x_20, 1); -lean_inc(x_21); +uint8_t x_30; +lean_dec(x_21); lean_dec_ref(x_20); -x_22 = lean_string_append(x_21, x_10); -x_23 = l_Lean_IR_EmitC_emitFnDeclAux___closed__0; -x_24 = lean_string_append(x_22, x_23); -x_25 = lean_string_append(x_24, x_14); -x_26 = l_Lean_IR_EmitC_emitMainFn___closed__2; -x_27 = lean_string_append(x_25, x_26); -x_28 = lean_string_append(x_27, x_14); -x_29 = l_Lean_IR_EmitC_emitReset___closed__2; -x_30 = lean_string_append(x_28, x_29); -x_31 = lean_string_append(x_30, x_10); -lean_dec_ref(x_10); -x_32 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_33 = lean_string_append(x_31, x_32); -x_34 = lean_string_append(x_33, x_14); -x_35 = lean_string_append(x_34, x_18); -x_36 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_35); -x_37 = !lean_is_exclusive(x_36); -if (x_37 == 0) +x_30 = !lean_is_exclusive(x_22); +if (x_30 == 0) { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_38 = lean_ctor_get(x_36, 1); -x_39 = lean_ctor_get(x_36, 0); -lean_dec(x_39); -x_40 = l_Lean_IR_EmitC_emitReset___closed__3; -x_41 = lean_string_append(x_38, x_40); -x_42 = lean_string_append(x_41, x_14); -x_43 = l_Lean_IR_EmitC_emitMainFn___closed__12; -x_44 = lean_string_append(x_42, x_43); -x_45 = lean_box(0); -x_46 = lean_string_append(x_44, x_14); -lean_ctor_set(x_36, 1, x_46); -lean_ctor_set(x_36, 0, x_45); -return x_36; +return x_22; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_47 = lean_ctor_get(x_36, 1); -lean_inc(x_47); -lean_dec(x_36); -x_48 = l_Lean_IR_EmitC_emitReset___closed__3; -x_49 = lean_string_append(x_47, x_48); -x_50 = lean_string_append(x_49, x_14); -x_51 = l_Lean_IR_EmitC_emitMainFn___closed__12; -x_52 = lean_string_append(x_50, x_51); -x_53 = lean_box(0); -x_54 = lean_string_append(x_52, x_14); -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_53); -lean_ctor_set(x_55, 1, x_54); -return x_55; +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_22, 0); +x_32 = lean_ctor_get(x_22, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_22); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitReset_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +default: { -lean_object* x_7; -x_7 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitReset_spec__0(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec_ref(x_5); -lean_dec(x_2); -return x_7; -} +uint8_t x_34; +x_34 = l_Lean_IR_FnBody_isTerminal(x_1); +if (x_34 == 0) +{ +lean_object* x_35; +x_35 = l_Lean_IR_FnBody_body(x_1); +lean_dec(x_1); +x_1 = x_35; +goto _start; } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitReset_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +else { -lean_object* x_5; -x_5 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitReset_spec__0___redArg(x_1, x_2, x_3, x_4); -lean_dec(x_2); -return x_5; +lean_object* x_37; lean_object* x_38; +lean_dec(x_1); +x_37 = lean_box(x_2); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_4); +return x_38; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitReset___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_IR_EmitC_declareVars___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_6; -x_6 = l_Lean_IR_EmitC_emitReset(x_1, x_2, x_3, x_4, x_5); -lean_dec_ref(x_4); +uint8_t x_5; lean_object* x_6; +x_5 = lean_unbox(x_2); +x_6 = l_Lean_IR_EmitC_declareVars(x_1, x_5, x_3, x_4); +lean_dec_ref(x_3); return x_6; } } -static lean_object* _init_l_Lean_IR_EmitC_emitReuse___closed__0() { +static lean_object* _init_l_Lean_IR_EmitC_emitTag___redArg___closed__0() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("if (lean_is_scalar(", 19, 19); +x_1 = lean_mk_string_unchecked(".tag", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitReuse___closed__1() { +static lean_object* _init_l_Lean_IR_EmitC_emitTag___redArg___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked(" lean_ctor_set_tag(", 19, 19); +x_1 = lean_mk_string_unchecked("lean_obj_tag(", 13, 13); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitReuse(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitTag___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_8; lean_object* x_9; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_16 = l_Lean_IR_EmitC_emitReuse___closed__0; -x_17 = lean_string_append(x_7, x_16); -x_18 = l_Lean_IR_EmitC_argToCString___closed__0; -x_19 = l_Nat_reprFast(x_2); -x_20 = lean_string_append(x_18, x_19); -lean_dec_ref(x_19); -x_21 = lean_string_append(x_17, x_20); -x_22 = l_Lean_IR_EmitC_emitReset___closed__1; -x_23 = lean_string_append(x_21, x_22); -x_24 = l_Lean_IR_EmitC_emitLn___closed__0; -x_25 = lean_string_append(x_23, x_24); -x_26 = l_Lean_IR_EmitC_emitFnDeclAux___closed__4; -x_27 = lean_string_append(x_25, x_26); -lean_inc(x_1); -x_28 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_27); -x_29 = lean_ctor_get(x_28, 1); -lean_inc(x_29); -lean_dec_ref(x_28); -lean_inc_ref(x_3); -x_30 = l_Lean_IR_EmitC_emitAllocCtor___redArg(x_3, x_29); -x_31 = lean_ctor_get(x_30, 1); -lean_inc(x_31); -lean_dec_ref(x_30); -x_32 = l_Lean_IR_EmitC_emitMainFn___closed__2; -x_33 = lean_string_append(x_31, x_32); -x_34 = lean_string_append(x_33, x_24); -x_35 = lean_string_append(x_34, x_26); -lean_inc(x_1); -x_36 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_35); -x_37 = lean_ctor_get(x_36, 1); -lean_inc(x_37); -lean_dec_ref(x_36); -x_38 = lean_string_append(x_37, x_20); -lean_dec_ref(x_20); -x_39 = l_Lean_IR_EmitC_emitFnDeclAux___closed__0; -x_40 = lean_string_append(x_38, x_39); -x_41 = lean_string_append(x_40, x_24); +uint8_t x_4; +x_4 = l_Lean_IR_IRType_isObj(x_2); if (x_4 == 0) { -lean_dec_ref(x_3); -x_8 = x_6; -x_9 = x_41; -goto block_15; +uint8_t x_5; +x_5 = l_Lean_IR_IRType_isStruct(x_2); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_6 = lean_box(0); +x_7 = l_Lean_IR_EmitC_argToCString___closed__0; +x_8 = l_Nat_reprFast(x_1); +x_9 = lean_string_append(x_7, x_8); +lean_dec_ref(x_8); +x_10 = lean_string_append(x_3, x_9); +lean_dec_ref(x_9); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_6); +lean_ctor_set(x_11, 1, x_10); +return x_11; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_42 = lean_ctor_get(x_3, 1); -lean_inc(x_42); -lean_dec_ref(x_3); -x_43 = l_Lean_IR_EmitC_emitReuse___closed__1; -x_44 = lean_string_append(x_41, x_43); -lean_inc(x_1); -x_45 = l_Nat_reprFast(x_1); -x_46 = lean_string_append(x_18, x_45); -lean_dec_ref(x_45); -x_47 = lean_string_append(x_44, x_46); -lean_dec_ref(x_46); -x_48 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___closed__0; -x_49 = lean_string_append(x_47, x_48); -x_50 = l_Nat_reprFast(x_42); -x_51 = lean_string_append(x_49, x_50); -lean_dec_ref(x_50); -x_52 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_53 = lean_string_append(x_51, x_52); -x_54 = lean_string_append(x_53, x_24); -x_8 = x_6; -x_9 = x_54; -goto block_15; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_12 = l_Lean_IR_EmitC_argToCString___closed__0; +x_13 = l_Nat_reprFast(x_1); +x_14 = lean_string_append(x_12, x_13); +lean_dec_ref(x_13); +x_15 = lean_string_append(x_3, x_14); +lean_dec_ref(x_14); +x_16 = l_Lean_IR_EmitC_emitTag___redArg___closed__0; +x_17 = lean_box(0); +x_18 = lean_string_append(x_15, x_16); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +return x_19; } -block_15: +} +else { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_10 = l_Lean_IR_EmitC_emitMainFn___closed__12; -x_11 = lean_string_append(x_9, x_10); -x_12 = l_Lean_IR_EmitC_emitLn___closed__0; -x_13 = lean_string_append(x_11, x_12); -x_14 = l_Lean_IR_EmitC_emitCtorSetArgs(x_1, x_5, x_8, x_13); -return x_14; +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_20 = l_Lean_IR_EmitC_emitTag___redArg___closed__1; +x_21 = lean_string_append(x_3, x_20); +x_22 = l_Lean_IR_EmitC_argToCString___closed__0; +x_23 = l_Nat_reprFast(x_1); +x_24 = lean_string_append(x_22, x_23); +lean_dec_ref(x_23); +x_25 = lean_string_append(x_21, x_24); +lean_dec_ref(x_24); +x_26 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__1; +x_27 = lean_box(0); +x_28 = lean_string_append(x_25, x_26); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; } } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitReuse___boxed(lean_object* x_1, lean_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_IR_EmitC_emitTag(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -uint8_t x_8; lean_object* x_9; -x_8 = lean_unbox(x_4); -x_9 = l_Lean_IR_EmitC_emitReuse(x_1, x_2, x_3, x_8, x_5, x_6, x_7); -lean_dec_ref(x_6); -lean_dec_ref(x_5); -return x_9; +lean_object* x_5; +x_5 = l_Lean_IR_EmitC_emitTag___redArg(x_1, x_2, x_4); +return x_5; } } -static lean_object* _init_l_Lean_IR_EmitC_emitProj___redArg___closed__0() { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitTag___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_ctor_get(", 14, 14); -return x_1; +lean_object* x_5; +x_5 = l_Lean_IR_EmitC_emitTag(x_1, x_2, x_3, x_4); +lean_dec_ref(x_3); +lean_dec(x_2); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitProj___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitTag___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_5; uint8_t x_6; -x_5 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_4); -x_6 = !lean_is_exclusive(x_5); -if (x_6 == 0) +lean_object* x_4; +x_4 = l_Lean_IR_EmitC_emitTag___redArg(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_isIf(lean_object* x_1) { +_start: { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_7 = lean_ctor_get(x_5, 1); -x_8 = lean_ctor_get(x_5, 0); -lean_dec(x_8); -x_9 = l_Lean_IR_EmitC_emitProj___redArg___closed__0; -x_10 = lean_string_append(x_7, x_9); -x_11 = l_Lean_IR_EmitC_argToCString___closed__0; -x_12 = l_Nat_reprFast(x_3); -x_13 = lean_string_append(x_11, x_12); -lean_dec_ref(x_12); -x_14 = lean_string_append(x_10, x_13); -lean_dec_ref(x_13); -x_15 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___closed__0; -x_16 = lean_string_append(x_14, x_15); -x_17 = l_Nat_reprFast(x_2); -x_18 = lean_string_append(x_16, x_17); -lean_dec_ref(x_17); -x_19 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_20 = lean_string_append(x_18, x_19); -x_21 = l_Lean_IR_EmitC_emitLn___closed__0; -x_22 = lean_box(0); -x_23 = lean_string_append(x_20, x_21); -lean_ctor_set(x_5, 1, x_23); -lean_ctor_set(x_5, 0, x_22); +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = lean_array_get_size(x_1); +x_3 = lean_unsigned_to_nat(2u); +x_4 = lean_nat_dec_eq(x_2, x_3); +if (x_4 == 0) +{ +lean_object* x_5; +x_5 = lean_box(0); return x_5; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_24 = lean_ctor_get(x_5, 1); -lean_inc(x_24); -lean_dec(x_5); -x_25 = l_Lean_IR_EmitC_emitProj___redArg___closed__0; -x_26 = lean_string_append(x_24, x_25); -x_27 = l_Lean_IR_EmitC_argToCString___closed__0; -x_28 = l_Nat_reprFast(x_3); -x_29 = lean_string_append(x_27, x_28); -lean_dec_ref(x_28); -x_30 = lean_string_append(x_26, x_29); -lean_dec_ref(x_29); -x_31 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___closed__0; -x_32 = lean_string_append(x_30, x_31); -x_33 = l_Nat_reprFast(x_2); -x_34 = lean_string_append(x_32, x_33); -lean_dec_ref(x_33); -x_35 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_36 = lean_string_append(x_34, x_35); -x_37 = l_Lean_IR_EmitC_emitLn___closed__0; -x_38 = lean_box(0); -x_39 = lean_string_append(x_36, x_37); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; +lean_object* x_6; lean_object* x_7; +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_fget(x_1, x_6); +if (lean_obj_tag(x_7) == 0) +{ +uint8_t x_8; +x_8 = !lean_is_exclusive(x_7); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; 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_7, 0); +x_10 = lean_ctor_get(x_7, 1); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec_ref(x_9); +x_12 = lean_unsigned_to_nat(1u); +x_13 = lean_array_fget_borrowed(x_1, x_12); +x_14 = l_Lean_IR_Alt_body(x_13); +lean_ctor_set(x_7, 1, x_14); +lean_ctor_set(x_7, 0, x_10); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_11); +lean_ctor_set(x_15, 1, x_7); +x_16 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_16, 0, x_15); +return x_16; } +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_17 = lean_ctor_get(x_7, 0); +x_18 = lean_ctor_get(x_7, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_7); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec_ref(x_17); +x_20 = lean_unsigned_to_nat(1u); +x_21 = lean_array_fget_borrowed(x_1, x_20); +x_22 = l_Lean_IR_Alt_body(x_21); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_18); +lean_ctor_set(x_23, 1, x_22); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_19); +lean_ctor_set(x_24, 1, x_23); +x_25 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_25, 0, x_24); +return x_25; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitProj(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +else { -lean_object* x_6; -x_6 = l_Lean_IR_EmitC_emitProj___redArg(x_1, x_2, x_3, x_5); -return x_6; +lean_object* x_26; +lean_dec(x_7); +x_26 = lean_box(0); +return x_26; +} } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitProj___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_IR_EmitC_isIf___boxed(lean_object* x_1) { _start: { -lean_object* x_6; -x_6 = l_Lean_IR_EmitC_emitProj(x_1, x_2, x_3, x_4, x_5); -lean_dec_ref(x_4); -return x_6; +lean_object* x_2; +x_2 = l_Lean_IR_EmitC_isIf(x_1); +lean_dec_ref(x_1); +return x_2; } } -static lean_object* _init_l_Lean_IR_EmitC_emitUProj___redArg___closed__0() { +LEAN_EXPORT uint8_t l_Lean_IR_EmitC_needsRC(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_ctor_get_usize(", 20, 20); -return x_1; +switch (lean_obj_tag(x_1)) { +case 7: +{ +uint8_t x_2; +x_2 = 1; +return x_2; } +case 8: +{ +uint8_t x_3; +x_3 = 1; +return x_3; } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUProj___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +case 11: { -lean_object* x_5; uint8_t x_6; -x_5 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_4); -x_6 = !lean_is_exclusive(x_5); -if (x_6 == 0) +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_4 = lean_ctor_get(x_1, 1); +x_5 = lean_unsigned_to_nat(0u); +x_6 = lean_array_get_size(x_4); +x_7 = lean_nat_dec_lt(x_5, x_6); +if (x_7 == 0) { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_7 = lean_ctor_get(x_5, 1); -x_8 = lean_ctor_get(x_5, 0); -lean_dec(x_8); -x_9 = l_Lean_IR_EmitC_emitUProj___redArg___closed__0; -x_10 = lean_string_append(x_7, x_9); -x_11 = l_Lean_IR_EmitC_argToCString___closed__0; -x_12 = l_Nat_reprFast(x_3); -x_13 = lean_string_append(x_11, x_12); -lean_dec_ref(x_12); -x_14 = lean_string_append(x_10, x_13); -lean_dec_ref(x_13); -x_15 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___closed__0; -x_16 = lean_string_append(x_14, x_15); -x_17 = l_Nat_reprFast(x_2); -x_18 = lean_string_append(x_16, x_17); -lean_dec_ref(x_17); -x_19 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_20 = lean_string_append(x_18, x_19); -x_21 = l_Lean_IR_EmitC_emitLn___closed__0; -x_22 = lean_box(0); -x_23 = lean_string_append(x_20, x_21); -lean_ctor_set(x_5, 1, x_23); -lean_ctor_set(x_5, 0, x_22); -return x_5; +return x_7; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_24 = lean_ctor_get(x_5, 1); -lean_inc(x_24); -lean_dec(x_5); -x_25 = l_Lean_IR_EmitC_emitUProj___redArg___closed__0; -x_26 = lean_string_append(x_24, x_25); -x_27 = l_Lean_IR_EmitC_argToCString___closed__0; -x_28 = l_Nat_reprFast(x_3); -x_29 = lean_string_append(x_27, x_28); -lean_dec_ref(x_28); -x_30 = lean_string_append(x_26, x_29); -lean_dec_ref(x_29); -x_31 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___closed__0; -x_32 = lean_string_append(x_30, x_31); -x_33 = l_Nat_reprFast(x_2); -x_34 = lean_string_append(x_32, x_33); -lean_dec_ref(x_33); -x_35 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_36 = lean_string_append(x_34, x_35); -x_37 = l_Lean_IR_EmitC_emitLn___closed__0; -x_38 = lean_box(0); -x_39 = lean_string_append(x_36, x_37); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; +if (x_7 == 0) +{ +return x_7; +} +else +{ +size_t x_8; size_t x_9; uint8_t x_10; +x_8 = 0; +x_9 = lean_usize_of_nat(x_6); +x_10 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_EmitC_needsRC_spec__0(x_4, x_8, x_9); +return x_10; +} +} +} +case 10: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_ctor_get(x_1, 1); +x_12 = lean_unsigned_to_nat(0u); +x_13 = lean_array_get_size(x_11); +x_14 = lean_nat_dec_lt(x_12, x_13); +if (x_14 == 0) +{ +return x_14; +} +else +{ +if (x_14 == 0) +{ +return x_14; +} +else +{ +size_t x_15; size_t x_16; uint8_t x_17; +x_15 = 0; +x_16 = lean_usize_of_nat(x_13); +x_17 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_EmitC_needsRC_spec__0(x_11, x_15, x_16); +return x_17; +} +} } +default: +{ +uint8_t x_18; +x_18 = 0; +return x_18; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUProj(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +} +} +LEAN_EXPORT uint8_t l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_EmitC_needsRC_spec__0(lean_object* x_1, size_t x_2, size_t x_3) { _start: { -lean_object* x_6; -x_6 = l_Lean_IR_EmitC_emitUProj___redArg(x_1, x_2, x_3, x_5); +uint8_t x_4; +x_4 = lean_usize_dec_eq(x_2, x_3); +if (x_4 == 0) +{ +lean_object* x_5; uint8_t x_6; +x_5 = lean_array_uget(x_1, x_2); +x_6 = l_Lean_IR_EmitC_needsRC(x_5); +lean_dec(x_5); +if (x_6 == 0) +{ +size_t x_7; size_t x_8; +x_7 = 1; +x_8 = lean_usize_add(x_2, x_7); +x_2 = x_8; +goto _start; +} +else +{ return x_6; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUProj___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +else +{ +uint8_t x_10; +x_10 = 0; +return x_10; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_EmitC_needsRC_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_6; -x_6 = l_Lean_IR_EmitC_emitUProj(x_1, x_2, x_3, x_4, x_5); -lean_dec_ref(x_4); -return x_6; +size_t x_4; size_t x_5; uint8_t x_6; lean_object* x_7; +x_4 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_5 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_6 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_EmitC_needsRC_spec__0(x_1, x_4, x_5); +lean_dec_ref(x_1); +x_7 = lean_box(x_6); +return x_7; } } -static lean_object* _init_l_Lean_IR_EmitC_emitSProj___redArg___closed__0() { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_needsRC___boxed(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_ctor_get_float", 19, 19); -return x_1; +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_IR_EmitC_needsRC(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; } } -static lean_object* _init_l_Lean_IR_EmitC_emitSProj___redArg___closed__1() { +static lean_object* _init_l_Lean_IR_EmitC_structIncFnPrefix___closed__0() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_ctor_get_float32", 21, 21); +x_1 = lean_mk_string_unchecked("_l_struct_inc_", 14, 14); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitSProj___redArg___closed__2() { +static lean_object* _init_l_Lean_IR_EmitC_structIncFnPrefix() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_ctor_get_uint8", 19, 19); +x_1 = l_Lean_IR_EmitC_structIncFnPrefix___closed__0; return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitSProj___redArg___closed__3() { +static lean_object* _init_l_Lean_IR_EmitC_structDecFnPrefix___closed__0() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_ctor_get_uint16", 20, 20); +x_1 = lean_mk_string_unchecked("_l_struct_dec_", 14, 14); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitSProj___redArg___closed__4() { +static lean_object* _init_l_Lean_IR_EmitC_structDecFnPrefix() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_ctor_get_uint32", 20, 20); +x_1 = l_Lean_IR_EmitC_structDecFnPrefix___closed__0; return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitSProj___redArg___closed__5() { +static lean_object* _init_l_Lean_IR_EmitC_structReshapeFnPrefix___closed__0() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_ctor_get_uint64", 20, 20); +x_1 = lean_mk_string_unchecked("_l_struct_reshape_", 18, 18); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSProj___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +static lean_object* _init_l_Lean_IR_EmitC_structReshapeFnPrefix() { _start: { -lean_object* x_7; lean_object* x_33; -x_33 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_6); -switch (lean_obj_tag(x_2)) { -case 0: -{ -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_33, 1); -lean_inc(x_34); -lean_dec_ref(x_33); -x_35 = l_Lean_IR_EmitC_emitSProj___redArg___closed__0; -x_36 = lean_string_append(x_34, x_35); -x_7 = x_36; -goto block_32; -} -case 9: -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_33, 1); -lean_inc(x_37); -lean_dec_ref(x_33); -x_38 = l_Lean_IR_EmitC_emitSProj___redArg___closed__1; -x_39 = lean_string_append(x_37, x_38); -x_7 = x_39; -goto block_32; +lean_object* x_1; +x_1 = l_Lean_IR_EmitC_structReshapeFnPrefix___closed__0; +return x_1; } -case 1: -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_33, 1); -lean_inc(x_40); -lean_dec_ref(x_33); -x_41 = l_Lean_IR_EmitC_emitSProj___redArg___closed__2; -x_42 = lean_string_append(x_40, x_41); -x_7 = x_42; -goto block_32; } -case 2: +static lean_object* _init_l_Lean_IR_EmitC_structBoxFnPrefix___closed__0() { +_start: { -lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_43 = lean_ctor_get(x_33, 1); -lean_inc(x_43); -lean_dec_ref(x_33); -x_44 = l_Lean_IR_EmitC_emitSProj___redArg___closed__3; -x_45 = lean_string_append(x_43, x_44); -x_7 = x_45; -goto block_32; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_l_struct_box_", 14, 14); +return x_1; } -case 3: -{ -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_33, 1); -lean_inc(x_46); -lean_dec_ref(x_33); -x_47 = l_Lean_IR_EmitC_emitSProj___redArg___closed__4; -x_48 = lean_string_append(x_46, x_47); -x_7 = x_48; -goto block_32; } -case 4: +static lean_object* _init_l_Lean_IR_EmitC_structBoxFnPrefix() { +_start: { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_33, 1); -lean_inc(x_49); -lean_dec_ref(x_33); -x_50 = l_Lean_IR_EmitC_emitSProj___redArg___closed__5; -x_51 = lean_string_append(x_49, x_50); -x_7 = x_51; -goto block_32; +lean_object* x_1; +x_1 = l_Lean_IR_EmitC_structBoxFnPrefix___closed__0; +return x_1; } -default: -{ -uint8_t x_52; -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_52 = !lean_is_exclusive(x_33); -if (x_52 == 0) -{ -lean_object* x_53; lean_object* x_54; -x_53 = lean_ctor_get(x_33, 0); -lean_dec(x_53); -x_54 = l_Lean_IR_EmitC_emitSSet___redArg___closed__6; -lean_ctor_set_tag(x_33, 1); -lean_ctor_set(x_33, 0, x_54); -return x_33; } -else +static lean_object* _init_l_Lean_IR_EmitC_structUnboxFnPrefix___closed__0() { +_start: { -lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_55 = lean_ctor_get(x_33, 1); -lean_inc(x_55); -lean_dec(x_33); -x_56 = l_Lean_IR_EmitC_emitSSet___redArg___closed__6; -x_57 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_55); -return x_57; -} +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_l_struct_unbox_", 16, 16); +return x_1; } } -block_32: -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_8 = l_Lean_IR_EmitC_emitFnDeclAux___closed__5; -x_9 = lean_string_append(x_7, x_8); -x_10 = l_Lean_IR_EmitC_argToCString___closed__0; -x_11 = l_Nat_reprFast(x_5); -x_12 = lean_string_append(x_10, x_11); -lean_dec_ref(x_11); -x_13 = lean_string_append(x_9, x_12); -lean_dec_ref(x_12); -x_14 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___closed__0; -x_15 = lean_string_append(x_13, x_14); -x_16 = l_Lean_IR_EmitC_emitOffset___redArg(x_3, x_4, x_15); -x_17 = !lean_is_exclusive(x_16); -if (x_17 == 0) +static lean_object* _init_l_Lean_IR_EmitC_structUnboxFnPrefix() { +_start: { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_18 = lean_ctor_get(x_16, 1); -x_19 = lean_ctor_get(x_16, 0); -lean_dec(x_19); -x_20 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_21 = lean_string_append(x_18, x_20); -x_22 = l_Lean_IR_EmitC_emitLn___closed__0; -x_23 = lean_box(0); -x_24 = lean_string_append(x_21, x_22); -lean_ctor_set(x_16, 1, x_24); -lean_ctor_set(x_16, 0, x_23); -return x_16; +lean_object* x_1; +x_1 = l_Lean_IR_EmitC_structUnboxFnPrefix___closed__0; +return x_1; } -else +} +static lean_object* _init_l_Lean_IR_EmitC_emitIncOfType___closed__0() { +_start: { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_25 = lean_ctor_get(x_16, 1); -lean_inc(x_25); -lean_dec(x_16); -x_26 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_27 = lean_string_append(x_25, x_26); -x_28 = l_Lean_IR_EmitC_emitLn___closed__0; -x_29 = lean_box(0); -x_30 = lean_string_append(x_27, x_28); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_30); -return x_31; +lean_object* x_1; +x_1 = lean_mk_string_unchecked(");", 2, 2); +return x_1; } } +static lean_object* _init_l_Lean_IR_EmitC_emitIncOfType___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_inc_ref_n", 14, 14); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSProj(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +static lean_object* _init_l_Lean_IR_EmitC_emitIncOfType___closed__2() { _start: { -lean_object* x_8; -x_8 = l_Lean_IR_EmitC_emitSProj___redArg(x_1, x_2, x_3, x_4, x_5, x_7); -return x_8; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_inc_ref", 12, 12); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSProj___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +static lean_object* _init_l_Lean_IR_EmitC_emitIncOfType___closed__3() { _start: { -lean_object* x_8; -x_8 = l_Lean_IR_EmitC_emitSProj(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec_ref(x_6); -lean_dec(x_2); -return x_8; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_inc_n", 10, 10); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSProj___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +static lean_object* _init_l_Lean_IR_EmitC_emitIncOfType___closed__4() { _start: { -lean_object* x_7; -x_7 = l_Lean_IR_EmitC_emitSProj___redArg(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_2); -return x_7; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_inc", 8, 8); +return x_1; } } -LEAN_EXPORT lean_object* l_List_mapTR_loop___at___00Lean_IR_EmitC_toStringArgs_spec__0(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitIncOfType(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -if (lean_obj_tag(x_1) == 0) +lean_object* x_8; lean_object* x_16; uint8_t x_21; lean_object* x_22; +x_21 = l_Lean_IR_IRType_isStruct(x_2); +if (x_21 == 0) { -lean_object* x_3; -x_3 = l_List_reverse___redArg(x_2); -return x_3; +if (x_4 == 0) +{ +lean_object* x_30; uint8_t x_31; +x_30 = lean_unsigned_to_nat(1u); +x_31 = lean_nat_dec_eq(x_3, x_30); +if (x_31 == 0) +{ +lean_object* x_32; +x_32 = l_Lean_IR_EmitC_emitIncOfType___closed__1; +x_22 = x_32; +goto block_29; } else { -uint8_t x_4; -x_4 = !lean_is_exclusive(x_1); -if (x_4 == 0) -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = lean_ctor_get(x_1, 0); -x_6 = lean_ctor_get(x_1, 1); -x_7 = l_Lean_IR_EmitC_argToCString(x_5); -lean_ctor_set(x_1, 1, x_2); -lean_ctor_set(x_1, 0, x_7); -{ -lean_object* _tmp_0 = x_6; -lean_object* _tmp_1 = x_1; -x_1 = _tmp_0; -x_2 = _tmp_1; +lean_object* x_33; +x_33 = l_Lean_IR_EmitC_emitIncOfType___closed__2; +x_22 = x_33; +goto block_29; } -goto _start; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_ctor_get(x_1, 0); -x_10 = lean_ctor_get(x_1, 1); -lean_inc(x_10); -lean_inc(x_9); -lean_dec(x_1); -x_11 = l_Lean_IR_EmitC_argToCString(x_9); -x_12 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_2); -x_1 = x_10; -x_2 = x_12; -goto _start; -} -} -} +lean_object* x_34; uint8_t x_35; +x_34 = lean_unsigned_to_nat(1u); +x_35 = lean_nat_dec_eq(x_3, x_34); +if (x_35 == 0) +{ +lean_object* x_36; +x_36 = l_Lean_IR_EmitC_emitIncOfType___closed__3; +x_22 = x_36; +goto block_29; } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_toStringArgs(lean_object* x_1) { -_start: +else { -lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = lean_array_to_list(x_1); -x_3 = lean_box(0); -x_4 = l_List_mapTR_loop___at___00Lean_IR_EmitC_toStringArgs_spec__0(x_2, x_3); -return x_4; +lean_object* x_37; +x_37 = l_Lean_IR_EmitC_emitIncOfType___closed__4; +x_22 = x_37; +goto block_29; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___at___00Lean_IR_EmitC_emitSimpleExternalCall_spec__0___redArg(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) { -_start: -{ -lean_object* x_8; uint8_t x_9; -x_8 = lean_unsigned_to_nat(0u); -x_9 = lean_nat_dec_eq(x_5, x_8); -if (x_9 == 1) -{ -lean_object* x_10; lean_object* x_11; -lean_dec(x_5); -lean_dec_ref(x_2); -x_10 = lean_box(x_6); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_10); -lean_ctor_set(x_11, 1, x_7); -return x_11; } else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; uint8_t x_23; lean_object* x_28; lean_object* x_29; uint8_t x_30; -x_12 = lean_unsigned_to_nat(1u); -x_13 = lean_nat_sub(x_5, x_12); -lean_dec(x_5); -x_14 = lean_nat_sub(x_4, x_13); -x_15 = lean_nat_sub(x_14, x_12); -lean_dec(x_14); -lean_inc_ref(x_2); -x_28 = lean_array_get_borrowed(x_2, x_3, x_15); -x_29 = lean_ctor_get(x_28, 1); -x_30 = l_Lean_IR_IRType_isErased(x_29); -if (x_30 == 0) +uint8_t x_38; +x_38 = l_Lean_IR_EmitC_needsRC(x_2); +if (x_38 == 0) { -uint8_t x_31; -x_31 = l_Lean_IR_IRType_isVoid(x_29); -x_23 = x_31; -goto block_27; +lean_object* x_39; lean_object* x_40; +x_39 = lean_box(0); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_7); +return x_40; } else { -x_23 = x_30; -goto block_27; +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_41 = lean_ctor_get(x_6, 6); +x_42 = lean_unsigned_to_nat(0u); +x_43 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0___redArg(x_42, x_41, x_2); +x_44 = l_Lean_IR_EmitC_structIncFnPrefix___closed__0; +x_45 = lean_string_append(x_7, x_44); +x_46 = l_Nat_reprFast(x_43); +x_47 = lean_string_append(x_45, x_46); +lean_dec_ref(x_46); +x_48 = l_Lean_IR_EmitC_emitFnDeclAux___closed__2; +x_49 = lean_string_append(x_47, x_48); +x_50 = lean_string_append(x_49, x_1); +x_51 = l_Lean_IR_EmitC_emitSpreadArg___closed__1; +x_52 = lean_string_append(x_50, x_51); +x_53 = lean_string_append(x_52, x_5); +x_54 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_55 = lean_string_append(x_53, x_54); +x_56 = l_Lean_IR_EmitC_emitLn___closed__0; +x_57 = lean_box(0); +x_58 = lean_string_append(x_55, x_56); +x_59 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_58); +return x_59; } -block_22: +} +block_15: { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_array_fget_borrowed(x_1, x_15); -lean_dec(x_15); -lean_inc(x_18); -x_19 = l_Lean_IR_EmitC_emitArg___redArg(x_18, x_17); -x_20 = lean_ctor_get(x_19, 1); -lean_inc(x_20); -lean_dec_ref(x_19); -x_5 = x_13; -x_6 = x_16; -x_7 = x_20; -goto _start; +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_9 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_10 = lean_string_append(x_8, x_9); +x_11 = l_Lean_IR_EmitC_emitLn___closed__0; +x_12 = lean_box(0); +x_13 = lean_string_append(x_10, x_11); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +return x_14; } -block_27: +block_20: { -if (x_23 == 0) +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = l_Lean_IR_EmitC_emitSpreadArg___closed__1; +x_18 = lean_string_append(x_16, x_17); +x_19 = lean_string_append(x_18, x_5); +x_8 = x_19; +goto block_15; +} +block_29: { -if (x_6 == 0) +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_23 = lean_string_append(x_7, x_22); +lean_dec_ref(x_22); +x_24 = l_Lean_IR_EmitC_emitFnDeclAux___closed__2; +x_25 = lean_string_append(x_23, x_24); +x_26 = lean_string_append(x_25, x_1); +x_27 = lean_unsigned_to_nat(1u); +x_28 = lean_nat_dec_eq(x_3, x_27); +if (x_28 == 0) { -lean_object* x_24; lean_object* x_25; -x_24 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___closed__0; -x_25 = lean_string_append(x_7, x_24); -x_16 = x_23; -x_17 = x_25; -goto block_22; +x_16 = x_26; +goto block_20; } else { -x_16 = x_23; -x_17 = x_7; -goto block_22; -} +if (x_21 == 0) +{ +x_8 = x_26; +goto block_15; } else { -lean_dec(x_15); -x_5 = x_13; -goto _start; -} -} +x_16 = x_26; +goto block_20; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___at___00Lean_IR_EmitC_emitSimpleExternalCall_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; -x_10 = l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___at___00Lean_IR_EmitC_emitSimpleExternalCall_spec__0___redArg(x_1, x_2, x_3, x_4, x_5, x_7, x_9); -return x_10; } } -static lean_object* _init_l_Lean_IR_EmitC_emitSimpleExternalCall___closed__0() { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitIncOfType___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_1; -x_1 = l_Lean_IR_instInhabitedParam_default; -return x_1; +uint8_t x_8; lean_object* x_9; +x_8 = lean_unbox(x_4); +x_9 = l_Lean_IR_EmitC_emitIncOfType(x_1, x_2, x_3, x_8, x_5, x_6, x_7); +lean_dec_ref(x_6); +lean_dec_ref(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec_ref(x_1); +return x_9; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSimpleExternalCall(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_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0_spec__0___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; -x_6 = l_Lean_IR_EmitC_emitSimpleExternalCall___closed__0; -x_7 = lean_string_append(x_5, x_1); -x_8 = l_Lean_IR_EmitC_emitFnDeclAux___closed__5; -x_9 = lean_string_append(x_7, x_8); -x_10 = lean_array_get_size(x_3); -x_11 = 1; -x_12 = l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___at___00Lean_IR_EmitC_emitSimpleExternalCall_spec__0___redArg(x_3, x_6, x_2, x_10, x_10, x_11, x_9); -if (lean_obj_tag(x_12) == 0) -{ -uint8_t x_13; -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_14 = lean_ctor_get(x_12, 1); -x_15 = lean_ctor_get(x_12, 0); -lean_dec(x_15); -x_16 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_17 = lean_string_append(x_14, x_16); -x_18 = l_Lean_IR_EmitC_emitLn___closed__0; -x_19 = lean_string_append(x_17, x_18); -x_20 = lean_box(0); -lean_ctor_set(x_12, 1, x_19); -lean_ctor_set(x_12, 0, x_20); -return x_12; -} -else +if (lean_obj_tag(x_3) == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_21 = lean_ctor_get(x_12, 1); -lean_inc(x_21); -lean_dec(x_12); -x_22 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_23 = lean_string_append(x_21, x_22); -x_24 = l_Lean_IR_EmitC_emitLn___closed__0; -x_25 = lean_string_append(x_23, x_24); -x_26 = lean_box(0); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_25); -return x_27; -} +lean_object* x_4; lean_object* x_5; +x_4 = l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg___closed__3; +x_5 = lean_panic_fn(x_1, x_4); +return x_5; } else { -uint8_t x_28; -x_28 = !lean_is_exclusive(x_12); -if (x_28 == 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 = l_Lean_IR_instBEqVarId_beq(x_6, x_2); +if (x_9 == 0) { -return x_12; +x_3 = x_8; +goto _start; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_12, 0); -x_30 = lean_ctor_get(x_12, 1); -lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_12); -x_31 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_30); -return x_31; +lean_dec(x_1); +lean_inc(x_7); +return x_7; } } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___at___00Lean_IR_EmitC_emitSimpleExternalCall_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -uint8_t x_10; lean_object* x_11; -x_10 = lean_unbox(x_7); -x_11 = l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___at___00Lean_IR_EmitC_emitSimpleExternalCall_spec__0(x_1, x_2, x_3, x_4, x_5, x_6, x_10, x_8, x_9); -lean_dec_ref(x_8); -lean_dec(x_4); -lean_dec_ref(x_3); -lean_dec_ref(x_1); -return x_11; +lean_object* x_5; +x_5 = l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0_spec__0___redArg(x_2, x_3, x_4); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSimpleExternalCall___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_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_6; -x_6 = l_Lean_IR_EmitC_emitSimpleExternalCall(x_1, x_2, x_3, x_4, x_5); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -return x_6; +lean_object* x_4; lean_object* x_5; uint64_t 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; size_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_array_get_size(x_4); +x_6 = l_Lean_IR_instHashableVarId_hash(x_3); +x_7 = 32; +x_8 = lean_uint64_shift_right(x_6, x_7); +x_9 = lean_uint64_xor(x_6, x_8); +x_10 = 16; +x_11 = lean_uint64_shift_right(x_9, x_10); +x_12 = lean_uint64_xor(x_9, x_11); +x_13 = lean_uint64_to_usize(x_12); +x_14 = lean_usize_of_nat(x_5); +x_15 = 1; +x_16 = lean_usize_sub(x_14, x_15); +x_17 = lean_usize_land(x_13, x_16); +x_18 = lean_array_uget(x_4, x_17); +x_19 = l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0_spec__0___redArg(x_1, x_3, x_18); +lean_dec(x_18); +return x_19; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___at___00Lean_IR_EmitC_emitSimpleExternalCall_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -uint8_t x_8; lean_object* x_9; -x_8 = lean_unbox(x_6); -x_9 = l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___at___00Lean_IR_EmitC_emitSimpleExternalCall_spec__0___redArg(x_1, x_2, x_3, x_4, x_5, x_8, x_7); -lean_dec(x_4); -lean_dec_ref(x_3); -lean_dec_ref(x_1); -return x_9; +lean_object* x_5; +x_5 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0___redArg(x_2, x_3, x_4); +return x_5; } } -static lean_object* _init_l_Lean_IR_EmitC_emitExternCall___closed__0() { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitInc(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("failed to emit extern application '", 35, 35); -return x_1; +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_6 = lean_ctor_get(x_4, 5); +x_7 = lean_box(0); +x_8 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0___redArg(x_7, x_6, x_1); +x_9 = l_Lean_IR_EmitC_argToCString___closed__0; +x_10 = l_Nat_reprFast(x_1); +x_11 = lean_string_append(x_9, x_10); +lean_dec_ref(x_10); +lean_inc(x_2); +x_12 = l_Nat_reprFast(x_2); +x_13 = l_Lean_IR_EmitC_emitIncOfType(x_11, x_8, x_2, x_3, x_12, x_4, x_5); +lean_dec_ref(x_12); +lean_dec(x_2); +lean_dec(x_8); +lean_dec_ref(x_11); +return x_13; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitExternCall(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_7; lean_object* x_16; lean_object* x_17; -x_16 = l_List_forM___at___00Lean_IR_EmitC_emitFnDecls_spec__3___closed__1; -x_17 = l_Lean_getExternEntryFor(x_3, x_16); -if (lean_obj_tag(x_17) == 1) -{ -lean_object* x_18; -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -lean_dec_ref(x_17); -switch (lean_obj_tag(x_18)) { -case 2: -{ -lean_object* x_19; lean_object* x_20; -lean_dec(x_1); -x_19 = lean_ctor_get(x_18, 1); -lean_inc_ref(x_19); -lean_dec_ref(x_18); -x_20 = l_Lean_IR_EmitC_emitSimpleExternalCall(x_19, x_2, x_4, x_5, x_6); -lean_dec_ref(x_4); -lean_dec_ref(x_19); -return x_20; +lean_object* x_5; +x_5 = l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0_spec__0(x_1, x_2, x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +return x_5; } -case 1: -{ -uint8_t x_21; -lean_dec(x_1); -x_21 = !lean_is_exclusive(x_18); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_22 = lean_ctor_get(x_18, 1); -x_23 = lean_ctor_get(x_18, 0); -lean_dec(x_23); -x_24 = l_Lean_IR_EmitC_toStringArgs(x_4); -x_25 = l_Lean_expandExternPattern(x_22, x_24); -lean_dec(x_24); -x_26 = lean_string_append(x_6, x_25); -lean_dec_ref(x_25); -x_27 = l_Lean_IR_EmitC_emitFnDeclAux___closed__0; -x_28 = lean_string_append(x_26, x_27); -x_29 = l_Lean_IR_EmitC_emitLn___closed__0; -x_30 = lean_box(0); -x_31 = lean_string_append(x_28, x_29); -lean_ctor_set_tag(x_18, 0); -lean_ctor_set(x_18, 1, x_31); -lean_ctor_set(x_18, 0, x_30); -return x_18; } -else +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_32 = lean_ctor_get(x_18, 1); -lean_inc(x_32); -lean_dec(x_18); -x_33 = l_Lean_IR_EmitC_toStringArgs(x_4); -x_34 = l_Lean_expandExternPattern(x_32, x_33); -lean_dec(x_33); -x_35 = lean_string_append(x_6, x_34); -lean_dec_ref(x_34); -x_36 = l_Lean_IR_EmitC_emitFnDeclAux___closed__0; -x_37 = lean_string_append(x_35, x_36); -x_38 = l_Lean_IR_EmitC_emitLn___closed__0; -x_39 = lean_box(0); -x_40 = lean_string_append(x_37, x_38); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); -return x_41; +lean_object* x_5; +x_5 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0(x_1, x_2, x_3, x_4); +lean_dec(x_4); +lean_dec_ref(x_3); +return x_5; } } -default: +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitInc___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -lean_dec(x_18); +uint8_t x_6; lean_object* x_7; +x_6 = lean_unbox(x_3); +x_7 = l_Lean_IR_EmitC_emitInc(x_1, x_2, x_6, x_4, x_5); lean_dec_ref(x_4); -x_7 = x_6; -goto block_15; -} +return x_7; } } -else +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_dec(x_17); -lean_dec_ref(x_4); -x_7 = x_6; -goto block_15; +lean_object* x_4; +x_4 = l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0_spec__0___redArg(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +return x_4; } -block_15: -{ -lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_8 = l_Lean_IR_EmitC_emitExternCall___closed__0; -x_9 = 1; -x_10 = l_Lean_Name_toStringWithToken___at___00Lean_Name_toString_spec__0(x_1, x_9); -x_11 = lean_string_append(x_8, x_10); -lean_dec_ref(x_10); -x_12 = l_Lean_IR_EmitC_getDecl___closed__1; -x_13 = lean_string_append(x_11, x_12); -x_14 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_7); -return x_14; } +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0___redArg(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec_ref(x_2); +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitExternCall___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +static lean_object* _init_l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDecOfType_spec__0___redArg___closed__0() { _start: { -lean_object* x_7; -x_7 = l_Lean_IR_EmitC_emitExternCall(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec_ref(x_5); -lean_dec_ref(x_2); -return x_7; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_dec_ref", 12, 12); +return x_1; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFullApp_spec__0(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +static lean_object* _init_l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDecOfType_spec__0___redArg___closed__1() { _start: { -lean_object* x_5; uint8_t x_10; -x_10 = lean_usize_dec_eq(x_2, x_3); -if (x_10 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_dec", 8, 8); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDecOfType_spec__0___redArg(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_11 = lean_array_uget(x_1, x_2); -x_12 = lean_ctor_get(x_11, 1); -lean_inc(x_12); -x_13 = lean_ctor_get(x_12, 1); -lean_inc(x_13); -lean_dec(x_12); -x_14 = l_Lean_IR_IRType_isVoid(x_13); -lean_dec(x_13); -if (x_14 == 0) +lean_object* x_5; uint8_t x_6; +x_5 = lean_unsigned_to_nat(0u); +x_6 = lean_nat_dec_eq(x_3, x_5); +if (x_6 == 1) { -lean_object* x_15; -x_15 = lean_array_push(x_4, x_11); -x_5 = x_15; -goto block_9; +lean_object* x_7; lean_object* x_8; +lean_dec(x_3); +x_7 = lean_box(0); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_4); +return x_8; } else { -lean_dec(x_11); -x_5 = x_4; -goto block_9; -} +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = lean_unsigned_to_nat(1u); +x_10 = lean_nat_sub(x_3, x_9); +lean_dec(x_3); +if (x_2 == 0) +{ +lean_object* x_22; +x_22 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDecOfType_spec__0___redArg___closed__0; +x_11 = x_22; +goto block_21; } else { -return x_4; +lean_object* x_23; +x_23 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDecOfType_spec__0___redArg___closed__1; +x_11 = x_23; +goto block_21; } -block_9: +block_21: { -size_t x_6; size_t x_7; -x_6 = 1; -x_7 = lean_usize_add(x_2, x_6); -x_2 = x_7; -x_4 = x_5; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_12 = lean_string_append(x_4, x_11); +lean_dec_ref(x_11); +x_13 = l_Lean_IR_EmitC_emitFnDeclAux___closed__2; +x_14 = lean_string_append(x_12, x_13); +x_15 = lean_string_append(x_14, x_1); +x_16 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_17 = lean_string_append(x_15, x_16); +x_18 = l_Lean_IR_EmitC_emitLn___closed__0; +x_19 = lean_string_append(x_17, x_18); +x_3 = x_10; +x_4 = x_19; goto _start; } } } -static lean_object* _init_l_Lean_IR_EmitC_emitFullApp___closed__0() { +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDecOfType_spec__0(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_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_object* x_8; +x_8 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDecOfType_spec__0___redArg(x_1, x_2, x_4, x_7); +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFullApp(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_IR_EmitC_emitDecOfType(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_6; lean_object* x_14; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_5); -x_23 = lean_ctor_get(x_22, 1); -lean_inc(x_23); -lean_dec_ref(x_22); -lean_inc_ref(x_4); -lean_inc(x_2); -x_24 = l_Lean_IR_EmitC_getDecl(x_2, x_4, x_23); -if (lean_obj_tag(x_24) == 0) -{ -lean_object* x_25; -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -if (lean_obj_tag(x_25) == 0) +uint8_t x_7; +x_7 = l_Lean_IR_IRType_isStruct(x_2); +if (x_7 == 0) { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_24, 1); -lean_inc(x_26); -lean_dec_ref(x_24); -x_27 = lean_ctor_get(x_25, 1); -lean_inc_ref(x_27); -lean_dec_ref(x_25); -lean_inc_ref(x_4); -x_28 = l_Lean_IR_EmitC_emitCName(x_2, x_4, x_26); -if (lean_obj_tag(x_28) == 0) +lean_object* x_8; +x_8 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDecOfType_spec__0___redArg(x_1, x_4, x_3, x_6); +return x_8; +} +else { -lean_object* x_29; lean_object* x_30; lean_object* x_40; lean_object* x_41; uint8_t x_42; -x_29 = lean_ctor_get(x_28, 1); -lean_inc(x_29); -lean_dec_ref(x_28); -x_40 = lean_unsigned_to_nat(0u); -x_41 = lean_array_get_size(x_3); -x_42 = lean_nat_dec_lt(x_40, x_41); -if (x_42 == 0) +uint8_t x_9; +lean_dec(x_3); +x_9 = l_Lean_IR_EmitC_needsRC(x_2); +if (x_9 == 0) { -lean_dec_ref(x_27); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -x_14 = x_29; -goto block_21; +lean_object* x_10; lean_object* x_11; +x_10 = lean_box(0); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_6); +return x_11; } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; -x_43 = l_Array_zip___redArg(x_3, x_27); -lean_dec_ref(x_27); -lean_dec_ref(x_3); -x_44 = lean_array_get_size(x_43); -x_45 = l_Lean_IR_EmitC_emitFullApp___closed__0; -x_46 = lean_nat_dec_lt(x_40, x_44); -if (x_46 == 0) -{ -lean_dec_ref(x_43); -x_30 = x_45; -goto block_39; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_12 = lean_ctor_get(x_5, 6); +x_13 = lean_unsigned_to_nat(0u); +x_14 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0___redArg(x_13, x_12, x_2); +x_15 = l_Lean_IR_EmitC_structDecFnPrefix___closed__0; +x_16 = lean_string_append(x_6, x_15); +x_17 = l_Nat_reprFast(x_14); +x_18 = lean_string_append(x_16, x_17); +lean_dec_ref(x_17); +x_19 = l_Lean_IR_EmitC_emitFnDeclAux___closed__2; +x_20 = lean_string_append(x_18, x_19); +x_21 = lean_string_append(x_20, x_1); +x_22 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_23 = lean_string_append(x_21, x_22); +x_24 = l_Lean_IR_EmitC_emitLn___closed__0; +x_25 = lean_box(0); +x_26 = lean_string_append(x_23, x_24); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; } -else -{ -uint8_t x_47; -x_47 = lean_nat_dec_le(x_44, x_44); -if (x_47 == 0) -{ -if (x_46 == 0) -{ -lean_dec_ref(x_43); -x_30 = x_45; -goto block_39; } -else -{ -size_t x_48; size_t x_49; lean_object* x_50; -x_48 = 0; -x_49 = lean_usize_of_nat(x_44); -x_50 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFullApp_spec__0(x_43, x_48, x_49, x_45); -lean_dec_ref(x_43); -x_30 = x_50; -goto block_39; } } -else +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDecOfType_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: { -size_t x_51; size_t x_52; lean_object* x_53; -x_51 = 0; -x_52 = lean_usize_of_nat(x_44); -x_53 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFullApp_spec__0(x_43, x_51, x_52, x_45); -lean_dec_ref(x_43); -x_30 = x_53; -goto block_39; -} +uint8_t x_8; lean_object* x_9; +x_8 = lean_unbox(x_2); +x_9 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDecOfType_spec__0(x_1, x_8, x_3, x_4, x_5, x_6, x_7); +lean_dec_ref(x_6); +lean_dec(x_3); +lean_dec_ref(x_1); +return x_9; } } -block_39: +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDecOfType_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_31 = l_Array_unzip___redArg(x_30); -lean_dec_ref(x_30); -x_32 = lean_ctor_get(x_31, 0); -lean_inc(x_32); -lean_dec_ref(x_31); -x_33 = l_Lean_IR_EmitC_emitFnDeclAux___closed__5; -x_34 = lean_string_append(x_29, x_33); -x_35 = l_Lean_IR_EmitC_emitArgs(x_32, x_4, x_34); -lean_dec_ref(x_4); -lean_dec(x_32); -x_36 = lean_ctor_get(x_35, 1); -lean_inc(x_36); -lean_dec_ref(x_35); -x_37 = l_Lean_IR_EmitC_emitFnDeclAux___closed__1; -x_38 = lean_string_append(x_36, x_37); -x_14 = x_38; -goto block_21; +uint8_t x_5; lean_object* x_6; +x_5 = lean_unbox(x_2); +x_6 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDecOfType_spec__0___redArg(x_1, x_5, x_3, x_4); +lean_dec_ref(x_1); +return x_6; } } -else +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitDecOfType___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -lean_dec_ref(x_27); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -return x_28; +uint8_t x_7; lean_object* x_8; +x_7 = lean_unbox(x_4); +x_8 = l_Lean_IR_EmitC_emitDecOfType(x_1, x_2, x_3, x_7, x_5, x_6); +lean_dec_ref(x_5); +lean_dec(x_2); +lean_dec_ref(x_1); +return x_8; } } -else -{ -lean_object* x_54; -x_54 = lean_ctor_get(x_25, 3); -if (lean_obj_tag(x_54) == 1) -{ -lean_object* x_55; -x_55 = lean_ctor_get(x_54, 0); -if (lean_obj_tag(x_55) == 3) -{ -lean_object* x_56; -x_56 = lean_ctor_get(x_54, 1); -if (lean_obj_tag(x_56) == 0) -{ -lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_57 = lean_ctor_get(x_24, 1); -lean_inc(x_57); -lean_dec_ref(x_24); -x_58 = lean_ctor_get(x_25, 1); -lean_inc_ref(x_58); -lean_dec_ref(x_25); -lean_inc_ref(x_4); -x_59 = l_Lean_IR_EmitC_emitCName(x_2, x_4, x_57); -if (lean_obj_tag(x_59) == 0) -{ -lean_object* x_60; lean_object* x_61; lean_object* x_71; lean_object* x_72; uint8_t x_73; -x_60 = lean_ctor_get(x_59, 1); -lean_inc(x_60); -lean_dec_ref(x_59); -x_71 = lean_unsigned_to_nat(0u); -x_72 = lean_array_get_size(x_3); -x_73 = lean_nat_dec_lt(x_71, x_72); -if (x_73 == 0) +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitDec(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5) { +_start: { -lean_dec_ref(x_58); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -x_6 = x_60; -goto block_13; +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_6 = lean_ctor_get(x_4, 5); +x_7 = lean_box(0); +x_8 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0___redArg(x_7, x_6, x_1); +x_9 = l_Lean_IR_EmitC_argToCString___closed__0; +x_10 = l_Nat_reprFast(x_1); +x_11 = lean_string_append(x_9, x_10); +lean_dec_ref(x_10); +x_12 = l_Lean_IR_EmitC_emitDecOfType(x_11, x_8, x_2, x_3, x_4, x_5); +lean_dec(x_8); +lean_dec_ref(x_11); +return x_12; } -else -{ -lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; -x_74 = l_Array_zip___redArg(x_3, x_58); -lean_dec_ref(x_58); -lean_dec_ref(x_3); -x_75 = lean_array_get_size(x_74); -x_76 = l_Lean_IR_EmitC_emitFullApp___closed__0; -x_77 = lean_nat_dec_lt(x_71, x_75); -if (x_77 == 0) -{ -lean_dec_ref(x_74); -x_61 = x_76; -goto block_70; } -else -{ -uint8_t x_78; -x_78 = lean_nat_dec_le(x_75, x_75); -if (x_78 == 0) -{ -if (x_77 == 0) +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitDec___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -lean_dec_ref(x_74); -x_61 = x_76; -goto block_70; +uint8_t x_6; lean_object* x_7; +x_6 = lean_unbox(x_3); +x_7 = l_Lean_IR_EmitC_emitDec(x_1, x_2, x_6, x_4, x_5); +lean_dec_ref(x_4); +return x_7; } -else +} +static lean_object* _init_l_Lean_IR_EmitC_emitDel___redArg___closed__0() { +_start: { -size_t x_79; size_t x_80; lean_object* x_81; -x_79 = 0; -x_80 = lean_usize_of_nat(x_75); -x_81 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFullApp_spec__0(x_74, x_79, x_80, x_76); -lean_dec_ref(x_74); -x_61 = x_81; -goto block_70; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_free_object(", 17, 17); +return x_1; } } -else +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitDel___redArg(lean_object* x_1, lean_object* x_2) { +_start: { -size_t x_82; size_t x_83; lean_object* x_84; -x_82 = 0; -x_83 = lean_usize_of_nat(x_75); -x_84 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFullApp_spec__0(x_74, x_82, x_83, x_76); -lean_dec_ref(x_74); -x_61 = x_84; -goto block_70; -} +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_3 = l_Lean_IR_EmitC_emitDel___redArg___closed__0; +x_4 = lean_string_append(x_2, x_3); +x_5 = l_Lean_IR_EmitC_argToCString___closed__0; +x_6 = l_Nat_reprFast(x_1); +x_7 = lean_string_append(x_5, x_6); +lean_dec_ref(x_6); +x_8 = lean_string_append(x_4, x_7); +lean_dec_ref(x_7); +x_9 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_10 = lean_string_append(x_8, x_9); +x_11 = l_Lean_IR_EmitC_emitLn___closed__0; +x_12 = lean_box(0); +x_13 = lean_string_append(x_10, x_11); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +return x_14; } } -block_70: +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitDel(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_62 = l_Array_unzip___redArg(x_61); -lean_dec_ref(x_61); -x_63 = lean_ctor_get(x_62, 0); -lean_inc(x_63); -lean_dec_ref(x_62); -x_64 = l_Lean_IR_EmitC_emitFnDeclAux___closed__5; -x_65 = lean_string_append(x_60, x_64); -x_66 = l_Lean_IR_EmitC_emitArgs(x_63, x_4, x_65); -lean_dec_ref(x_4); -lean_dec(x_63); -x_67 = lean_ctor_get(x_66, 1); -lean_inc(x_67); -lean_dec_ref(x_66); -x_68 = l_Lean_IR_EmitC_emitFnDeclAux___closed__1; -x_69 = lean_string_append(x_67, x_68); -x_6 = x_69; -goto block_13; +lean_object* x_4; +x_4 = l_Lean_IR_EmitC_emitDel___redArg(x_1, x_3); +return x_4; } } -else +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitDel___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_dec_ref(x_58); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -return x_59; +lean_object* x_4; +x_4 = l_Lean_IR_EmitC_emitDel(x_1, x_2, x_3); +lean_dec_ref(x_2); +return x_4; } } -else +static lean_object* _init_l_Lean_IR_EmitC_emitSetTag___redArg___closed__0() { +_start: { -lean_object* x_85; lean_object* x_86; lean_object* x_87; -lean_inc_ref(x_54); -x_85 = lean_ctor_get(x_24, 1); -lean_inc(x_85); -lean_dec_ref(x_24); -x_86 = lean_ctor_get(x_25, 1); -lean_inc_ref(x_86); -lean_dec_ref(x_25); -x_87 = l_Lean_IR_EmitC_emitExternCall(x_2, x_86, x_54, x_3, x_4, x_85); -lean_dec_ref(x_4); -lean_dec_ref(x_86); -return x_87; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_ctor_set_tag(", 18, 18); +return x_1; } } -else +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSetTag___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_88; lean_object* x_89; lean_object* x_90; -lean_inc_ref(x_54); -x_88 = lean_ctor_get(x_24, 1); -lean_inc(x_88); -lean_dec_ref(x_24); -x_89 = lean_ctor_get(x_25, 1); -lean_inc_ref(x_89); -lean_dec_ref(x_25); -x_90 = l_Lean_IR_EmitC_emitExternCall(x_2, x_89, x_54, x_3, x_4, x_88); -lean_dec_ref(x_4); -lean_dec_ref(x_89); -return x_90; +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_4 = l_Lean_IR_EmitC_emitSetTag___redArg___closed__0; +x_5 = lean_string_append(x_3, x_4); +x_6 = l_Lean_IR_EmitC_argToCString___closed__0; +x_7 = l_Nat_reprFast(x_1); +x_8 = lean_string_append(x_6, x_7); +lean_dec_ref(x_7); +x_9 = lean_string_append(x_5, x_8); +lean_dec_ref(x_8); +x_10 = l_Lean_IR_EmitC_emitSpreadArg___closed__1; +x_11 = lean_string_append(x_9, x_10); +x_12 = l_Nat_reprFast(x_2); +x_13 = lean_string_append(x_11, x_12); +lean_dec_ref(x_12); +x_14 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_15 = lean_string_append(x_13, x_14); +x_16 = l_Lean_IR_EmitC_emitLn___closed__0; +x_17 = lean_box(0); +x_18 = lean_string_append(x_15, x_16); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +return x_19; } } -else +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSetTag(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_91; lean_object* x_92; lean_object* x_93; -lean_inc(x_54); -x_91 = lean_ctor_get(x_24, 1); -lean_inc(x_91); -lean_dec_ref(x_24); -x_92 = lean_ctor_get(x_25, 1); -lean_inc_ref(x_92); -lean_dec_ref(x_25); -x_93 = l_Lean_IR_EmitC_emitExternCall(x_2, x_92, x_54, x_3, x_4, x_91); -lean_dec_ref(x_4); -lean_dec_ref(x_92); -return x_93; -} +lean_object* x_5; +x_5 = l_Lean_IR_EmitC_emitSetTag___redArg(x_1, x_2, x_4); +return x_5; } } -else +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSetTag___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -uint8_t x_94; -lean_dec_ref(x_4); +lean_object* x_5; +x_5 = l_Lean_IR_EmitC_emitSetTag(x_1, x_2, x_3, x_4); lean_dec_ref(x_3); -lean_dec(x_2); -x_94 = !lean_is_exclusive(x_24); -if (x_94 == 0) -{ -return x_24; -} -else -{ -lean_object* x_95; lean_object* x_96; lean_object* x_97; -x_95 = lean_ctor_get(x_24, 0); -x_96 = lean_ctor_get(x_24, 1); -lean_inc(x_96); -lean_inc(x_95); -lean_dec(x_24); -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; +return x_5; } } -block_13: +static lean_object* _init_l_Lean_IR_EmitC_emitSet___redArg___closed__0() { +_start: { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_7 = l_Lean_IR_EmitC_emitFnDeclAux___closed__0; -x_8 = lean_string_append(x_6, x_7); -x_9 = l_Lean_IR_EmitC_emitLn___closed__0; -x_10 = lean_box(0); -x_11 = lean_string_append(x_8, x_9); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_10); -lean_ctor_set(x_12, 1, x_11); -return x_12; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_ctor_set(", 14, 14); +return x_1; } -block_21: +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSet___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_15 = l_Lean_IR_EmitC_emitFnDeclAux___closed__0; -x_16 = lean_string_append(x_14, x_15); -x_17 = l_Lean_IR_EmitC_emitLn___closed__0; -x_18 = lean_box(0); -x_19 = lean_string_append(x_16, x_17); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_18); -lean_ctor_set(x_20, 1, x_19); -return x_20; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_5 = l_Lean_IR_EmitC_emitSet___redArg___closed__0; +x_6 = lean_string_append(x_4, x_5); +x_7 = l_Lean_IR_EmitC_argToCString___closed__0; +x_8 = l_Nat_reprFast(x_1); +x_9 = lean_string_append(x_7, x_8); +lean_dec_ref(x_8); +x_10 = lean_string_append(x_6, x_9); +lean_dec_ref(x_9); +x_11 = l_Lean_IR_EmitC_emitSpreadArg___closed__1; +x_12 = lean_string_append(x_10, x_11); +x_13 = l_Nat_reprFast(x_2); +x_14 = lean_string_append(x_12, x_13); +lean_dec_ref(x_13); +x_15 = lean_string_append(x_14, x_11); +x_16 = l_Lean_IR_EmitC_emitArg___redArg(x_3, x_15); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_18 = lean_ctor_get(x_16, 1); +x_19 = lean_ctor_get(x_16, 0); +lean_dec(x_19); +x_20 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_21 = lean_string_append(x_18, x_20); +x_22 = l_Lean_IR_EmitC_emitLn___closed__0; +x_23 = lean_box(0); +x_24 = lean_string_append(x_21, x_22); +lean_ctor_set(x_16, 1, x_24); +lean_ctor_set(x_16, 0, x_23); +return x_16; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_25 = lean_ctor_get(x_16, 1); +lean_inc(x_25); +lean_dec(x_16); +x_26 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_27 = lean_string_append(x_25, x_26); +x_28 = l_Lean_IR_EmitC_emitLn___closed__0; +x_29 = lean_box(0); +x_30 = lean_string_append(x_27, x_28); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFullApp_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSet(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -size_t x_5; size_t x_6; lean_object* x_7; -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFullApp_spec__0(x_1, x_5, x_6, x_4); -lean_dec_ref(x_1); -return x_7; +lean_object* x_6; +x_6 = l_Lean_IR_EmitC_emitSet___redArg(x_1, x_2, x_3, x_5); +return x_6; } } -static lean_object* _init_l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitPartialApp_spec__0___redArg___closed__0() { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSet___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_IR_EmitC_emitSet(x_1, x_2, x_3, x_4, x_5); +lean_dec_ref(x_4); +return x_6; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitOffset___redArg___closed__0() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_closure_set(", 17, 17); +x_1 = lean_mk_string_unchecked("sizeof(void*)*", 14, 14); return x_1; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitPartialApp_spec__0___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +static lean_object* _init_l_Lean_IR_EmitC_emitOffset___redArg___closed__1() { _start: { -lean_object* x_8; uint8_t x_9; -x_8 = lean_unsigned_to_nat(0u); -x_9 = lean_nat_dec_eq(x_6, x_8); -if (x_9 == 1) +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" + ", 3, 3); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitOffset___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_10; lean_object* x_11; -lean_dec(x_6); +lean_object* x_4; uint8_t x_5; +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_nat_dec_lt(x_4, x_1); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +lean_dec(x_1); +x_6 = lean_box(0); +x_7 = l_Nat_reprFast(x_2); +x_8 = lean_string_append(x_3, x_7); +lean_dec_ref(x_7); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_6); +lean_ctor_set(x_9, 1, x_8); +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_10 = l_Lean_IR_EmitC_emitOffset___redArg___closed__0; +x_11 = lean_string_append(x_3, x_10); +x_12 = l_Nat_reprFast(x_1); +x_13 = lean_string_append(x_11, x_12); +lean_dec_ref(x_12); +x_14 = lean_nat_dec_lt(x_4, x_2); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_dec(x_2); -x_10 = lean_box(0); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_10); -lean_ctor_set(x_11, 1, x_7); -return x_11; +x_15 = lean_box(0); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_13); +return x_16; } else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_12 = lean_unsigned_to_nat(1u); -x_13 = lean_nat_sub(x_6, x_12); -lean_dec(x_6); -x_14 = lean_nat_sub(x_5, x_13); -x_15 = lean_nat_sub(x_14, x_12); -lean_dec(x_14); -x_16 = lean_array_fget_borrowed(x_1, x_15); -x_17 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitPartialApp_spec__0___redArg___closed__0; -x_18 = lean_string_append(x_7, x_17); -x_19 = l_Lean_IR_EmitC_argToCString___closed__0; -lean_inc(x_2); +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_17 = l_Lean_IR_EmitC_emitOffset___redArg___closed__1; +x_18 = lean_string_append(x_13, x_17); +x_19 = lean_box(0); x_20 = l_Nat_reprFast(x_2); -x_21 = lean_string_append(x_19, x_20); +x_21 = lean_string_append(x_18, x_20); lean_dec_ref(x_20); -x_22 = lean_string_append(x_18, x_21); -lean_dec_ref(x_21); -x_23 = lean_string_append(x_22, x_3); -x_24 = l_Nat_reprFast(x_15); -x_25 = lean_string_append(x_23, x_24); -lean_dec_ref(x_24); -x_26 = lean_string_append(x_25, x_3); -lean_inc(x_16); -x_27 = l_Lean_IR_EmitC_emitArg___redArg(x_16, x_26); -x_28 = lean_ctor_get(x_27, 1); -lean_inc(x_28); -lean_dec_ref(x_27); -x_29 = lean_string_append(x_28, x_4); -x_30 = l_Lean_IR_EmitC_emitLn___closed__0; -x_31 = lean_string_append(x_29, x_30); -x_6 = x_13; -x_7 = x_31; -goto _start; +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_19); +lean_ctor_set(x_22, 1, x_21); +return x_22; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitPartialApp_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitOffset(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_10; -x_10 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitPartialApp_spec__0___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_9); -return x_10; +lean_object* x_5; +x_5 = l_Lean_IR_EmitC_emitOffset___redArg(x_1, x_2, x_4); +return x_5; } } -static lean_object* _init_l_Lean_IR_EmitC_emitPartialApp___closed__0() { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitOffset___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_IR_EmitC_emitOffset(x_1, x_2, x_3, x_4); +lean_dec_ref(x_3); +return x_5; +} +} +static lean_object* _init_l_panic___at___00Lean_IR_EmitC_emitUSet_spec__0___closed__0() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__1; +x_2 = l_EStateM_instInhabited___redArg(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_EmitC_emitUSet_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = l_panic___at___00Lean_IR_EmitC_emitUSet_spec__0___closed__0; +x_5 = lean_alloc_closure((void*)(l_instInhabitedForall___redArg___lam__0___boxed), 2, 1); +lean_closure_set(x_5, 0, x_4); +x_6 = lean_panic_fn(x_5, x_1); +x_7 = lean_apply_2(x_6, x_2, x_3); +return x_7; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitUSet___closed__0() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_alloc_closure((void*)(", 27, 27); +x_1 = lean_mk_string_unchecked(".cs.c", 5, 5); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitPartialApp___closed__1() { +static lean_object* _init_l_Lean_IR_EmitC_emitUSet___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("), ", 3, 3); +x_1 = lean_mk_string_unchecked(".u[", 3, 3); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitPartialApp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +static lean_object* _init_l_Lean_IR_EmitC_emitUSet___closed__2() { _start: { -lean_object* x_6; -lean_inc_ref(x_4); -lean_inc(x_2); -x_6 = l_Lean_IR_EmitC_getDecl(x_2, x_4, x_5); -if (lean_obj_tag(x_6) == 0) -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_7 = lean_ctor_get(x_6, 0); -lean_inc(x_7); -x_8 = lean_ctor_get(x_6, 1); -lean_inc(x_8); -lean_dec_ref(x_6); -lean_inc(x_1); -x_9 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_8); -x_10 = lean_ctor_get(x_9, 1); -lean_inc(x_10); -lean_dec_ref(x_9); -x_11 = l_Lean_IR_EmitC_emitPartialApp___closed__0; -x_12 = lean_string_append(x_10, x_11); -x_13 = l_Lean_IR_EmitC_emitCName(x_2, x_4, x_12); -if (lean_obj_tag(x_13) == 0) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec_ref(x_13); -x_15 = l_Lean_IR_Decl_params(x_7); -lean_dec(x_7); -x_16 = lean_array_get_size(x_15); -lean_dec_ref(x_15); -x_17 = l_Lean_IR_EmitC_emitPartialApp___closed__1; -x_18 = lean_string_append(x_14, x_17); -x_19 = l_Nat_reprFast(x_16); -x_20 = lean_string_append(x_18, x_19); -lean_dec_ref(x_19); -x_21 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___closed__0; -x_22 = lean_string_append(x_20, x_21); -x_23 = lean_array_get_size(x_3); -x_24 = l_Nat_reprFast(x_23); -x_25 = lean_string_append(x_22, x_24); -lean_dec_ref(x_24); -x_26 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_27 = lean_string_append(x_25, x_26); -x_28 = l_Lean_IR_EmitC_emitLn___closed__0; -x_29 = lean_string_append(x_27, x_28); -x_30 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitPartialApp_spec__0___redArg(x_3, x_1, x_21, x_26, x_23, x_23, x_29); -return x_30; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("] = ", 4, 4); +return x_1; } -else +} +static lean_object* _init_l_Lean_IR_EmitC_emitUSet___closed__3() { +_start: { -lean_dec(x_7); -lean_dec(x_1); -return x_13; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Compiler.IR.EmitC", 22, 22); +return x_1; } } +static lean_object* _init_l_Lean_IR_EmitC_emitUSet___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.IR.EmitC.emitUSet", 22, 22); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitUSet___closed__5() { +_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_IR_EmitC_emitUSet___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_IR_EmitC_emitUSet___closed__5; +x_2 = lean_unsigned_to_nat(42u); +x_3 = lean_unsigned_to_nat(449u); +x_4 = l_Lean_IR_EmitC_emitUSet___closed__4; +x_5 = l_Lean_IR_EmitC_emitUSet___closed__3; +x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); +return x_6; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitUSet___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_ctor_set_usize(", 20, 20); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUSet(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_ctor_get(x_5, 5); +x_8 = lean_box(0); +x_9 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0___redArg(x_8, x_7, x_1); +if (lean_obj_tag(x_9) == 11) +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_9, 1); +x_12 = lean_ctor_get(x_9, 0); +lean_dec(x_12); +x_13 = lean_array_get(x_8, x_11, x_2); +lean_dec_ref(x_11); +if (lean_obj_tag(x_13) == 10) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +lean_dec_ref(x_5); +x_14 = lean_ctor_get(x_13, 1); +lean_inc_ref(x_14); +lean_dec_ref(x_13); +x_15 = l_Lean_IR_EmitC_argToCString___closed__0; +x_16 = l_Nat_reprFast(x_1); +x_17 = lean_string_append(x_15, x_16); +lean_dec_ref(x_16); +x_18 = lean_string_append(x_6, x_17); +lean_dec_ref(x_17); +x_19 = l_Lean_IR_EmitC_emitUSet___closed__0; +x_20 = lean_string_append(x_18, x_19); +x_21 = l_Nat_reprFast(x_2); +x_22 = lean_string_append(x_20, x_21); +lean_dec_ref(x_21); +x_23 = l_Lean_IR_EmitC_emitUSet___closed__1; +x_24 = lean_string_append(x_22, x_23); +x_25 = lean_array_get_size(x_14); +lean_dec_ref(x_14); +x_26 = lean_nat_sub(x_3, x_25); +lean_dec(x_3); +x_27 = l_Nat_reprFast(x_26); +x_28 = lean_string_append(x_24, x_27); +lean_dec_ref(x_27); +x_29 = l_Lean_IR_EmitC_emitUSet___closed__2; +x_30 = lean_string_append(x_28, x_29); +x_31 = l_Nat_reprFast(x_4); +x_32 = lean_string_append(x_15, x_31); +lean_dec_ref(x_31); +x_33 = lean_string_append(x_30, x_32); +lean_dec_ref(x_32); +x_34 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_35 = lean_string_append(x_33, x_34); +x_36 = l_Lean_IR_EmitC_emitLn___closed__0; +x_37 = lean_box(0); +x_38 = lean_string_append(x_35, x_36); +lean_ctor_set_tag(x_9, 0); +lean_ctor_set(x_9, 1, x_38); +lean_ctor_set(x_9, 0, x_37); +return x_9; +} else { -uint8_t x_31; -lean_dec_ref(x_4); +lean_object* x_39; lean_object* x_40; +lean_dec(x_13); +lean_free_object(x_9); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_31 = !lean_is_exclusive(x_6); -if (x_31 == 0) +x_39 = l_Lean_IR_EmitC_emitUSet___closed__6; +x_40 = l_panic___at___00Lean_IR_EmitC_emitUSet_spec__0(x_39, x_5, x_6); +return x_40; +} +} +else { -return x_6; +lean_object* x_41; lean_object* x_42; +x_41 = lean_ctor_get(x_9, 1); +lean_inc(x_41); +lean_dec(x_9); +x_42 = lean_array_get(x_8, x_41, x_2); +lean_dec_ref(x_41); +if (lean_obj_tag(x_42) == 10) +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +lean_dec_ref(x_5); +x_43 = lean_ctor_get(x_42, 1); +lean_inc_ref(x_43); +lean_dec_ref(x_42); +x_44 = l_Lean_IR_EmitC_argToCString___closed__0; +x_45 = l_Nat_reprFast(x_1); +x_46 = lean_string_append(x_44, x_45); +lean_dec_ref(x_45); +x_47 = lean_string_append(x_6, x_46); +lean_dec_ref(x_46); +x_48 = l_Lean_IR_EmitC_emitUSet___closed__0; +x_49 = lean_string_append(x_47, x_48); +x_50 = l_Nat_reprFast(x_2); +x_51 = lean_string_append(x_49, x_50); +lean_dec_ref(x_50); +x_52 = l_Lean_IR_EmitC_emitUSet___closed__1; +x_53 = lean_string_append(x_51, x_52); +x_54 = lean_array_get_size(x_43); +lean_dec_ref(x_43); +x_55 = lean_nat_sub(x_3, x_54); +lean_dec(x_3); +x_56 = l_Nat_reprFast(x_55); +x_57 = lean_string_append(x_53, x_56); +lean_dec_ref(x_56); +x_58 = l_Lean_IR_EmitC_emitUSet___closed__2; +x_59 = lean_string_append(x_57, x_58); +x_60 = l_Nat_reprFast(x_4); +x_61 = lean_string_append(x_44, x_60); +lean_dec_ref(x_60); +x_62 = lean_string_append(x_59, x_61); +lean_dec_ref(x_61); +x_63 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_64 = lean_string_append(x_62, x_63); +x_65 = l_Lean_IR_EmitC_emitLn___closed__0; +x_66 = lean_box(0); +x_67 = lean_string_append(x_64, x_65); +x_68 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_68, 0, x_66); +lean_ctor_set(x_68, 1, x_67); +return x_68; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_6, 0); -x_33 = lean_ctor_get(x_6, 1); -lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_6); -x_34 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_34, 0, x_32); -lean_ctor_set(x_34, 1, x_33); -return x_34; +lean_object* x_69; lean_object* x_70; +lean_dec(x_42); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_69 = l_Lean_IR_EmitC_emitUSet___closed__6; +x_70 = l_panic___at___00Lean_IR_EmitC_emitUSet_spec__0(x_69, x_5, x_6); +return x_70; +} +} +} +else +{ +lean_dec_ref(x_5); +lean_dec(x_2); +if (lean_obj_tag(x_9) == 10) +{ +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_71 = lean_ctor_get(x_9, 1); +lean_inc_ref(x_71); +lean_dec_ref(x_9); +x_72 = l_Lean_IR_EmitC_argToCString___closed__0; +x_73 = l_Nat_reprFast(x_1); +x_74 = lean_string_append(x_72, x_73); +lean_dec_ref(x_73); +x_75 = lean_string_append(x_6, x_74); +lean_dec_ref(x_74); +x_76 = l_Lean_IR_EmitC_emitUSet___closed__1; +x_77 = lean_string_append(x_75, x_76); +x_78 = lean_array_get_size(x_71); +lean_dec_ref(x_71); +x_79 = lean_nat_sub(x_3, x_78); +lean_dec(x_3); +x_80 = l_Nat_reprFast(x_79); +x_81 = lean_string_append(x_77, x_80); +lean_dec_ref(x_80); +x_82 = l_Lean_IR_EmitC_emitUSet___closed__2; +x_83 = lean_string_append(x_81, x_82); +x_84 = l_Nat_reprFast(x_4); +x_85 = lean_string_append(x_72, x_84); +lean_dec_ref(x_84); +x_86 = lean_string_append(x_83, x_85); +lean_dec_ref(x_85); +x_87 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_88 = lean_string_append(x_86, x_87); +x_89 = l_Lean_IR_EmitC_emitLn___closed__0; +x_90 = lean_box(0); +x_91 = lean_string_append(x_88, x_89); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_91); +return x_92; +} +else +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +lean_dec(x_9); +x_93 = l_Lean_IR_EmitC_emitUSet___closed__7; +x_94 = lean_string_append(x_6, x_93); +x_95 = l_Lean_IR_EmitC_argToCString___closed__0; +x_96 = l_Nat_reprFast(x_1); +x_97 = lean_string_append(x_95, x_96); +lean_dec_ref(x_96); +x_98 = lean_string_append(x_94, x_97); +lean_dec_ref(x_97); +x_99 = l_Lean_IR_EmitC_emitSpreadArg___closed__1; +x_100 = lean_string_append(x_98, x_99); +x_101 = l_Nat_reprFast(x_3); +x_102 = lean_string_append(x_100, x_101); +lean_dec_ref(x_101); +x_103 = lean_string_append(x_102, x_99); +x_104 = l_Nat_reprFast(x_4); +x_105 = lean_string_append(x_95, x_104); +lean_dec_ref(x_104); +x_106 = lean_string_append(x_103, x_105); +lean_dec_ref(x_105); +x_107 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_108 = lean_string_append(x_106, x_107); +x_109 = l_Lean_IR_EmitC_emitLn___closed__0; +x_110 = lean_box(0); +x_111 = lean_string_append(x_108, x_109); +x_112 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_112, 0, x_110); +lean_ctor_set(x_112, 1, x_111); +return x_112; } } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitPartialApp_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_IR_EmitC_emitSSet___closed__0() { _start: { -lean_object* x_10; -x_10 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitPartialApp_spec__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec_ref(x_8); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -lean_dec_ref(x_1); -return x_10; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("*((", 3, 3); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitPartialApp___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +static lean_object* _init_l_Lean_IR_EmitC_emitSSet___closed__1() { _start: { -lean_object* x_6; -x_6 = l_Lean_IR_EmitC_emitPartialApp(x_1, x_2, x_3, x_4, x_5); -lean_dec_ref(x_3); -return x_6; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("*)(", 3, 3); +return x_1; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitPartialApp_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +static lean_object* _init_l_Lean_IR_EmitC_emitSSet___closed__2() { _start: { -lean_object* x_8; -x_8 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitPartialApp_spec__0___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -lean_dec_ref(x_1); -return x_8; +lean_object* x_1; +x_1 = lean_mk_string_unchecked(".s+", 3, 3); +return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitApp___closed__0() { +static lean_object* _init_l_Lean_IR_EmitC_emitSSet___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_apply_", 11, 11); +x_1 = lean_mk_string_unchecked(")) = ", 5, 5); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitApp___closed__1() { +static lean_object* _init_l_Lean_IR_EmitC_emitSSet___closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("{ lean_object* _aargs[] = {", 27, 27); +x_1 = lean_mk_string_unchecked("lean_ctor_set_float", 19, 19); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitApp___closed__2() { +static lean_object* _init_l_Lean_IR_EmitC_emitSSet___closed__5() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("};", 2, 2); +x_1 = lean_mk_string_unchecked("lean_ctor_set_float32", 21, 21); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitApp___closed__3() { +static lean_object* _init_l_Lean_IR_EmitC_emitSSet___closed__6() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_apply_m(", 13, 13); +x_1 = lean_mk_string_unchecked("lean_ctor_set_uint8", 19, 19); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitApp___closed__4() { +static lean_object* _init_l_Lean_IR_EmitC_emitSSet___closed__7() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked(", _aargs); }", 12, 12); +x_1 = lean_mk_string_unchecked("lean_ctor_set_uint16", 20, 20); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitApp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +static lean_object* _init_l_Lean_IR_EmitC_emitSSet___closed__8() { _start: { -lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_6 = l_Lean_IR_EmitC_emitFnDeclAux___closed__3; -x_7 = lean_array_get_size(x_3); -x_8 = lean_nat_dec_lt(x_6, x_7); -if (x_8 == 0) -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_9 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_5); -x_10 = lean_ctor_get(x_9, 1); -lean_inc(x_10); -lean_dec_ref(x_9); -x_11 = l_Lean_IR_EmitC_emitApp___closed__0; -x_12 = lean_string_append(x_10, x_11); -x_13 = l_Nat_reprFast(x_7); -x_14 = lean_string_append(x_12, x_13); -lean_dec_ref(x_13); -x_15 = l_Lean_IR_EmitC_emitFnDeclAux___closed__5; -x_16 = lean_string_append(x_14, x_15); -x_17 = l_Lean_IR_EmitC_argToCString___closed__0; -x_18 = l_Nat_reprFast(x_2); -x_19 = lean_string_append(x_17, x_18); -lean_dec_ref(x_18); -x_20 = lean_string_append(x_16, x_19); -lean_dec_ref(x_19); -x_21 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___closed__0; -x_22 = lean_string_append(x_20, x_21); -x_23 = l_Lean_IR_EmitC_emitArgs(x_3, x_4, x_22); -x_24 = !lean_is_exclusive(x_23); -if (x_24 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_ctor_set_uint32", 20, 20); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitSSet___closed__9() { +_start: { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_25 = lean_ctor_get(x_23, 1); -x_26 = lean_ctor_get(x_23, 0); -lean_dec(x_26); -x_27 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_28 = lean_string_append(x_25, x_27); -x_29 = l_Lean_IR_EmitC_emitLn___closed__0; -x_30 = lean_box(0); -x_31 = lean_string_append(x_28, x_29); -lean_ctor_set(x_23, 1, x_31); -lean_ctor_set(x_23, 0, x_30); -return x_23; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_ctor_set_uint64", 20, 20); +return x_1; } -else +} +static lean_object* _init_l_Lean_IR_EmitC_emitSSet___closed__10() { +_start: { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_32 = lean_ctor_get(x_23, 1); -lean_inc(x_32); -lean_dec(x_23); -x_33 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_34 = lean_string_append(x_32, x_33); -x_35 = l_Lean_IR_EmitC_emitLn___closed__0; -x_36 = lean_box(0); -x_37 = lean_string_append(x_34, x_35); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -return x_38; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid instruction", 19, 19); +return x_1; } } -else +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSSet(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; -x_39 = l_Lean_IR_EmitC_emitApp___closed__1; -x_40 = lean_string_append(x_5, x_39); -x_41 = l_Lean_IR_EmitC_emitArgs(x_3, x_4, x_40); -x_42 = lean_ctor_get(x_41, 1); -lean_inc(x_42); -lean_dec_ref(x_41); -x_43 = l_Lean_IR_EmitC_emitApp___closed__2; -x_44 = lean_string_append(x_42, x_43); -x_45 = l_Lean_IR_EmitC_emitLn___closed__0; -x_46 = lean_string_append(x_44, x_45); -x_47 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_46); -x_48 = !lean_is_exclusive(x_47); -if (x_48 == 0) +lean_object* x_9; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_7, 5); +x_44 = lean_box(0); +x_45 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0___redArg(x_44, x_43, x_1); +if (lean_obj_tag(x_45) == 11) { -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_49 = lean_ctor_get(x_47, 1); -x_50 = lean_ctor_get(x_47, 0); +lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; +lean_dec_ref(x_45); +lean_dec(x_3); +x_46 = l_Lean_IR_EmitC_emitSSet___closed__0; +x_47 = lean_string_append(x_8, x_46); +x_48 = l_Lean_IR_EmitC_toCType(x_6, x_7, x_47); +x_49 = !lean_is_exclusive(x_48); +if (x_49 == 0) +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_50 = lean_ctor_get(x_48, 0); +x_51 = lean_ctor_get(x_48, 1); +x_52 = lean_string_append(x_51, x_50); lean_dec(x_50); -x_51 = l_Lean_IR_EmitC_emitApp___closed__3; -x_52 = lean_string_append(x_49, x_51); -x_53 = l_Lean_IR_EmitC_argToCString___closed__0; -x_54 = l_Nat_reprFast(x_2); -x_55 = lean_string_append(x_53, x_54); -lean_dec_ref(x_54); -x_56 = lean_string_append(x_52, x_55); -lean_dec_ref(x_55); -x_57 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___closed__0; -x_58 = lean_string_append(x_56, x_57); -x_59 = l_Nat_reprFast(x_7); +x_53 = l_Lean_IR_EmitC_emitSSet___closed__1; +x_54 = lean_string_append(x_52, x_53); +x_55 = l_Lean_IR_EmitC_argToCString___closed__0; +x_56 = l_Nat_reprFast(x_1); +x_57 = lean_string_append(x_55, x_56); +lean_dec_ref(x_56); +x_58 = lean_string_append(x_54, x_57); +lean_dec_ref(x_57); +x_59 = l_Lean_IR_EmitC_emitUSet___closed__0; x_60 = lean_string_append(x_58, x_59); -lean_dec_ref(x_59); -x_61 = l_Lean_IR_EmitC_emitApp___closed__4; +x_61 = l_Nat_reprFast(x_2); x_62 = lean_string_append(x_60, x_61); -x_63 = lean_box(0); -x_64 = lean_string_append(x_62, x_45); -lean_ctor_set(x_47, 1, x_64); -lean_ctor_set(x_47, 0, x_63); -return x_47; -} -else -{ -lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_65 = lean_ctor_get(x_47, 1); -lean_inc(x_65); -lean_dec(x_47); -x_66 = l_Lean_IR_EmitC_emitApp___closed__3; -x_67 = lean_string_append(x_65, x_66); -x_68 = l_Lean_IR_EmitC_argToCString___closed__0; -x_69 = l_Nat_reprFast(x_2); -x_70 = lean_string_append(x_68, x_69); +lean_dec_ref(x_61); +x_63 = l_Lean_IR_EmitC_emitSSet___closed__2; +x_64 = lean_string_append(x_62, x_63); +x_65 = l_Nat_reprFast(x_4); +x_66 = lean_string_append(x_64, x_65); +lean_dec_ref(x_65); +x_67 = l_Lean_IR_EmitC_emitSSet___closed__3; +x_68 = lean_string_append(x_66, x_67); +x_69 = l_Nat_reprFast(x_5); +x_70 = lean_string_append(x_55, x_69); lean_dec_ref(x_69); -x_71 = lean_string_append(x_67, x_70); +x_71 = lean_string_append(x_68, x_70); lean_dec_ref(x_70); -x_72 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___closed__0; +x_72 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; x_73 = lean_string_append(x_71, x_72); -x_74 = l_Nat_reprFast(x_7); +x_74 = l_Lean_IR_EmitC_emitLn___closed__0; x_75 = lean_string_append(x_73, x_74); -lean_dec_ref(x_74); -x_76 = l_Lean_IR_EmitC_emitApp___closed__4; -x_77 = lean_string_append(x_75, x_76); -x_78 = lean_box(0); -x_79 = lean_string_append(x_77, x_45); -x_80 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_80, 0, x_78); -lean_ctor_set(x_80, 1, x_79); -return x_80; -} -} +x_76 = lean_box(0); +lean_ctor_set(x_48, 1, x_75); +lean_ctor_set(x_48, 0, x_76); +return x_48; } -} -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitApp___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +else { -lean_object* x_6; -x_6 = l_Lean_IR_EmitC_emitApp(x_1, x_2, x_3, x_4, x_5); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -return x_6; +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_77 = lean_ctor_get(x_48, 0); +x_78 = lean_ctor_get(x_48, 1); +lean_inc(x_78); +lean_inc(x_77); +lean_dec(x_48); +x_79 = lean_string_append(x_78, x_77); +lean_dec(x_77); +x_80 = l_Lean_IR_EmitC_emitSSet___closed__1; +x_81 = lean_string_append(x_79, x_80); +x_82 = l_Lean_IR_EmitC_argToCString___closed__0; +x_83 = l_Nat_reprFast(x_1); +x_84 = lean_string_append(x_82, x_83); +lean_dec_ref(x_83); +x_85 = lean_string_append(x_81, x_84); +lean_dec_ref(x_84); +x_86 = l_Lean_IR_EmitC_emitUSet___closed__0; +x_87 = lean_string_append(x_85, x_86); +x_88 = l_Nat_reprFast(x_2); +x_89 = lean_string_append(x_87, x_88); +lean_dec_ref(x_88); +x_90 = l_Lean_IR_EmitC_emitSSet___closed__2; +x_91 = lean_string_append(x_89, x_90); +x_92 = l_Nat_reprFast(x_4); +x_93 = lean_string_append(x_91, x_92); +lean_dec_ref(x_92); +x_94 = l_Lean_IR_EmitC_emitSSet___closed__3; +x_95 = lean_string_append(x_93, x_94); +x_96 = l_Nat_reprFast(x_5); +x_97 = lean_string_append(x_82, x_96); +lean_dec_ref(x_96); +x_98 = lean_string_append(x_95, x_97); +lean_dec_ref(x_97); +x_99 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_100 = lean_string_append(x_98, x_99); +x_101 = l_Lean_IR_EmitC_emitLn___closed__0; +x_102 = lean_string_append(x_100, x_101); +x_103 = lean_box(0); +x_104 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_104, 0, x_103); +lean_ctor_set(x_104, 1, x_102); +return x_104; } } -static lean_object* _init_l_Lean_IR_EmitC_emitBoxFn___redArg___closed__0() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_box_usize", 14, 14); -return x_1; -} -} -static lean_object* _init_l_Lean_IR_EmitC_emitBoxFn___redArg___closed__1() { -_start: +lean_dec(x_2); +if (lean_obj_tag(x_45) == 10) { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_box_uint32", 15, 15); -return x_1; -} -} -static lean_object* _init_l_Lean_IR_EmitC_emitBoxFn___redArg___closed__2() { -_start: +lean_object* x_105; lean_object* x_106; lean_object* x_107; uint8_t x_108; +lean_dec_ref(x_45); +lean_dec(x_3); +x_105 = l_Lean_IR_EmitC_emitSSet___closed__0; +x_106 = lean_string_append(x_8, x_105); +x_107 = l_Lean_IR_EmitC_toCType(x_6, x_7, x_106); +x_108 = !lean_is_exclusive(x_107); +if (x_108 == 0) { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_box_uint64", 15, 15); -return x_1; +lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; +x_109 = lean_ctor_get(x_107, 0); +x_110 = lean_ctor_get(x_107, 1); +x_111 = lean_string_append(x_110, x_109); +lean_dec(x_109); +x_112 = l_Lean_IR_EmitC_emitSSet___closed__1; +x_113 = lean_string_append(x_111, x_112); +x_114 = l_Lean_IR_EmitC_argToCString___closed__0; +x_115 = l_Nat_reprFast(x_1); +x_116 = lean_string_append(x_114, x_115); +lean_dec_ref(x_115); +x_117 = lean_string_append(x_113, x_116); +lean_dec_ref(x_116); +x_118 = l_Lean_IR_EmitC_emitSSet___closed__2; +x_119 = lean_string_append(x_117, x_118); +x_120 = l_Nat_reprFast(x_4); +x_121 = lean_string_append(x_119, x_120); +lean_dec_ref(x_120); +x_122 = l_Lean_IR_EmitC_emitSSet___closed__3; +x_123 = lean_string_append(x_121, x_122); +x_124 = l_Nat_reprFast(x_5); +x_125 = lean_string_append(x_114, x_124); +lean_dec_ref(x_124); +x_126 = lean_string_append(x_123, x_125); +lean_dec_ref(x_125); +x_127 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_128 = lean_string_append(x_126, x_127); +x_129 = l_Lean_IR_EmitC_emitLn___closed__0; +x_130 = lean_string_append(x_128, x_129); +x_131 = lean_box(0); +lean_ctor_set(x_107, 1, x_130); +lean_ctor_set(x_107, 0, x_131); +return x_107; +} +else +{ +lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; +x_132 = lean_ctor_get(x_107, 0); +x_133 = lean_ctor_get(x_107, 1); +lean_inc(x_133); +lean_inc(x_132); +lean_dec(x_107); +x_134 = lean_string_append(x_133, x_132); +lean_dec(x_132); +x_135 = l_Lean_IR_EmitC_emitSSet___closed__1; +x_136 = lean_string_append(x_134, x_135); +x_137 = l_Lean_IR_EmitC_argToCString___closed__0; +x_138 = l_Nat_reprFast(x_1); +x_139 = lean_string_append(x_137, x_138); +lean_dec_ref(x_138); +x_140 = lean_string_append(x_136, x_139); +lean_dec_ref(x_139); +x_141 = l_Lean_IR_EmitC_emitSSet___closed__2; +x_142 = lean_string_append(x_140, x_141); +x_143 = l_Nat_reprFast(x_4); +x_144 = lean_string_append(x_142, x_143); +lean_dec_ref(x_143); +x_145 = l_Lean_IR_EmitC_emitSSet___closed__3; +x_146 = lean_string_append(x_144, x_145); +x_147 = l_Nat_reprFast(x_5); +x_148 = lean_string_append(x_137, x_147); +lean_dec_ref(x_147); +x_149 = lean_string_append(x_146, x_148); +lean_dec_ref(x_148); +x_150 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_151 = lean_string_append(x_149, x_150); +x_152 = l_Lean_IR_EmitC_emitLn___closed__0; +x_153 = lean_string_append(x_151, x_152); +x_154 = lean_box(0); +x_155 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_155, 0, x_154); +lean_ctor_set(x_155, 1, x_153); +return x_155; } } -static lean_object* _init_l_Lean_IR_EmitC_emitBoxFn___redArg___closed__3() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_box_float", 14, 14); -return x_1; -} -} -static lean_object* _init_l_Lean_IR_EmitC_emitBoxFn___redArg___closed__4() { -_start: +lean_dec(x_45); +switch (lean_obj_tag(x_6)) { +case 0: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_box_float32", 16, 16); -return x_1; -} +lean_object* x_156; lean_object* x_157; +x_156 = l_Lean_IR_EmitC_emitSSet___closed__4; +x_157 = lean_string_append(x_8, x_156); +x_9 = x_157; +goto block_42; } -static lean_object* _init_l_Lean_IR_EmitC_emitBoxFn___redArg___closed__5() { -_start: +case 9: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_box", 8, 8); -return x_1; -} +lean_object* x_158; lean_object* x_159; +x_158 = l_Lean_IR_EmitC_emitSSet___closed__5; +x_159 = lean_string_append(x_8, x_158); +x_9 = x_159; +goto block_42; } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitBoxFn___redArg(lean_object* x_1, lean_object* x_2) { -_start: +case 1: { -switch (lean_obj_tag(x_1)) { -case 5: +lean_object* x_160; lean_object* x_161; +x_160 = l_Lean_IR_EmitC_emitSSet___closed__6; +x_161 = lean_string_append(x_8, x_160); +x_9 = x_161; +goto block_42; +} +case 2: { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_3 = l_Lean_IR_EmitC_emitBoxFn___redArg___closed__0; -x_4 = lean_box(0); -x_5 = lean_string_append(x_2, x_3); -x_6 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_6, 0, x_4); -lean_ctor_set(x_6, 1, x_5); -return x_6; +lean_object* x_162; lean_object* x_163; +x_162 = l_Lean_IR_EmitC_emitSSet___closed__7; +x_163 = lean_string_append(x_8, x_162); +x_9 = x_163; +goto block_42; } case 3: { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_7 = l_Lean_IR_EmitC_emitBoxFn___redArg___closed__1; -x_8 = lean_box(0); -x_9 = lean_string_append(x_2, x_7); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_8); -lean_ctor_set(x_10, 1, x_9); -return x_10; +lean_object* x_164; lean_object* x_165; +x_164 = l_Lean_IR_EmitC_emitSSet___closed__8; +x_165 = lean_string_append(x_8, x_164); +x_9 = x_165; +goto block_42; } case 4: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_11 = l_Lean_IR_EmitC_emitBoxFn___redArg___closed__2; -x_12 = lean_box(0); -x_13 = lean_string_append(x_2, x_11); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_12); -lean_ctor_set(x_14, 1, x_13); -return x_14; -} -case 0: -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_15 = l_Lean_IR_EmitC_emitBoxFn___redArg___closed__3; -x_16 = lean_box(0); -x_17 = lean_string_append(x_2, x_15); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_16); -lean_ctor_set(x_18, 1, x_17); -return x_18; -} -case 9: -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_19 = l_Lean_IR_EmitC_emitBoxFn___redArg___closed__4; -x_20 = lean_box(0); -x_21 = lean_string_append(x_2, x_19); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -return x_22; +lean_object* x_166; lean_object* x_167; +x_166 = l_Lean_IR_EmitC_emitSSet___closed__9; +x_167 = lean_string_append(x_8, x_166); +x_9 = x_167; +goto block_42; } default: { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_23 = l_Lean_IR_EmitC_emitBoxFn___redArg___closed__5; -x_24 = lean_box(0); -x_25 = lean_string_append(x_2, x_23); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -return x_26; +lean_object* x_168; lean_object* x_169; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_168 = l_Lean_IR_EmitC_emitSSet___closed__10; +x_169 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_169, 0, x_168); +lean_ctor_set(x_169, 1, x_8); +return x_169; } } } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitBoxFn(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +block_42: { -lean_object* x_4; -x_4 = l_Lean_IR_EmitC_emitBoxFn___redArg(x_1, x_3); -return x_4; -} +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_10 = l_Lean_IR_EmitC_emitFnDeclAux___closed__2; +x_11 = lean_string_append(x_9, x_10); +x_12 = l_Lean_IR_EmitC_argToCString___closed__0; +x_13 = l_Nat_reprFast(x_1); +x_14 = lean_string_append(x_12, x_13); +lean_dec_ref(x_13); +x_15 = lean_string_append(x_11, x_14); +lean_dec_ref(x_14); +x_16 = l_Lean_IR_EmitC_emitSpreadArg___closed__1; +x_17 = lean_string_append(x_15, x_16); +x_18 = l_Lean_IR_EmitC_emitOffset___redArg(x_3, x_4, x_17); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_20 = lean_ctor_get(x_18, 1); +x_21 = lean_ctor_get(x_18, 0); +lean_dec(x_21); +x_22 = lean_string_append(x_20, x_16); +x_23 = l_Nat_reprFast(x_5); +x_24 = lean_string_append(x_12, x_23); +lean_dec_ref(x_23); +x_25 = lean_string_append(x_22, x_24); +lean_dec_ref(x_24); +x_26 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_27 = lean_string_append(x_25, x_26); +x_28 = l_Lean_IR_EmitC_emitLn___closed__0; +x_29 = lean_box(0); +x_30 = lean_string_append(x_27, x_28); +lean_ctor_set(x_18, 1, x_30); +lean_ctor_set(x_18, 0, x_29); +return x_18; } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitBoxFn___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -lean_object* x_4; -x_4 = l_Lean_IR_EmitC_emitBoxFn(x_1, x_2, x_3); -lean_dec_ref(x_2); -lean_dec(x_1); -return x_4; +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_31 = lean_ctor_get(x_18, 1); +lean_inc(x_31); +lean_dec(x_18); +x_32 = lean_string_append(x_31, x_16); +x_33 = l_Nat_reprFast(x_5); +x_34 = lean_string_append(x_12, x_33); +lean_dec_ref(x_33); +x_35 = lean_string_append(x_32, x_34); +lean_dec_ref(x_34); +x_36 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_37 = lean_string_append(x_35, x_36); +x_38 = l_Lean_IR_EmitC_emitLn___closed__0; +x_39 = lean_box(0); +x_40 = lean_string_append(x_37, x_38); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +return x_41; +} } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitBoxFn___redArg___boxed(lean_object* x_1, lean_object* x_2) { +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSSet___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_3; -x_3 = l_Lean_IR_EmitC_emitBoxFn___redArg(x_1, x_2); -lean_dec(x_1); -return x_3; +lean_object* x_9; +x_9 = l_Lean_IR_EmitC_emitSSet(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec_ref(x_7); +lean_dec(x_6); +return x_9; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitBox___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_5 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_4); -x_6 = lean_ctor_get(x_5, 1); -lean_inc(x_6); -lean_dec_ref(x_5); -x_7 = l_Lean_IR_EmitC_emitBoxFn___redArg(x_3, x_6); -x_8 = !lean_is_exclusive(x_7); -if (x_8 == 0) +lean_object* x_6; uint8_t x_7; +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_nat_dec_eq(x_4, x_6); +if (x_7 == 1) { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_9 = lean_ctor_get(x_7, 1); -x_10 = lean_ctor_get(x_7, 0); -lean_dec(x_10); -x_11 = l_Lean_IR_EmitC_emitFnDeclAux___closed__5; -x_12 = lean_string_append(x_9, x_11); -x_13 = l_Lean_IR_EmitC_argToCString___closed__0; -x_14 = l_Nat_reprFast(x_2); -x_15 = lean_string_append(x_13, x_14); -lean_dec_ref(x_14); -x_16 = lean_string_append(x_12, x_15); -lean_dec_ref(x_15); -x_17 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_18 = lean_string_append(x_16, x_17); -x_19 = l_Lean_IR_EmitC_emitLn___closed__0; -x_20 = lean_box(0); -x_21 = lean_string_append(x_18, x_19); -lean_ctor_set(x_7, 1, x_21); -lean_ctor_set(x_7, 0, x_20); -return x_7; -} +lean_object* x_8; lean_object* x_9; +lean_dec(x_4); +x_8 = lean_box(0); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_5); +return x_9; +} else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_22 = lean_ctor_get(x_7, 1); -lean_inc(x_22); +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; 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_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_sub(x_4, x_10); +lean_dec(x_4); +x_12 = lean_nat_sub(x_3, x_11); +x_13 = lean_nat_sub(x_12, x_10); +lean_dec(x_12); +x_14 = lean_array_fget_borrowed(x_1, x_13); +x_15 = lean_ctor_get(x_14, 0); +x_16 = lean_array_fget_borrowed(x_2, x_13); +lean_dec(x_13); +x_17 = l_Lean_IR_EmitC_argToCString___closed__0; +lean_inc(x_15); +x_18 = l_Nat_reprFast(x_15); +x_19 = lean_string_append(x_17, x_18); +lean_dec_ref(x_18); +x_20 = lean_string_append(x_5, x_19); +lean_dec_ref(x_19); +x_21 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__1; +x_22 = lean_string_append(x_20, x_21); +lean_inc(x_16); +x_23 = l_Lean_IR_EmitC_emitArg___redArg(x_16, x_22); +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec_ref(x_23); +x_25 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_26 = lean_string_append(x_24, x_25); +x_27 = l_Lean_IR_EmitC_emitLn___closed__0; +x_28 = lean_string_append(x_26, x_27); +x_4 = x_11; +x_5 = x_28; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0___redArg(x_1, x_2, x_3, x_4, x_7); +return x_8; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitJmp___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid goto", 12, 12); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitJmp___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("goto ", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitJmp___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("block_", 6, 6); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitJmp(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_IR_EmitC_getJPParams(x_1, x_3, x_4); +if (lean_obj_tag(x_5) == 0) +{ +uint8_t x_6; +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_ctor_get(x_5, 1); +x_9 = lean_array_get_size(x_2); +x_10 = lean_array_get_size(x_7); +x_11 = lean_nat_dec_eq(x_9, x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_dec(x_7); -x_23 = l_Lean_IR_EmitC_emitFnDeclAux___closed__5; +lean_dec(x_1); +x_12 = l_Lean_IR_EmitC_emitJmp___closed__0; +lean_ctor_set_tag(x_5, 1); +lean_ctor_set(x_5, 0, x_12); +return x_5; +} +else +{ +lean_object* x_13; uint8_t x_14; +lean_free_object(x_5); +x_13 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0___redArg(x_7, x_2, x_9, x_9, x_8); +lean_dec(x_7); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_15 = lean_ctor_get(x_13, 1); +x_16 = lean_ctor_get(x_13, 0); +lean_dec(x_16); +x_17 = l_Lean_IR_EmitC_emitJmp___closed__1; +x_18 = lean_string_append(x_15, x_17); +x_19 = l_Lean_IR_EmitC_emitJmp___closed__2; +x_20 = l_Nat_reprFast(x_1); +x_21 = lean_string_append(x_19, x_20); +lean_dec_ref(x_20); +x_22 = lean_string_append(x_18, x_21); +lean_dec_ref(x_21); +x_23 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; x_24 = lean_string_append(x_22, x_23); -x_25 = l_Lean_IR_EmitC_argToCString___closed__0; -x_26 = l_Nat_reprFast(x_2); -x_27 = lean_string_append(x_25, x_26); -lean_dec_ref(x_26); -x_28 = lean_string_append(x_24, x_27); -lean_dec_ref(x_27); -x_29 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; +x_25 = l_Lean_IR_EmitC_emitLn___closed__0; +x_26 = lean_box(0); +x_27 = lean_string_append(x_24, x_25); +lean_ctor_set(x_13, 1, x_27); +lean_ctor_set(x_13, 0, x_26); +return x_13; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_28 = lean_ctor_get(x_13, 1); +lean_inc(x_28); +lean_dec(x_13); +x_29 = l_Lean_IR_EmitC_emitJmp___closed__1; x_30 = lean_string_append(x_28, x_29); -x_31 = l_Lean_IR_EmitC_emitLn___closed__0; -x_32 = lean_box(0); -x_33 = lean_string_append(x_30, x_31); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_32); -lean_ctor_set(x_34, 1, x_33); -return x_34; +x_31 = l_Lean_IR_EmitC_emitJmp___closed__2; +x_32 = l_Nat_reprFast(x_1); +x_33 = lean_string_append(x_31, x_32); +lean_dec_ref(x_32); +x_34 = lean_string_append(x_30, x_33); +lean_dec_ref(x_33); +x_35 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_36 = lean_string_append(x_34, x_35); +x_37 = l_Lean_IR_EmitC_emitLn___closed__0; +x_38 = lean_box(0); +x_39 = lean_string_append(x_36, x_37); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set(x_40, 1, x_39); +return x_40; +} +} +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; +x_41 = lean_ctor_get(x_5, 0); +x_42 = lean_ctor_get(x_5, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_5); +x_43 = lean_array_get_size(x_2); +x_44 = lean_array_get_size(x_41); +x_45 = lean_nat_dec_eq(x_43, x_44); +if (x_45 == 0) +{ +lean_object* x_46; lean_object* x_47; +lean_dec(x_41); +lean_dec(x_1); +x_46 = l_Lean_IR_EmitC_emitJmp___closed__0; +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_42); +return x_47; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_48 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0___redArg(x_41, x_2, x_43, x_43, x_42); +lean_dec(x_41); +x_49 = lean_ctor_get(x_48, 1); +lean_inc(x_49); +if (lean_is_exclusive(x_48)) { + lean_ctor_release(x_48, 0); + lean_ctor_release(x_48, 1); + x_50 = x_48; +} else { + lean_dec_ref(x_48); + x_50 = lean_box(0); +} +x_51 = l_Lean_IR_EmitC_emitJmp___closed__1; +x_52 = lean_string_append(x_49, x_51); +x_53 = l_Lean_IR_EmitC_emitJmp___closed__2; +x_54 = l_Nat_reprFast(x_1); +x_55 = lean_string_append(x_53, x_54); +lean_dec_ref(x_54); +x_56 = lean_string_append(x_52, x_55); +lean_dec_ref(x_55); +x_57 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_58 = lean_string_append(x_56, x_57); +x_59 = l_Lean_IR_EmitC_emitLn___closed__0; +x_60 = lean_box(0); +x_61 = lean_string_append(x_58, x_59); +if (lean_is_scalar(x_50)) { + x_62 = lean_alloc_ctor(0, 2, 0); +} else { + x_62 = x_50; +} +lean_ctor_set(x_62, 0, x_60); +lean_ctor_set(x_62, 1, x_61); +return x_62; +} +} +} +else +{ +uint8_t x_63; +lean_dec(x_1); +x_63 = !lean_is_exclusive(x_5); +if (x_63 == 0) +{ +return x_5; +} +else +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_5, 0); +x_65 = lean_ctor_get(x_5, 1); +lean_inc(x_65); +lean_inc(x_64); +lean_dec(x_5); +x_66 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +return x_66; +} } } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitBox(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_6; -x_6 = l_Lean_IR_EmitC_emitBox___redArg(x_1, x_2, x_3, x_5); -return x_6; +lean_object* x_8; +x_8 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec_ref(x_6); +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitBox___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitJmp___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_IR_EmitC_emitJmp(x_1, x_2, x_3, x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0___redArg(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLhs___redArg(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_Lean_IR_EmitC_argToCString___closed__0; +x_4 = l_Nat_reprFast(x_1); +x_5 = lean_string_append(x_3, x_4); +lean_dec_ref(x_4); +x_6 = lean_string_append(x_2, x_5); +lean_dec_ref(x_5); +x_7 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__1; +x_8 = lean_box(0); +x_9 = lean_string_append(x_6, x_7); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_8); +lean_ctor_set(x_10, 1, x_9); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLhs(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLhs___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_IR_EmitC_emitLhs(x_1, x_2, x_3); +lean_dec_ref(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitArgs_spec__0___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; uint8_t x_6; +x_5 = lean_unsigned_to_nat(0u); +x_6 = lean_nat_dec_eq(x_3, x_5); +if (x_6 == 1) +{ +lean_object* x_7; lean_object* x_8; +lean_dec(x_3); +x_7 = lean_box(0); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_4); +return x_8; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_19; +x_9 = lean_unsigned_to_nat(1u); +x_10 = lean_nat_sub(x_3, x_9); +lean_dec(x_3); +x_11 = lean_nat_sub(x_2, x_10); +x_12 = lean_nat_sub(x_11, x_9); +lean_dec(x_11); +x_19 = lean_nat_dec_lt(x_5, x_12); +if (x_19 == 0) +{ +x_13 = x_4; +goto block_18; +} +else +{ +lean_object* x_20; lean_object* x_21; +x_20 = l_Lean_IR_EmitC_emitSpreadArg___closed__1; +x_21 = lean_string_append(x_4, x_20); +x_13 = x_21; +goto block_18; +} +block_18: +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_array_fget_borrowed(x_1, x_12); +lean_dec(x_12); +lean_inc(x_14); +x_15 = l_Lean_IR_EmitC_emitArg___redArg(x_14, x_13); +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec_ref(x_15); +x_3 = x_10; +x_4 = x_16; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitArgs_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitArgs_spec__0___redArg(x_1, x_2, x_3, x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitArgs(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_array_get_size(x_1); +x_5 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitArgs_spec__0___redArg(x_1, x_4, x_4, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitArgs_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitArgs_spec__0(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec_ref(x_5); +lean_dec(x_2); +lean_dec_ref(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitArgs___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_IR_EmitC_emitArgs(x_1, x_2, x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitArgs_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitArgs_spec__0___redArg(x_1, x_2, x_3, x_4); +lean_dec(x_2); +lean_dec_ref(x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitCtorScalarSize___redArg___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("sizeof(size_t)*", 15, 15); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCtorScalarSize___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_nat_dec_eq(x_1, x_4); +if (x_5 == 0) +{ +uint8_t x_6; +x_6 = lean_nat_dec_eq(x_2, x_4); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_7 = l_Lean_IR_EmitC_emitCtorScalarSize___redArg___closed__0; +x_8 = lean_string_append(x_3, x_7); +x_9 = l_Nat_reprFast(x_1); +x_10 = lean_string_append(x_8, x_9); +lean_dec_ref(x_9); +x_11 = l_Lean_IR_EmitC_emitOffset___redArg___closed__1; +x_12 = lean_string_append(x_10, x_11); +x_13 = lean_box(0); +x_14 = l_Nat_reprFast(x_2); +x_15 = lean_string_append(x_12, x_14); +lean_dec_ref(x_14); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_13); +lean_ctor_set(x_16, 1, x_15); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_dec(x_2); +x_17 = l_Lean_IR_EmitC_emitCtorScalarSize___redArg___closed__0; +x_18 = lean_string_append(x_3, x_17); +x_19 = lean_box(0); +x_20 = l_Nat_reprFast(x_1); +x_21 = lean_string_append(x_18, x_20); +lean_dec_ref(x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_19); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_dec(x_1); +x_23 = lean_box(0); +x_24 = l_Nat_reprFast(x_2); +x_25 = lean_string_append(x_3, x_24); +lean_dec_ref(x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_23); +lean_ctor_set(x_26, 1, x_25); +return x_26; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCtorScalarSize(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_IR_EmitC_emitCtorScalarSize___redArg(x_1, x_2, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCtorScalarSize___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_IR_EmitC_emitCtorScalarSize(x_1, x_2, x_3, x_4); +lean_dec_ref(x_3); +return x_5; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitAllocCtor___redArg___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_alloc_ctor(", 16, 16); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitAllocCtor___redArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +x_5 = lean_ctor_get(x_1, 3); +lean_inc(x_5); +x_6 = lean_ctor_get(x_1, 4); +lean_inc(x_6); +lean_dec_ref(x_1); +x_7 = l_Lean_IR_EmitC_emitAllocCtor___redArg___closed__0; +x_8 = lean_string_append(x_2, x_7); +x_9 = l_Nat_reprFast(x_3); +x_10 = lean_string_append(x_8, x_9); +lean_dec_ref(x_9); +x_11 = l_Lean_IR_EmitC_emitSpreadArg___closed__1; +x_12 = lean_string_append(x_10, x_11); +x_13 = l_Nat_reprFast(x_4); +x_14 = lean_string_append(x_12, x_13); +lean_dec_ref(x_13); +x_15 = lean_string_append(x_14, x_11); +x_16 = l_Lean_IR_EmitC_emitCtorScalarSize___redArg(x_5, x_6, x_15); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_18 = lean_ctor_get(x_16, 1); +x_19 = lean_ctor_get(x_16, 0); +lean_dec(x_19); +x_20 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_21 = lean_string_append(x_18, x_20); +x_22 = l_Lean_IR_EmitC_emitLn___closed__0; +x_23 = lean_box(0); +x_24 = lean_string_append(x_21, x_22); +lean_ctor_set(x_16, 1, x_24); +lean_ctor_set(x_16, 0, x_23); +return x_16; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_25 = lean_ctor_get(x_16, 1); +lean_inc(x_25); +lean_dec(x_16); +x_26 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_27 = lean_string_append(x_25, x_26); +x_28 = l_Lean_IR_EmitC_emitLn___closed__0; +x_29 = lean_box(0); +x_30 = lean_string_append(x_27, x_28); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitAllocCtor(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_IR_EmitC_emitAllocCtor___redArg(x_1, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitAllocCtor___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_IR_EmitC_emitAllocCtor(x_1, x_2, x_3); +lean_dec_ref(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitCtorSetArgs_spec__0___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_nat_dec_eq(x_4, x_6); +if (x_7 == 1) +{ +lean_object* x_8; lean_object* x_9; +lean_dec(x_4); +lean_dec(x_1); +x_8 = lean_box(0); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_5); +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_sub(x_4, x_10); +lean_dec(x_4); +x_12 = lean_nat_sub(x_3, x_11); +x_13 = lean_nat_sub(x_12, x_10); +lean_dec(x_12); +x_14 = l_Lean_IR_EmitC_emitSet___redArg___closed__0; +x_15 = lean_string_append(x_5, x_14); +x_16 = l_Lean_IR_EmitC_argToCString___closed__0; +lean_inc(x_1); +x_17 = l_Nat_reprFast(x_1); +x_18 = lean_string_append(x_16, x_17); +lean_dec_ref(x_17); +x_19 = lean_string_append(x_15, x_18); +lean_dec_ref(x_18); +x_20 = l_Lean_IR_EmitC_emitSpreadArg___closed__1; +x_21 = lean_string_append(x_19, x_20); +lean_inc(x_13); +x_22 = l_Nat_reprFast(x_13); +x_23 = lean_string_append(x_21, x_22); +lean_dec_ref(x_22); +x_24 = lean_string_append(x_23, x_20); +x_25 = lean_array_fget_borrowed(x_2, x_13); +lean_dec(x_13); +lean_inc(x_25); +x_26 = l_Lean_IR_EmitC_emitArg___redArg(x_25, x_24); +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +lean_dec_ref(x_26); +x_28 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_29 = lean_string_append(x_27, x_28); +x_30 = l_Lean_IR_EmitC_emitLn___closed__0; +x_31 = lean_string_append(x_29, x_30); +x_4 = x_11; +x_5 = x_31; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitCtorSetArgs_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitCtorSetArgs_spec__0___redArg(x_1, x_2, x_3, x_4, x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCtorSetArgs(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_array_get_size(x_2); +x_6 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitCtorSetArgs_spec__0___redArg(x_1, x_2, x_5, x_5, x_4); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitCtorSetArgs_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitCtorSetArgs_spec__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec_ref(x_6); +lean_dec(x_3); +lean_dec_ref(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCtorSetArgs___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_IR_EmitC_emitCtorSetArgs(x_1, x_2, x_3, x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitCtorSetArgs_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitCtorSetArgs_spec__0___redArg(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_3); +lean_dec_ref(x_2); +return x_6; +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitCtor_spec__0___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; uint8_t x_18; +x_18 = lean_nat_dec_lt(x_9, x_1); +if (x_18 == 0) +{ +lean_object* x_19; +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_10); +lean_ctor_set(x_19, 1, x_11); +return x_19; +} +else +{ +lean_object* x_20; +lean_inc(x_2); +x_20 = lean_array_get_borrowed(x_2, x_3, x_9); +switch (lean_obj_tag(x_20)) { +case 6: +{ +x_12 = x_4; +x_13 = x_11; +goto block_17; +} +case 13: +{ +x_12 = x_4; +x_13 = x_11; +goto block_17; +} +default: +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; 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_21 = l_Lean_IR_EmitC_argToCString___closed__0; +lean_inc(x_5); +x_22 = l_Nat_reprFast(x_5); +x_23 = lean_string_append(x_21, x_22); +lean_dec_ref(x_22); +x_24 = lean_string_append(x_11, x_23); +lean_dec_ref(x_23); +x_25 = l_Lean_IR_EmitC_emitUSet___closed__0; +x_26 = lean_string_append(x_24, x_25); +lean_inc(x_6); +x_27 = l_Nat_reprFast(x_6); +x_28 = lean_string_append(x_26, x_27); +lean_dec_ref(x_27); +x_29 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__0; +x_30 = lean_string_append(x_28, x_29); +lean_inc(x_9); +x_31 = l_Nat_reprFast(x_9); +x_32 = lean_string_append(x_30, x_31); +lean_dec_ref(x_31); +x_33 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__1; +x_34 = lean_string_append(x_32, x_33); +x_35 = lean_array_fget_borrowed(x_7, x_9); +lean_inc(x_35); +x_36 = l_Lean_IR_EmitC_emitArg___redArg(x_35, x_34); +if (lean_obj_tag(x_36) == 0) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_37 = lean_ctor_get(x_36, 1); +lean_inc(x_37); +lean_dec_ref(x_36); +x_38 = lean_string_append(x_37, x_8); +x_39 = l_Lean_IR_EmitC_emitLn___closed__0; +x_40 = lean_string_append(x_38, x_39); +x_12 = x_4; +x_13 = x_40; +goto block_17; +} +else +{ +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +return x_36; +} +} +} +} +block_17: +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_unsigned_to_nat(1u); +x_15 = lean_nat_add(x_9, x_14); +lean_dec(x_9); +x_9 = x_15; +x_10 = x_12; +x_11 = x_13; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitCtor_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +lean_object* x_16; +x_16 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitCtor_spec__0___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_11, x_12, x_15); +return x_16; +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitCtor_spec__1___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; uint8_t x_16; +x_16 = lean_nat_dec_lt(x_7, x_1); +if (x_16 == 0) +{ +lean_object* x_17; +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_2); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_8); +lean_ctor_set(x_17, 1, x_9); +return x_17; +} +else +{ +lean_object* x_18; +lean_inc(x_2); +x_18 = lean_array_get_borrowed(x_2, x_3, x_7); +switch (lean_obj_tag(x_18)) { +case 6: +{ +x_10 = x_4; +x_11 = x_9; +goto block_15; +} +case 13: +{ +x_10 = x_4; +x_11 = x_9; +goto block_15; +} +default: +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_19 = l_Lean_IR_EmitC_argToCString___closed__0; +lean_inc(x_5); +x_20 = l_Nat_reprFast(x_5); +x_21 = lean_string_append(x_19, x_20); +lean_dec_ref(x_20); +x_22 = lean_string_append(x_9, x_21); +lean_dec_ref(x_21); +x_23 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__0; +x_24 = lean_string_append(x_22, x_23); +lean_inc(x_7); +x_25 = l_Nat_reprFast(x_7); +x_26 = lean_string_append(x_24, x_25); +lean_dec_ref(x_25); +x_27 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__1; +x_28 = lean_string_append(x_26, x_27); +x_29 = lean_array_fget_borrowed(x_6, x_7); +lean_inc(x_29); +x_30 = l_Lean_IR_EmitC_emitArg___redArg(x_29, x_28); +if (lean_obj_tag(x_30) == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec_ref(x_30); +x_32 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_33 = lean_string_append(x_31, x_32); +x_34 = l_Lean_IR_EmitC_emitLn___closed__0; +x_35 = lean_string_append(x_33, x_34); +x_10 = x_4; +x_11 = x_35; +goto block_15; +} +else +{ +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_2); +return x_30; +} +} +} +} +block_15: +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_unsigned_to_nat(1u); +x_13 = lean_nat_add(x_7, x_12); +lean_dec(x_7); +x_7 = x_13; +x_8 = x_10; +x_9 = x_11; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitCtor_spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitCtor_spec__1___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_9, x_10, x_13); +return x_14; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitCtor___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_box(", 9, 9); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitCtor___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(".tag = ", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitCtor___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.IR.EmitC.emitCtor", 22, 22); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitCtor___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_IR_EmitC_emitUSet___closed__5; +x_2 = lean_unsigned_to_nat(44u); +x_3 = lean_unsigned_to_nat(511u); +x_4 = l_Lean_IR_EmitC_emitCtor___closed__2; +x_5 = l_Lean_IR_EmitC_emitUSet___closed__3; +x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCtor(lean_object* x_1, 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_13; lean_object* x_14; uint8_t x_15; lean_object* x_31; lean_object* x_32; lean_object* x_41; +x_41 = lean_box(0); +if (lean_obj_tag(x_2) == 11) +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_42 = lean_ctor_get(x_2, 1); +x_43 = lean_ctor_get(x_3, 1); +x_44 = l_Lean_IR_EmitC_argToCString___closed__0; +lean_inc(x_1); +x_45 = l_Nat_reprFast(x_1); +x_46 = lean_string_append(x_44, x_45); +lean_dec_ref(x_45); +x_47 = lean_string_append(x_6, x_46); +lean_dec_ref(x_46); +x_48 = l_Lean_IR_EmitC_emitCtor___closed__1; +x_49 = lean_string_append(x_47, x_48); +lean_inc(x_43); +x_50 = l_Nat_reprFast(x_43); +x_51 = lean_string_append(x_49, x_50); +lean_dec_ref(x_50); +x_52 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_53 = lean_string_append(x_51, x_52); +x_54 = l_Lean_IR_EmitC_emitLn___closed__0; +x_55 = lean_string_append(x_53, x_54); +x_56 = lean_array_get_borrowed(x_41, x_42, x_43); +if (lean_obj_tag(x_56) == 10) +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +lean_inc(x_43); +lean_dec_ref(x_5); +lean_dec_ref(x_3); +x_57 = lean_ctor_get(x_56, 1); +x_58 = lean_array_get_size(x_4); +x_59 = lean_unsigned_to_nat(0u); +x_60 = lean_box(0); +x_61 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitCtor_spec__0___redArg(x_58, x_41, x_57, x_60, x_1, x_43, x_4, x_52, x_59, x_60, x_55); +if (lean_obj_tag(x_61) == 0) +{ +uint8_t x_62; +x_62 = !lean_is_exclusive(x_61); +if (x_62 == 0) +{ +lean_object* x_63; +x_63 = lean_ctor_get(x_61, 0); +lean_dec(x_63); +lean_ctor_set(x_61, 0, x_60); +return x_61; +} +else +{ +lean_object* x_64; lean_object* x_65; +x_64 = lean_ctor_get(x_61, 1); +lean_inc(x_64); +lean_dec(x_61); +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_60); +lean_ctor_set(x_65, 1, x_64); +return x_65; +} +} +else +{ +return x_61; +} +} +else +{ +lean_object* x_66; lean_object* x_67; +x_66 = l_Lean_IR_EmitC_emitCtor___closed__3; +lean_inc_ref(x_5); +x_67 = l_panic___at___00Lean_IR_EmitC_emitUSet_spec__0(x_66, x_5, x_55); +if (lean_obj_tag(x_67) == 0) +{ +lean_object* x_68; +x_68 = lean_ctor_get(x_67, 1); +lean_inc(x_68); +lean_dec_ref(x_67); +x_31 = x_5; +x_32 = x_68; +goto block_40; +} +else +{ +lean_dec_ref(x_5); +lean_dec_ref(x_3); +lean_dec(x_1); +return x_67; +} +} +} +else +{ +if (lean_obj_tag(x_2) == 10) +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; +lean_dec_ref(x_5); +lean_dec_ref(x_3); +x_69 = lean_ctor_get(x_2, 1); +x_70 = lean_array_get_size(x_4); +x_71 = lean_unsigned_to_nat(0u); +x_72 = lean_box(0); +x_73 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitCtor_spec__1___redArg(x_70, x_41, x_69, x_72, x_1, x_4, x_71, x_72, x_6); +if (lean_obj_tag(x_73) == 0) +{ +uint8_t x_74; +x_74 = !lean_is_exclusive(x_73); +if (x_74 == 0) +{ +lean_object* x_75; +x_75 = lean_ctor_get(x_73, 0); +lean_dec(x_75); +lean_ctor_set(x_73, 0, x_72); +return x_73; +} +else +{ +lean_object* x_76; lean_object* x_77; +x_76 = lean_ctor_get(x_73, 1); +lean_inc(x_76); +lean_dec(x_73); +x_77 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_77, 0, x_72); +lean_ctor_set(x_77, 1, x_76); +return x_77; +} +} +else +{ +return x_73; +} +} +else +{ +x_31 = x_5; +x_32 = x_6; +goto block_40; +} +} +block_12: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = l_Lean_IR_EmitC_emitAllocCtor___redArg(x_3, x_8); +x_10 = lean_ctor_get(x_9, 1); +lean_inc(x_10); +lean_dec_ref(x_9); +x_11 = l_Lean_IR_EmitC_emitCtorSetArgs(x_1, x_4, x_7, x_10); +lean_dec_ref(x_7); +return x_11; +} +block_30: +{ +if (x_15 == 0) +{ +x_7 = x_14; +x_8 = x_13; +goto block_12; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_16 = lean_ctor_get(x_3, 1); +x_17 = lean_ctor_get(x_3, 4); +x_18 = lean_unsigned_to_nat(0u); +x_19 = lean_nat_dec_eq(x_17, x_18); +if (x_19 == 0) +{ +x_7 = x_14; +x_8 = x_13; +goto block_12; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_inc(x_16); +lean_dec_ref(x_14); +lean_dec_ref(x_3); +lean_dec(x_1); +x_20 = l_Lean_IR_EmitC_emitCtor___closed__0; +x_21 = lean_string_append(x_13, x_20); +x_22 = l_Nat_reprFast(x_16); +x_23 = lean_string_append(x_21, x_22); +lean_dec_ref(x_22); +x_24 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_25 = lean_string_append(x_23, x_24); +x_26 = l_Lean_IR_EmitC_emitLn___closed__0; +x_27 = lean_box(0); +x_28 = lean_string_append(x_25, x_26); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; +} +} +} +block_40: +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; +lean_inc(x_1); +x_33 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_32); +x_34 = lean_ctor_get(x_33, 1); +lean_inc(x_34); +lean_dec_ref(x_33); +x_35 = lean_ctor_get(x_3, 2); +x_36 = lean_ctor_get(x_3, 3); +x_37 = lean_unsigned_to_nat(0u); +x_38 = lean_nat_dec_eq(x_35, x_37); +if (x_38 == 0) +{ +x_13 = x_34; +x_14 = x_31; +x_15 = x_38; +goto block_30; +} +else +{ +uint8_t x_39; +x_39 = lean_nat_dec_eq(x_36, x_37); +x_13 = x_34; +x_14 = x_31; +x_15 = x_39; +goto block_30; +} +} +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitCtor_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +lean_object* x_16; +x_16 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitCtor_spec__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec_ref(x_14); +lean_dec_ref(x_8); +lean_dec_ref(x_7); +lean_dec_ref(x_3); +lean_dec(x_1); +return x_16; +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitCtor_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitCtor_spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec_ref(x_12); +lean_dec_ref(x_6); +lean_dec_ref(x_3); +lean_dec(x_1); +return x_14; +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitCtor_spec__1___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitCtor_spec__1___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec_ref(x_6); +lean_dec_ref(x_3); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitCtor_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitCtor_spec__0___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec_ref(x_8); +lean_dec_ref(x_7); +lean_dec_ref(x_3); +lean_dec(x_1); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCtor___boxed(lean_object* x_1, 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_IR_EmitC_emitCtor(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec_ref(x_4); +lean_dec(x_2); +return x_7; +} +} +static lean_object* _init_l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitReset_spec__0___redArg___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" lean_ctor_release(", 19, 19); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitReset_spec__0___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; uint8_t x_6; +x_5 = lean_unsigned_to_nat(0u); +x_6 = lean_nat_dec_eq(x_3, x_5); +if (x_6 == 1) +{ +lean_object* x_7; lean_object* x_8; +lean_dec(x_3); +lean_dec(x_1); +x_7 = lean_box(0); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_4); +return x_8; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_9 = lean_unsigned_to_nat(1u); +x_10 = lean_nat_sub(x_3, x_9); +lean_dec(x_3); +x_11 = lean_nat_sub(x_2, x_10); +x_12 = lean_nat_sub(x_11, x_9); +lean_dec(x_11); +x_13 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitReset_spec__0___redArg___closed__0; +x_14 = lean_string_append(x_4, x_13); +x_15 = l_Lean_IR_EmitC_argToCString___closed__0; +lean_inc(x_1); +x_16 = l_Nat_reprFast(x_1); +x_17 = lean_string_append(x_15, x_16); +lean_dec_ref(x_16); +x_18 = lean_string_append(x_14, x_17); +lean_dec_ref(x_17); +x_19 = l_Lean_IR_EmitC_emitSpreadArg___closed__1; +x_20 = lean_string_append(x_18, x_19); +x_21 = l_Nat_reprFast(x_12); +x_22 = lean_string_append(x_20, x_21); +lean_dec_ref(x_21); +x_23 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_24 = lean_string_append(x_22, x_23); +x_25 = l_Lean_IR_EmitC_emitLn___closed__0; +x_26 = lean_string_append(x_24, x_25); +x_3 = x_10; +x_4 = x_26; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitReset_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitReset_spec__0___redArg(x_1, x_2, x_3, x_6); +return x_7; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitReset___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("if (lean_is_exclusive(", 22, 22); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitReset___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(")) {", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitReset___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" lean_dec_ref(", 14, 14); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitReset___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_box(0);", 12, 12); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitReset(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_6 = l_Lean_IR_EmitC_emitReset___closed__0; +x_7 = lean_string_append(x_5, x_6); +x_8 = l_Lean_IR_EmitC_argToCString___closed__0; +lean_inc(x_3); +x_9 = l_Nat_reprFast(x_3); +x_10 = lean_string_append(x_8, x_9); +lean_dec_ref(x_9); +x_11 = lean_string_append(x_7, x_10); +x_12 = l_Lean_IR_EmitC_emitReset___closed__1; +x_13 = lean_string_append(x_11, x_12); +x_14 = l_Lean_IR_EmitC_emitLn___closed__0; +x_15 = lean_string_append(x_13, x_14); +lean_inc(x_2); +x_16 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitReset_spec__0___redArg(x_3, x_2, x_2, x_15); +lean_dec(x_2); +x_17 = lean_ctor_get(x_16, 1); +lean_inc(x_17); +lean_dec_ref(x_16); +x_18 = l_Lean_IR_EmitC_emitSpreadArg___closed__0; +x_19 = lean_string_append(x_17, x_18); +lean_inc(x_1); +x_20 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_19); +x_21 = lean_ctor_get(x_20, 1); +lean_inc(x_21); +lean_dec_ref(x_20); +x_22 = lean_string_append(x_21, x_10); +x_23 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_24 = lean_string_append(x_22, x_23); +x_25 = lean_string_append(x_24, x_14); +x_26 = l_Lean_IR_EmitC_emitMainFn___closed__2; +x_27 = lean_string_append(x_25, x_26); +x_28 = lean_string_append(x_27, x_14); +x_29 = l_Lean_IR_EmitC_emitReset___closed__2; +x_30 = lean_string_append(x_28, x_29); +x_31 = lean_string_append(x_30, x_10); +lean_dec_ref(x_10); +x_32 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_33 = lean_string_append(x_31, x_32); +x_34 = lean_string_append(x_33, x_14); +x_35 = lean_string_append(x_34, x_18); +x_36 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_35); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_38 = lean_ctor_get(x_36, 1); +x_39 = lean_ctor_get(x_36, 0); +lean_dec(x_39); +x_40 = l_Lean_IR_EmitC_emitReset___closed__3; +x_41 = lean_string_append(x_38, x_40); +x_42 = lean_string_append(x_41, x_14); +x_43 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; +x_44 = lean_string_append(x_42, x_43); +x_45 = lean_box(0); +x_46 = lean_string_append(x_44, x_14); +lean_ctor_set(x_36, 1, x_46); +lean_ctor_set(x_36, 0, x_45); +return x_36; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_47 = lean_ctor_get(x_36, 1); +lean_inc(x_47); +lean_dec(x_36); +x_48 = l_Lean_IR_EmitC_emitReset___closed__3; +x_49 = lean_string_append(x_47, x_48); +x_50 = lean_string_append(x_49, x_14); +x_51 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; +x_52 = lean_string_append(x_50, x_51); +x_53 = lean_box(0); +x_54 = lean_string_append(x_52, x_14); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +return x_55; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitReset_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitReset_spec__0(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec_ref(x_5); +lean_dec(x_2); +return x_7; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitReset_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitReset_spec__0___redArg(x_1, x_2, x_3, x_4); +lean_dec(x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitReset___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_IR_EmitC_emitReset(x_1, x_2, x_3, x_4, x_5); +lean_dec_ref(x_4); +return x_6; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitReuse___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("if (lean_is_scalar(", 19, 19); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitReuse___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" lean_ctor_set_tag(", 19, 19); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitReuse(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_16 = l_Lean_IR_EmitC_emitReuse___closed__0; +x_17 = lean_string_append(x_7, x_16); +x_18 = l_Lean_IR_EmitC_argToCString___closed__0; +x_19 = l_Nat_reprFast(x_2); +x_20 = lean_string_append(x_18, x_19); +lean_dec_ref(x_19); +x_21 = lean_string_append(x_17, x_20); +x_22 = l_Lean_IR_EmitC_emitReset___closed__1; +x_23 = lean_string_append(x_21, x_22); +x_24 = l_Lean_IR_EmitC_emitLn___closed__0; +x_25 = lean_string_append(x_23, x_24); +x_26 = l_Lean_IR_EmitC_emitSpreadArg___closed__0; +x_27 = lean_string_append(x_25, x_26); +lean_inc(x_1); +x_28 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_27); +x_29 = lean_ctor_get(x_28, 1); +lean_inc(x_29); +lean_dec_ref(x_28); +lean_inc_ref(x_3); +x_30 = l_Lean_IR_EmitC_emitAllocCtor___redArg(x_3, x_29); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec_ref(x_30); +x_32 = l_Lean_IR_EmitC_emitMainFn___closed__2; +x_33 = lean_string_append(x_31, x_32); +x_34 = lean_string_append(x_33, x_24); +x_35 = lean_string_append(x_34, x_26); +lean_inc(x_1); +x_36 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_35); +x_37 = lean_ctor_get(x_36, 1); +lean_inc(x_37); +lean_dec_ref(x_36); +x_38 = lean_string_append(x_37, x_20); +lean_dec_ref(x_20); +x_39 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_40 = lean_string_append(x_38, x_39); +x_41 = lean_string_append(x_40, x_24); +if (x_4 == 0) +{ +lean_dec_ref(x_3); +x_8 = x_6; +x_9 = x_41; +goto block_15; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_42 = lean_ctor_get(x_3, 1); +lean_inc(x_42); +lean_dec_ref(x_3); +x_43 = l_Lean_IR_EmitC_emitReuse___closed__1; +x_44 = lean_string_append(x_41, x_43); +lean_inc(x_1); +x_45 = l_Nat_reprFast(x_1); +x_46 = lean_string_append(x_18, x_45); +lean_dec_ref(x_45); +x_47 = lean_string_append(x_44, x_46); +lean_dec_ref(x_46); +x_48 = l_Lean_IR_EmitC_emitSpreadArg___closed__1; +x_49 = lean_string_append(x_47, x_48); +x_50 = l_Nat_reprFast(x_42); +x_51 = lean_string_append(x_49, x_50); +lean_dec_ref(x_50); +x_52 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_53 = lean_string_append(x_51, x_52); +x_54 = lean_string_append(x_53, x_24); +x_8 = x_6; +x_9 = x_54; +goto block_15; +} +block_15: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_10 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; +x_11 = lean_string_append(x_9, x_10); +x_12 = l_Lean_IR_EmitC_emitLn___closed__0; +x_13 = lean_string_append(x_11, x_12); +x_14 = l_Lean_IR_EmitC_emitCtorSetArgs(x_1, x_5, x_8, x_13); +return x_14; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitReuse___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +uint8_t x_8; lean_object* x_9; +x_8 = lean_unbox(x_4); +x_9 = l_Lean_IR_EmitC_emitReuse(x_1, x_2, x_3, x_8, x_5, x_6, x_7); +lean_dec_ref(x_6); +lean_dec_ref(x_5); +return x_9; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitProj___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_ctor_get(", 14, 14); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitProj(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; uint8_t x_8; +x_7 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_6); +x_8 = !lean_is_exclusive(x_7); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_9 = lean_ctor_get(x_7, 1); +x_10 = lean_ctor_get(x_7, 0); +lean_dec(x_10); +x_11 = lean_ctor_get(x_5, 5); +x_12 = lean_box(0); +x_13 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0___redArg(x_12, x_11, x_4); +if (lean_obj_tag(x_13) == 10) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_dec_ref(x_13); +lean_dec(x_2); +x_14 = l_Lean_IR_EmitC_argToCString___closed__0; +x_15 = l_Nat_reprFast(x_4); +x_16 = lean_string_append(x_14, x_15); +lean_dec_ref(x_15); +x_17 = lean_string_append(x_9, x_16); +lean_dec_ref(x_16); +x_18 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__0; +x_19 = lean_string_append(x_17, x_18); +x_20 = l_Nat_reprFast(x_3); +x_21 = lean_string_append(x_19, x_20); +lean_dec_ref(x_20); +x_22 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_23 = lean_string_append(x_21, x_22); +x_24 = l_Lean_IR_EmitC_emitLn___closed__0; +x_25 = lean_box(0); +x_26 = lean_string_append(x_23, x_24); +lean_ctor_set(x_7, 1, x_26); +lean_ctor_set(x_7, 0, x_25); +return x_7; +} +else +{ +if (lean_obj_tag(x_13) == 11) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +lean_dec_ref(x_13); +x_27 = l_Lean_IR_EmitC_argToCString___closed__0; +x_28 = l_Nat_reprFast(x_4); +x_29 = lean_string_append(x_27, x_28); +lean_dec_ref(x_28); +x_30 = lean_string_append(x_9, x_29); +lean_dec_ref(x_29); +x_31 = l_Lean_IR_EmitC_emitUSet___closed__0; +x_32 = lean_string_append(x_30, x_31); +x_33 = l_Nat_reprFast(x_2); +x_34 = lean_string_append(x_32, x_33); +lean_dec_ref(x_33); +x_35 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__0; +x_36 = lean_string_append(x_34, x_35); +x_37 = l_Nat_reprFast(x_3); +x_38 = lean_string_append(x_36, x_37); +lean_dec_ref(x_37); +x_39 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_40 = lean_string_append(x_38, x_39); +x_41 = l_Lean_IR_EmitC_emitLn___closed__0; +x_42 = lean_box(0); +x_43 = lean_string_append(x_40, x_41); +lean_ctor_set(x_7, 1, x_43); +lean_ctor_set(x_7, 0, x_42); +return x_7; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +lean_dec(x_13); +lean_dec(x_2); +x_44 = l_Lean_IR_EmitC_emitProj___closed__0; +x_45 = lean_string_append(x_9, x_44); +x_46 = l_Lean_IR_EmitC_argToCString___closed__0; +x_47 = l_Nat_reprFast(x_4); +x_48 = lean_string_append(x_46, x_47); +lean_dec_ref(x_47); +x_49 = lean_string_append(x_45, x_48); +lean_dec_ref(x_48); +x_50 = l_Lean_IR_EmitC_emitSpreadArg___closed__1; +x_51 = lean_string_append(x_49, x_50); +x_52 = l_Nat_reprFast(x_3); +x_53 = lean_string_append(x_51, x_52); +lean_dec_ref(x_52); +x_54 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_55 = lean_string_append(x_53, x_54); +x_56 = l_Lean_IR_EmitC_emitLn___closed__0; +x_57 = lean_box(0); +x_58 = lean_string_append(x_55, x_56); +lean_ctor_set(x_7, 1, x_58); +lean_ctor_set(x_7, 0, x_57); +return x_7; +} +} +} +else +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_59 = lean_ctor_get(x_7, 1); +lean_inc(x_59); +lean_dec(x_7); +x_60 = lean_ctor_get(x_5, 5); +x_61 = lean_box(0); +x_62 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0___redArg(x_61, x_60, x_4); +if (lean_obj_tag(x_62) == 10) +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +lean_dec_ref(x_62); +lean_dec(x_2); +x_63 = l_Lean_IR_EmitC_argToCString___closed__0; +x_64 = l_Nat_reprFast(x_4); +x_65 = lean_string_append(x_63, x_64); +lean_dec_ref(x_64); +x_66 = lean_string_append(x_59, x_65); +lean_dec_ref(x_65); +x_67 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__0; +x_68 = lean_string_append(x_66, x_67); +x_69 = l_Nat_reprFast(x_3); +x_70 = lean_string_append(x_68, x_69); +lean_dec_ref(x_69); +x_71 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_72 = lean_string_append(x_70, x_71); +x_73 = l_Lean_IR_EmitC_emitLn___closed__0; +x_74 = lean_box(0); +x_75 = lean_string_append(x_72, x_73); +x_76 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_76, 0, x_74); +lean_ctor_set(x_76, 1, x_75); +return x_76; +} +else +{ +if (lean_obj_tag(x_62) == 11) +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +lean_dec_ref(x_62); +x_77 = l_Lean_IR_EmitC_argToCString___closed__0; +x_78 = l_Nat_reprFast(x_4); +x_79 = lean_string_append(x_77, x_78); +lean_dec_ref(x_78); +x_80 = lean_string_append(x_59, x_79); +lean_dec_ref(x_79); +x_81 = l_Lean_IR_EmitC_emitUSet___closed__0; +x_82 = lean_string_append(x_80, x_81); +x_83 = l_Nat_reprFast(x_2); +x_84 = lean_string_append(x_82, x_83); +lean_dec_ref(x_83); +x_85 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__0; +x_86 = lean_string_append(x_84, x_85); +x_87 = l_Nat_reprFast(x_3); +x_88 = lean_string_append(x_86, x_87); +lean_dec_ref(x_87); +x_89 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_90 = lean_string_append(x_88, x_89); +x_91 = l_Lean_IR_EmitC_emitLn___closed__0; +x_92 = lean_box(0); +x_93 = lean_string_append(x_90, x_91); +x_94 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_94, 0, x_92); +lean_ctor_set(x_94, 1, x_93); +return x_94; +} +else +{ +lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; +lean_dec(x_62); +lean_dec(x_2); +x_95 = l_Lean_IR_EmitC_emitProj___closed__0; +x_96 = lean_string_append(x_59, x_95); +x_97 = l_Lean_IR_EmitC_argToCString___closed__0; +x_98 = l_Nat_reprFast(x_4); +x_99 = lean_string_append(x_97, x_98); +lean_dec_ref(x_98); +x_100 = lean_string_append(x_96, x_99); +lean_dec_ref(x_99); +x_101 = l_Lean_IR_EmitC_emitSpreadArg___closed__1; +x_102 = lean_string_append(x_100, x_101); +x_103 = l_Nat_reprFast(x_3); +x_104 = lean_string_append(x_102, x_103); +lean_dec_ref(x_103); +x_105 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_106 = lean_string_append(x_104, x_105); +x_107 = l_Lean_IR_EmitC_emitLn___closed__0; +x_108 = lean_box(0); +x_109 = lean_string_append(x_106, x_107); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_108); +lean_ctor_set(x_110, 1, x_109); +return x_110; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitProj___boxed(lean_object* x_1, 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_IR_EmitC_emitProj(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec_ref(x_5); +return x_7; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitUProj___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("];", 2, 2); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitUProj___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.IR.EmitC.emitUProj", 23, 23); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitUProj___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_IR_EmitC_emitUSet___closed__5; +x_2 = lean_unsigned_to_nat(42u); +x_3 = lean_unsigned_to_nat(563u); +x_4 = l_Lean_IR_EmitC_emitUProj___closed__1; +x_5 = l_Lean_IR_EmitC_emitUSet___closed__3; +x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); +return x_6; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitUProj___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_ctor_get_usize(", 20, 20); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUProj(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; uint8_t x_8; +x_7 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_6); +x_8 = !lean_is_exclusive(x_7); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_9 = lean_ctor_get(x_7, 1); +x_10 = lean_ctor_get(x_7, 0); +lean_dec(x_10); +x_11 = lean_ctor_get(x_5, 5); +x_12 = lean_box(0); +x_13 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0___redArg(x_12, x_11, x_4); +if (lean_obj_tag(x_13) == 11) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_13, 1); +lean_inc_ref(x_14); +lean_dec_ref(x_13); +x_15 = lean_array_get(x_12, x_14, x_2); +lean_dec_ref(x_14); +if (lean_obj_tag(x_15) == 10) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +lean_dec_ref(x_5); +x_16 = lean_ctor_get(x_15, 1); +lean_inc_ref(x_16); +lean_dec_ref(x_15); +x_17 = l_Lean_IR_EmitC_argToCString___closed__0; +x_18 = l_Nat_reprFast(x_4); +x_19 = lean_string_append(x_17, x_18); +lean_dec_ref(x_18); +x_20 = lean_string_append(x_9, x_19); +lean_dec_ref(x_19); +x_21 = l_Lean_IR_EmitC_emitUSet___closed__0; +x_22 = lean_string_append(x_20, x_21); +x_23 = l_Nat_reprFast(x_2); +x_24 = lean_string_append(x_22, x_23); +lean_dec_ref(x_23); +x_25 = l_Lean_IR_EmitC_emitUSet___closed__1; +x_26 = lean_string_append(x_24, x_25); +x_27 = lean_array_get_size(x_16); +lean_dec_ref(x_16); +x_28 = lean_nat_sub(x_3, x_27); +lean_dec(x_3); +x_29 = l_Nat_reprFast(x_28); +x_30 = lean_string_append(x_26, x_29); +lean_dec_ref(x_29); +x_31 = l_Lean_IR_EmitC_emitUProj___closed__0; +x_32 = lean_string_append(x_30, x_31); +x_33 = l_Lean_IR_EmitC_emitLn___closed__0; +x_34 = lean_box(0); +x_35 = lean_string_append(x_32, x_33); +lean_ctor_set(x_7, 1, x_35); +lean_ctor_set(x_7, 0, x_34); +return x_7; +} +else +{ +lean_object* x_36; lean_object* x_37; +lean_dec(x_15); +lean_free_object(x_7); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_36 = l_Lean_IR_EmitC_emitUProj___closed__2; +x_37 = l_panic___at___00Lean_IR_EmitC_emitUSet_spec__0(x_36, x_5, x_9); +return x_37; +} +} +else +{ +lean_dec_ref(x_5); +lean_dec(x_2); +if (lean_obj_tag(x_13) == 10) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_38 = lean_ctor_get(x_13, 1); +lean_inc_ref(x_38); +lean_dec_ref(x_13); +x_39 = l_Lean_IR_EmitC_argToCString___closed__0; +x_40 = l_Nat_reprFast(x_4); +x_41 = lean_string_append(x_39, x_40); +lean_dec_ref(x_40); +x_42 = lean_string_append(x_9, x_41); +lean_dec_ref(x_41); +x_43 = l_Lean_IR_EmitC_emitUSet___closed__1; +x_44 = lean_string_append(x_42, x_43); +x_45 = lean_array_get_size(x_38); +lean_dec_ref(x_38); +x_46 = lean_nat_sub(x_3, x_45); +lean_dec(x_3); +x_47 = l_Nat_reprFast(x_46); +x_48 = lean_string_append(x_44, x_47); +lean_dec_ref(x_47); +x_49 = l_Lean_IR_EmitC_emitUProj___closed__0; +x_50 = lean_string_append(x_48, x_49); +x_51 = l_Lean_IR_EmitC_emitLn___closed__0; +x_52 = lean_box(0); +x_53 = lean_string_append(x_50, x_51); +lean_ctor_set(x_7, 1, x_53); +lean_ctor_set(x_7, 0, x_52); +return x_7; +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +lean_dec(x_13); +x_54 = l_Lean_IR_EmitC_emitUProj___closed__3; +x_55 = lean_string_append(x_9, x_54); +x_56 = l_Lean_IR_EmitC_argToCString___closed__0; +x_57 = l_Nat_reprFast(x_4); +x_58 = lean_string_append(x_56, x_57); +lean_dec_ref(x_57); +x_59 = lean_string_append(x_55, x_58); +lean_dec_ref(x_58); +x_60 = l_Lean_IR_EmitC_emitSpreadArg___closed__1; +x_61 = lean_string_append(x_59, x_60); +x_62 = l_Nat_reprFast(x_3); +x_63 = lean_string_append(x_61, x_62); +lean_dec_ref(x_62); +x_64 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_65 = lean_string_append(x_63, x_64); +x_66 = l_Lean_IR_EmitC_emitLn___closed__0; +x_67 = lean_box(0); +x_68 = lean_string_append(x_65, x_66); +lean_ctor_set(x_7, 1, x_68); +lean_ctor_set(x_7, 0, x_67); +return x_7; +} +} +} +else +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_69 = lean_ctor_get(x_7, 1); +lean_inc(x_69); +lean_dec(x_7); +x_70 = lean_ctor_get(x_5, 5); +x_71 = lean_box(0); +x_72 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0___redArg(x_71, x_70, x_4); +if (lean_obj_tag(x_72) == 11) +{ +lean_object* x_73; lean_object* x_74; +x_73 = lean_ctor_get(x_72, 1); +lean_inc_ref(x_73); +lean_dec_ref(x_72); +x_74 = lean_array_get(x_71, x_73, x_2); +lean_dec_ref(x_73); +if (lean_obj_tag(x_74) == 10) +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +lean_dec_ref(x_5); +x_75 = lean_ctor_get(x_74, 1); +lean_inc_ref(x_75); +lean_dec_ref(x_74); +x_76 = l_Lean_IR_EmitC_argToCString___closed__0; +x_77 = l_Nat_reprFast(x_4); +x_78 = lean_string_append(x_76, x_77); +lean_dec_ref(x_77); +x_79 = lean_string_append(x_69, x_78); +lean_dec_ref(x_78); +x_80 = l_Lean_IR_EmitC_emitUSet___closed__0; +x_81 = lean_string_append(x_79, x_80); +x_82 = l_Nat_reprFast(x_2); +x_83 = lean_string_append(x_81, x_82); +lean_dec_ref(x_82); +x_84 = l_Lean_IR_EmitC_emitUSet___closed__1; +x_85 = lean_string_append(x_83, x_84); +x_86 = lean_array_get_size(x_75); +lean_dec_ref(x_75); +x_87 = lean_nat_sub(x_3, x_86); +lean_dec(x_3); +x_88 = l_Nat_reprFast(x_87); +x_89 = lean_string_append(x_85, x_88); +lean_dec_ref(x_88); +x_90 = l_Lean_IR_EmitC_emitUProj___closed__0; +x_91 = lean_string_append(x_89, x_90); +x_92 = l_Lean_IR_EmitC_emitLn___closed__0; +x_93 = lean_box(0); +x_94 = lean_string_append(x_91, x_92); +x_95 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_95, 0, x_93); +lean_ctor_set(x_95, 1, x_94); +return x_95; +} +else +{ +lean_object* x_96; lean_object* x_97; +lean_dec(x_74); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_96 = l_Lean_IR_EmitC_emitUProj___closed__2; +x_97 = l_panic___at___00Lean_IR_EmitC_emitUSet_spec__0(x_96, x_5, x_69); +return x_97; +} +} +else +{ +lean_dec_ref(x_5); +lean_dec(x_2); +if (lean_obj_tag(x_72) == 10) +{ +lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; +x_98 = lean_ctor_get(x_72, 1); +lean_inc_ref(x_98); +lean_dec_ref(x_72); +x_99 = l_Lean_IR_EmitC_argToCString___closed__0; +x_100 = l_Nat_reprFast(x_4); +x_101 = lean_string_append(x_99, x_100); +lean_dec_ref(x_100); +x_102 = lean_string_append(x_69, x_101); +lean_dec_ref(x_101); +x_103 = l_Lean_IR_EmitC_emitUSet___closed__1; +x_104 = lean_string_append(x_102, x_103); +x_105 = lean_array_get_size(x_98); +lean_dec_ref(x_98); +x_106 = lean_nat_sub(x_3, x_105); +lean_dec(x_3); +x_107 = l_Nat_reprFast(x_106); +x_108 = lean_string_append(x_104, x_107); +lean_dec_ref(x_107); +x_109 = l_Lean_IR_EmitC_emitUProj___closed__0; +x_110 = lean_string_append(x_108, x_109); +x_111 = l_Lean_IR_EmitC_emitLn___closed__0; +x_112 = lean_box(0); +x_113 = lean_string_append(x_110, x_111); +x_114 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_114, 0, x_112); +lean_ctor_set(x_114, 1, x_113); +return x_114; +} +else +{ +lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; +lean_dec(x_72); +x_115 = l_Lean_IR_EmitC_emitUProj___closed__3; +x_116 = lean_string_append(x_69, x_115); +x_117 = l_Lean_IR_EmitC_argToCString___closed__0; +x_118 = l_Nat_reprFast(x_4); +x_119 = lean_string_append(x_117, x_118); +lean_dec_ref(x_118); +x_120 = lean_string_append(x_116, x_119); +lean_dec_ref(x_119); +x_121 = l_Lean_IR_EmitC_emitSpreadArg___closed__1; +x_122 = lean_string_append(x_120, x_121); +x_123 = l_Nat_reprFast(x_3); +x_124 = lean_string_append(x_122, x_123); +lean_dec_ref(x_123); +x_125 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_126 = lean_string_append(x_124, x_125); +x_127 = l_Lean_IR_EmitC_emitLn___closed__0; +x_128 = lean_box(0); +x_129 = lean_string_append(x_126, x_127); +x_130 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_130, 0, x_128); +lean_ctor_set(x_130, 1, x_129); +return x_130; +} +} +} +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitSProj___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("));", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitSProj___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_ctor_get_float", 19, 19); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitSProj___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_ctor_get_float32", 21, 21); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitSProj___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_ctor_get_uint8", 19, 19); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitSProj___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_ctor_get_uint16", 20, 20); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitSProj___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_ctor_get_uint32", 20, 20); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitSProj___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_ctor_get_uint64", 20, 20); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSProj(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_35; uint8_t x_36; +x_35 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_8); +x_36 = !lean_is_exclusive(x_35); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_37 = lean_ctor_get(x_35, 1); +x_38 = lean_ctor_get(x_35, 0); +lean_dec(x_38); +x_39 = lean_ctor_get(x_7, 5); +x_40 = lean_box(0); +x_41 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0___redArg(x_40, x_39, x_6); +if (lean_obj_tag(x_41) == 11) +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; +lean_dec_ref(x_41); +lean_free_object(x_35); +lean_dec(x_4); +x_42 = l_Lean_IR_EmitC_emitSSet___closed__0; +x_43 = lean_string_append(x_37, x_42); +x_44 = l_Lean_IR_EmitC_toCType(x_2, x_7, x_43); +x_45 = !lean_is_exclusive(x_44); +if (x_45 == 0) +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_46 = lean_ctor_get(x_44, 0); +x_47 = lean_ctor_get(x_44, 1); +x_48 = lean_string_append(x_47, x_46); +lean_dec(x_46); +x_49 = l_Lean_IR_EmitC_emitSSet___closed__1; +x_50 = lean_string_append(x_48, x_49); +x_51 = l_Lean_IR_EmitC_argToCString___closed__0; +x_52 = l_Nat_reprFast(x_6); +x_53 = lean_string_append(x_51, x_52); +lean_dec_ref(x_52); +x_54 = lean_string_append(x_50, x_53); +lean_dec_ref(x_53); +x_55 = l_Lean_IR_EmitC_emitUSet___closed__0; +x_56 = lean_string_append(x_54, x_55); +x_57 = l_Nat_reprFast(x_3); +x_58 = lean_string_append(x_56, x_57); +lean_dec_ref(x_57); +x_59 = l_Lean_IR_EmitC_emitSSet___closed__2; +x_60 = lean_string_append(x_58, x_59); +x_61 = l_Nat_reprFast(x_5); +x_62 = lean_string_append(x_60, x_61); +lean_dec_ref(x_61); +x_63 = l_Lean_IR_EmitC_emitSProj___closed__0; +x_64 = lean_string_append(x_62, x_63); +x_65 = l_Lean_IR_EmitC_emitLn___closed__0; +x_66 = lean_string_append(x_64, x_65); +x_67 = lean_box(0); +lean_ctor_set(x_44, 1, x_66); +lean_ctor_set(x_44, 0, x_67); +return x_44; +} +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_68 = lean_ctor_get(x_44, 0); +x_69 = lean_ctor_get(x_44, 1); +lean_inc(x_69); +lean_inc(x_68); +lean_dec(x_44); +x_70 = lean_string_append(x_69, x_68); +lean_dec(x_68); +x_71 = l_Lean_IR_EmitC_emitSSet___closed__1; +x_72 = lean_string_append(x_70, x_71); +x_73 = l_Lean_IR_EmitC_argToCString___closed__0; +x_74 = l_Nat_reprFast(x_6); +x_75 = lean_string_append(x_73, x_74); +lean_dec_ref(x_74); +x_76 = lean_string_append(x_72, x_75); +lean_dec_ref(x_75); +x_77 = l_Lean_IR_EmitC_emitUSet___closed__0; +x_78 = lean_string_append(x_76, x_77); +x_79 = l_Nat_reprFast(x_3); +x_80 = lean_string_append(x_78, x_79); +lean_dec_ref(x_79); +x_81 = l_Lean_IR_EmitC_emitSSet___closed__2; +x_82 = lean_string_append(x_80, x_81); +x_83 = l_Nat_reprFast(x_5); +x_84 = lean_string_append(x_82, x_83); +lean_dec_ref(x_83); +x_85 = l_Lean_IR_EmitC_emitSProj___closed__0; +x_86 = lean_string_append(x_84, x_85); +x_87 = l_Lean_IR_EmitC_emitLn___closed__0; +x_88 = lean_string_append(x_86, x_87); +x_89 = lean_box(0); +x_90 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_90, 0, x_89); +lean_ctor_set(x_90, 1, x_88); +return x_90; +} +} +else +{ +lean_dec(x_3); +if (lean_obj_tag(x_41) == 10) +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; uint8_t x_94; +lean_dec_ref(x_41); +lean_free_object(x_35); +lean_dec(x_4); +x_91 = l_Lean_IR_EmitC_emitSSet___closed__0; +x_92 = lean_string_append(x_37, x_91); +x_93 = l_Lean_IR_EmitC_toCType(x_2, x_7, x_92); +x_94 = !lean_is_exclusive(x_93); +if (x_94 == 0) +{ +lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_95 = lean_ctor_get(x_93, 0); +x_96 = lean_ctor_get(x_93, 1); +x_97 = lean_string_append(x_96, x_95); +lean_dec(x_95); +x_98 = l_Lean_IR_EmitC_emitSSet___closed__1; +x_99 = lean_string_append(x_97, x_98); +x_100 = l_Lean_IR_EmitC_argToCString___closed__0; +x_101 = l_Nat_reprFast(x_6); +x_102 = lean_string_append(x_100, x_101); +lean_dec_ref(x_101); +x_103 = lean_string_append(x_99, x_102); +lean_dec_ref(x_102); +x_104 = l_Lean_IR_EmitC_emitSSet___closed__2; +x_105 = lean_string_append(x_103, x_104); +x_106 = l_Nat_reprFast(x_5); +x_107 = lean_string_append(x_105, x_106); +lean_dec_ref(x_106); +x_108 = l_Lean_IR_EmitC_emitSProj___closed__0; +x_109 = lean_string_append(x_107, x_108); +x_110 = l_Lean_IR_EmitC_emitLn___closed__0; +x_111 = lean_string_append(x_109, x_110); +x_112 = lean_box(0); +lean_ctor_set(x_93, 1, x_111); +lean_ctor_set(x_93, 0, x_112); +return x_93; +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; +x_113 = lean_ctor_get(x_93, 0); +x_114 = lean_ctor_get(x_93, 1); +lean_inc(x_114); +lean_inc(x_113); +lean_dec(x_93); +x_115 = lean_string_append(x_114, x_113); +lean_dec(x_113); +x_116 = l_Lean_IR_EmitC_emitSSet___closed__1; +x_117 = lean_string_append(x_115, x_116); +x_118 = l_Lean_IR_EmitC_argToCString___closed__0; +x_119 = l_Nat_reprFast(x_6); +x_120 = lean_string_append(x_118, x_119); +lean_dec_ref(x_119); +x_121 = lean_string_append(x_117, x_120); +lean_dec_ref(x_120); +x_122 = l_Lean_IR_EmitC_emitSSet___closed__2; +x_123 = lean_string_append(x_121, x_122); +x_124 = l_Nat_reprFast(x_5); +x_125 = lean_string_append(x_123, x_124); +lean_dec_ref(x_124); +x_126 = l_Lean_IR_EmitC_emitSProj___closed__0; +x_127 = lean_string_append(x_125, x_126); +x_128 = l_Lean_IR_EmitC_emitLn___closed__0; +x_129 = lean_string_append(x_127, x_128); +x_130 = lean_box(0); +x_131 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_131, 0, x_130); +lean_ctor_set(x_131, 1, x_129); +return x_131; +} +} +else +{ +lean_dec(x_41); +switch (lean_obj_tag(x_2)) { +case 0: +{ +lean_object* x_132; lean_object* x_133; +lean_free_object(x_35); +x_132 = l_Lean_IR_EmitC_emitSProj___closed__1; +x_133 = lean_string_append(x_37, x_132); +x_9 = x_133; +goto block_34; +} +case 9: +{ +lean_object* x_134; lean_object* x_135; +lean_free_object(x_35); +x_134 = l_Lean_IR_EmitC_emitSProj___closed__2; +x_135 = lean_string_append(x_37, x_134); +x_9 = x_135; +goto block_34; +} +case 1: +{ +lean_object* x_136; lean_object* x_137; +lean_free_object(x_35); +x_136 = l_Lean_IR_EmitC_emitSProj___closed__3; +x_137 = lean_string_append(x_37, x_136); +x_9 = x_137; +goto block_34; +} +case 2: +{ +lean_object* x_138; lean_object* x_139; +lean_free_object(x_35); +x_138 = l_Lean_IR_EmitC_emitSProj___closed__4; +x_139 = lean_string_append(x_37, x_138); +x_9 = x_139; +goto block_34; +} +case 3: +{ +lean_object* x_140; lean_object* x_141; +lean_free_object(x_35); +x_140 = l_Lean_IR_EmitC_emitSProj___closed__5; +x_141 = lean_string_append(x_37, x_140); +x_9 = x_141; +goto block_34; +} +case 4: +{ +lean_object* x_142; lean_object* x_143; +lean_free_object(x_35); +x_142 = l_Lean_IR_EmitC_emitSProj___closed__6; +x_143 = lean_string_append(x_37, x_142); +x_9 = x_143; +goto block_34; +} +default: +{ +lean_object* x_144; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_144 = l_Lean_IR_EmitC_emitSSet___closed__10; +lean_ctor_set_tag(x_35, 1); +lean_ctor_set(x_35, 0, x_144); +return x_35; +} +} +} +} +} +else +{ +lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; +x_145 = lean_ctor_get(x_35, 1); +lean_inc(x_145); +lean_dec(x_35); +x_146 = lean_ctor_get(x_7, 5); +x_147 = lean_box(0); +x_148 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_emitInc_spec__0___redArg(x_147, x_146, x_6); +if (lean_obj_tag(x_148) == 11) +{ +lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; +lean_dec_ref(x_148); +lean_dec(x_4); +x_149 = l_Lean_IR_EmitC_emitSSet___closed__0; +x_150 = lean_string_append(x_145, x_149); +x_151 = l_Lean_IR_EmitC_toCType(x_2, x_7, x_150); +x_152 = lean_ctor_get(x_151, 0); +lean_inc(x_152); +x_153 = lean_ctor_get(x_151, 1); +lean_inc(x_153); +if (lean_is_exclusive(x_151)) { + lean_ctor_release(x_151, 0); + lean_ctor_release(x_151, 1); + x_154 = x_151; +} else { + lean_dec_ref(x_151); + x_154 = lean_box(0); +} +x_155 = lean_string_append(x_153, x_152); +lean_dec(x_152); +x_156 = l_Lean_IR_EmitC_emitSSet___closed__1; +x_157 = lean_string_append(x_155, x_156); +x_158 = l_Lean_IR_EmitC_argToCString___closed__0; +x_159 = l_Nat_reprFast(x_6); +x_160 = lean_string_append(x_158, x_159); +lean_dec_ref(x_159); +x_161 = lean_string_append(x_157, x_160); +lean_dec_ref(x_160); +x_162 = l_Lean_IR_EmitC_emitUSet___closed__0; +x_163 = lean_string_append(x_161, x_162); +x_164 = l_Nat_reprFast(x_3); +x_165 = lean_string_append(x_163, x_164); +lean_dec_ref(x_164); +x_166 = l_Lean_IR_EmitC_emitSSet___closed__2; +x_167 = lean_string_append(x_165, x_166); +x_168 = l_Nat_reprFast(x_5); +x_169 = lean_string_append(x_167, x_168); +lean_dec_ref(x_168); +x_170 = l_Lean_IR_EmitC_emitSProj___closed__0; +x_171 = lean_string_append(x_169, x_170); +x_172 = l_Lean_IR_EmitC_emitLn___closed__0; +x_173 = lean_string_append(x_171, x_172); +x_174 = lean_box(0); +if (lean_is_scalar(x_154)) { + x_175 = lean_alloc_ctor(0, 2, 0); +} else { + x_175 = x_154; +} +lean_ctor_set(x_175, 0, x_174); +lean_ctor_set(x_175, 1, x_173); +return x_175; +} +else +{ +lean_dec(x_3); +if (lean_obj_tag(x_148) == 10) +{ +lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; +lean_dec_ref(x_148); +lean_dec(x_4); +x_176 = l_Lean_IR_EmitC_emitSSet___closed__0; +x_177 = lean_string_append(x_145, x_176); +x_178 = l_Lean_IR_EmitC_toCType(x_2, x_7, x_177); +x_179 = lean_ctor_get(x_178, 0); +lean_inc(x_179); +x_180 = lean_ctor_get(x_178, 1); +lean_inc(x_180); +if (lean_is_exclusive(x_178)) { + lean_ctor_release(x_178, 0); + lean_ctor_release(x_178, 1); + x_181 = x_178; +} else { + lean_dec_ref(x_178); + x_181 = lean_box(0); +} +x_182 = lean_string_append(x_180, x_179); +lean_dec(x_179); +x_183 = l_Lean_IR_EmitC_emitSSet___closed__1; +x_184 = lean_string_append(x_182, x_183); +x_185 = l_Lean_IR_EmitC_argToCString___closed__0; +x_186 = l_Nat_reprFast(x_6); +x_187 = lean_string_append(x_185, x_186); +lean_dec_ref(x_186); +x_188 = lean_string_append(x_184, x_187); +lean_dec_ref(x_187); +x_189 = l_Lean_IR_EmitC_emitSSet___closed__2; +x_190 = lean_string_append(x_188, x_189); +x_191 = l_Nat_reprFast(x_5); +x_192 = lean_string_append(x_190, x_191); +lean_dec_ref(x_191); +x_193 = l_Lean_IR_EmitC_emitSProj___closed__0; +x_194 = lean_string_append(x_192, x_193); +x_195 = l_Lean_IR_EmitC_emitLn___closed__0; +x_196 = lean_string_append(x_194, x_195); +x_197 = lean_box(0); +if (lean_is_scalar(x_181)) { + x_198 = lean_alloc_ctor(0, 2, 0); +} else { + x_198 = x_181; +} +lean_ctor_set(x_198, 0, x_197); +lean_ctor_set(x_198, 1, x_196); +return x_198; +} +else +{ +lean_dec(x_148); +switch (lean_obj_tag(x_2)) { +case 0: +{ +lean_object* x_199; lean_object* x_200; +x_199 = l_Lean_IR_EmitC_emitSProj___closed__1; +x_200 = lean_string_append(x_145, x_199); +x_9 = x_200; +goto block_34; +} +case 9: +{ +lean_object* x_201; lean_object* x_202; +x_201 = l_Lean_IR_EmitC_emitSProj___closed__2; +x_202 = lean_string_append(x_145, x_201); +x_9 = x_202; +goto block_34; +} +case 1: +{ +lean_object* x_203; lean_object* x_204; +x_203 = l_Lean_IR_EmitC_emitSProj___closed__3; +x_204 = lean_string_append(x_145, x_203); +x_9 = x_204; +goto block_34; +} +case 2: +{ +lean_object* x_205; lean_object* x_206; +x_205 = l_Lean_IR_EmitC_emitSProj___closed__4; +x_206 = lean_string_append(x_145, x_205); +x_9 = x_206; +goto block_34; +} +case 3: +{ +lean_object* x_207; lean_object* x_208; +x_207 = l_Lean_IR_EmitC_emitSProj___closed__5; +x_208 = lean_string_append(x_145, x_207); +x_9 = x_208; +goto block_34; +} +case 4: +{ +lean_object* x_209; lean_object* x_210; +x_209 = l_Lean_IR_EmitC_emitSProj___closed__6; +x_210 = lean_string_append(x_145, x_209); +x_9 = x_210; +goto block_34; +} +default: +{ +lean_object* x_211; lean_object* x_212; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_211 = l_Lean_IR_EmitC_emitSSet___closed__10; +x_212 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_212, 0, x_211); +lean_ctor_set(x_212, 1, x_145); +return x_212; +} +} +} +} +} +block_34: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_10 = l_Lean_IR_EmitC_emitFnDeclAux___closed__2; +x_11 = lean_string_append(x_9, x_10); +x_12 = l_Lean_IR_EmitC_argToCString___closed__0; +x_13 = l_Nat_reprFast(x_6); +x_14 = lean_string_append(x_12, x_13); +lean_dec_ref(x_13); +x_15 = lean_string_append(x_11, x_14); +lean_dec_ref(x_14); +x_16 = l_Lean_IR_EmitC_emitSpreadArg___closed__1; +x_17 = lean_string_append(x_15, x_16); +x_18 = l_Lean_IR_EmitC_emitOffset___redArg(x_4, x_5, x_17); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_20 = lean_ctor_get(x_18, 1); +x_21 = lean_ctor_get(x_18, 0); +lean_dec(x_21); +x_22 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_23 = lean_string_append(x_20, x_22); +x_24 = l_Lean_IR_EmitC_emitLn___closed__0; +x_25 = lean_box(0); +x_26 = lean_string_append(x_23, x_24); +lean_ctor_set(x_18, 1, x_26); +lean_ctor_set(x_18, 0, x_25); +return x_18; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_27 = lean_ctor_get(x_18, 1); +lean_inc(x_27); +lean_dec(x_18); +x_28 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_29 = lean_string_append(x_27, x_28); +x_30 = l_Lean_IR_EmitC_emitLn___closed__0; +x_31 = lean_box(0); +x_32 = lean_string_append(x_29, x_30); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSProj___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_IR_EmitC_emitSProj(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec_ref(x_7); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at___00Lean_IR_EmitC_toStringArgs_spec__0(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; +x_3 = l_List_reverse___redArg(x_2); +return x_3; +} +else +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_1, 0); +x_6 = lean_ctor_get(x_1, 1); +x_7 = l_Lean_IR_EmitC_argToCString(x_5); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set(x_1, 0, x_7); +{ +lean_object* _tmp_0 = x_6; +lean_object* _tmp_1 = x_1; +x_1 = _tmp_0; +x_2 = _tmp_1; +} +goto _start; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_1, 0); +x_10 = lean_ctor_get(x_1, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_1); +x_11 = l_Lean_IR_EmitC_argToCString(x_9); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_2); +x_1 = x_10; +x_2 = x_12; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_toStringArgs(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_array_to_list(x_1); +x_3 = lean_box(0); +x_4 = l_List_mapTR_loop___at___00Lean_IR_EmitC_toStringArgs_spec__0(x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___at___00Lean_IR_EmitC_emitSimpleExternalCall_spec__0___redArg(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) { +_start: +{ +lean_object* x_8; uint8_t x_9; +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_nat_dec_eq(x_5, x_8); +if (x_9 == 1) +{ +lean_object* x_10; lean_object* x_11; +lean_dec(x_5); +lean_dec_ref(x_2); +x_10 = lean_box(x_6); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_7); +return x_11; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; uint8_t x_23; lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_12 = lean_unsigned_to_nat(1u); +x_13 = lean_nat_sub(x_5, x_12); +lean_dec(x_5); +x_14 = lean_nat_sub(x_4, x_13); +x_15 = lean_nat_sub(x_14, x_12); +lean_dec(x_14); +lean_inc_ref(x_2); +x_28 = lean_array_get_borrowed(x_2, x_3, x_15); +x_29 = lean_ctor_get(x_28, 1); +x_30 = l_Lean_IR_IRType_isErased(x_29); +if (x_30 == 0) +{ +uint8_t x_31; +x_31 = l_Lean_IR_IRType_isVoid(x_29); +x_23 = x_31; +goto block_27; +} +else +{ +x_23 = x_30; +goto block_27; +} +block_22: +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_array_fget_borrowed(x_1, x_15); +lean_dec(x_15); +lean_inc(x_18); +x_19 = l_Lean_IR_EmitC_emitArg___redArg(x_18, x_17); +x_20 = lean_ctor_get(x_19, 1); +lean_inc(x_20); +lean_dec_ref(x_19); +x_5 = x_13; +x_6 = x_16; +x_7 = x_20; +goto _start; +} +block_27: +{ +if (x_23 == 0) +{ +if (x_6 == 0) +{ +lean_object* x_24; lean_object* x_25; +x_24 = l_Lean_IR_EmitC_emitSpreadArg___closed__1; +x_25 = lean_string_append(x_7, x_24); +x_16 = x_23; +x_17 = x_25; +goto block_22; +} +else +{ +x_16 = x_23; +x_17 = x_7; +goto block_22; +} +} +else +{ +lean_dec(x_15); +x_5 = x_13; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___at___00Lean_IR_EmitC_emitSimpleExternalCall_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___at___00Lean_IR_EmitC_emitSimpleExternalCall_spec__0___redArg(x_1, x_2, x_3, x_4, x_5, x_7, x_9); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSimpleExternalCall(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; +x_6 = l_Lean_IR_EmitC_emitFullAppArgs___closed__0; +x_7 = lean_string_append(x_5, x_1); +x_8 = l_Lean_IR_EmitC_emitFnDeclAux___closed__2; +x_9 = lean_string_append(x_7, x_8); +x_10 = lean_array_get_size(x_3); +x_11 = 1; +x_12 = l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___at___00Lean_IR_EmitC_emitSimpleExternalCall_spec__0___redArg(x_3, x_6, x_2, x_10, x_10, x_11, x_9); +if (lean_obj_tag(x_12) == 0) +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_14 = lean_ctor_get(x_12, 1); +x_15 = lean_ctor_get(x_12, 0); +lean_dec(x_15); +x_16 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_17 = lean_string_append(x_14, x_16); +x_18 = l_Lean_IR_EmitC_emitLn___closed__0; +x_19 = lean_string_append(x_17, x_18); +x_20 = lean_box(0); +lean_ctor_set(x_12, 1, x_19); +lean_ctor_set(x_12, 0, x_20); +return x_12; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_21 = lean_ctor_get(x_12, 1); +lean_inc(x_21); +lean_dec(x_12); +x_22 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_23 = lean_string_append(x_21, x_22); +x_24 = l_Lean_IR_EmitC_emitLn___closed__0; +x_25 = lean_string_append(x_23, x_24); +x_26 = lean_box(0); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +return x_27; +} +} +else +{ +uint8_t x_28; +x_28 = !lean_is_exclusive(x_12); +if (x_28 == 0) +{ +return x_12; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_12, 0); +x_30 = lean_ctor_get(x_12, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_12); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___at___00Lean_IR_EmitC_emitSimpleExternalCall_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_7); +x_11 = l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___at___00Lean_IR_EmitC_emitSimpleExternalCall_spec__0(x_1, x_2, x_3, x_4, x_5, x_6, x_10, x_8, x_9); +lean_dec_ref(x_8); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSimpleExternalCall___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_IR_EmitC_emitSimpleExternalCall(x_1, x_2, x_3, x_4, x_5); +lean_dec_ref(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___at___00Lean_IR_EmitC_emitSimpleExternalCall_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +uint8_t x_8; lean_object* x_9; +x_8 = lean_unbox(x_6); +x_9 = l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___at___00Lean_IR_EmitC_emitSimpleExternalCall_spec__0___redArg(x_1, x_2, x_3, x_4, x_5, x_8, x_7); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_1); +return x_9; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitExternCall___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("failed to emit extern application '", 35, 35); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitExternCall(lean_object* x_1, 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_16; lean_object* x_17; +x_16 = l_List_forM___at___00Lean_IR_EmitC_emitFnDecls_spec__3___closed__1; +x_17 = l_Lean_getExternEntryFor(x_3, x_16); +if (lean_obj_tag(x_17) == 1) +{ +lean_object* x_18; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +lean_dec_ref(x_17); +switch (lean_obj_tag(x_18)) { +case 2: +{ +lean_object* x_19; lean_object* x_20; +lean_dec(x_1); +x_19 = lean_ctor_get(x_18, 1); +lean_inc_ref(x_19); +lean_dec_ref(x_18); +x_20 = l_Lean_IR_EmitC_emitSimpleExternalCall(x_19, x_2, x_4, x_5, x_6); +lean_dec_ref(x_4); +lean_dec_ref(x_19); +return x_20; +} +case 1: +{ +uint8_t x_21; +lean_dec(x_1); +x_21 = !lean_is_exclusive(x_18); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_22 = lean_ctor_get(x_18, 1); +x_23 = lean_ctor_get(x_18, 0); +lean_dec(x_23); +x_24 = l_Lean_IR_EmitC_toStringArgs(x_4); +x_25 = l_Lean_expandExternPattern(x_22, x_24); +lean_dec(x_24); +x_26 = lean_string_append(x_6, x_25); +lean_dec_ref(x_25); +x_27 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_28 = lean_string_append(x_26, x_27); +x_29 = l_Lean_IR_EmitC_emitLn___closed__0; +x_30 = lean_box(0); +x_31 = lean_string_append(x_28, x_29); +lean_ctor_set_tag(x_18, 0); +lean_ctor_set(x_18, 1, x_31); +lean_ctor_set(x_18, 0, x_30); +return x_18; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_32 = lean_ctor_get(x_18, 1); +lean_inc(x_32); +lean_dec(x_18); +x_33 = l_Lean_IR_EmitC_toStringArgs(x_4); +x_34 = l_Lean_expandExternPattern(x_32, x_33); +lean_dec(x_33); +x_35 = lean_string_append(x_6, x_34); +lean_dec_ref(x_34); +x_36 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_37 = lean_string_append(x_35, x_36); +x_38 = l_Lean_IR_EmitC_emitLn___closed__0; +x_39 = lean_box(0); +x_40 = lean_string_append(x_37, x_38); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +return x_41; +} +} +default: +{ +lean_dec(x_18); +lean_dec_ref(x_4); +x_7 = x_6; +goto block_15; +} +} +} +else +{ +lean_dec(x_17); +lean_dec_ref(x_4); +x_7 = x_6; +goto block_15; +} +block_15: +{ +lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_8 = l_Lean_IR_EmitC_emitExternCall___closed__0; +x_9 = 1; +x_10 = l_Lean_Name_toStringWithToken___at___00Lean_Name_toString_spec__0(x_1, x_9); +x_11 = lean_string_append(x_8, x_10); +lean_dec_ref(x_10); +x_12 = l_Lean_IR_EmitC_getDecl___closed__1; +x_13 = lean_string_append(x_11, x_12); +x_14 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_7); +return x_14; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitExternCall___boxed(lean_object* x_1, 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_IR_EmitC_emitExternCall(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec_ref(x_5); +lean_dec_ref(x_2); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFullApp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; lean_object* x_14; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_5); +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec_ref(x_22); +lean_inc_ref(x_4); +lean_inc(x_2); +x_24 = l_Lean_IR_EmitC_getDecl(x_2, x_4, x_23); +if (lean_obj_tag(x_24) == 0) +{ +lean_object* x_25; +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec_ref(x_24); +x_27 = lean_ctor_get(x_25, 1); +lean_inc_ref(x_27); +lean_dec_ref(x_25); +lean_inc_ref(x_4); +x_28 = l_Lean_IR_EmitC_emitCName(x_2, x_4, x_26); +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_29 = lean_ctor_get(x_28, 1); +lean_inc(x_29); +lean_dec_ref(x_28); +x_30 = lean_unsigned_to_nat(0u); +x_31 = lean_array_get_size(x_3); +x_32 = lean_nat_dec_lt(x_30, x_31); +if (x_32 == 0) +{ +lean_dec_ref(x_27); +lean_dec_ref(x_4); +lean_dec_ref(x_3); +x_6 = x_29; +goto block_13; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = l_Lean_IR_EmitC_emitFnDeclAux___closed__2; +x_34 = lean_string_append(x_29, x_33); +x_35 = l_Lean_IR_EmitC_emitFullAppArgs(x_27, x_3, x_4, x_34); +lean_dec_ref(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_27); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_35, 1); +lean_inc(x_36); +lean_dec_ref(x_35); +x_37 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__1; +x_38 = lean_string_append(x_36, x_37); +x_6 = x_38; +goto block_13; +} +else +{ +return x_35; +} +} +} +else +{ +lean_dec_ref(x_27); +lean_dec_ref(x_4); +lean_dec_ref(x_3); +return x_28; +} +} +else +{ +lean_object* x_39; +x_39 = lean_ctor_get(x_25, 3); +if (lean_obj_tag(x_39) == 1) +{ +lean_object* x_40; +x_40 = lean_ctor_get(x_39, 0); +if (lean_obj_tag(x_40) == 3) +{ +lean_object* x_41; +x_41 = lean_ctor_get(x_39, 1); +if (lean_obj_tag(x_41) == 0) +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_24, 1); +lean_inc(x_42); +lean_dec_ref(x_24); +x_43 = lean_ctor_get(x_25, 1); +lean_inc_ref(x_43); +lean_dec_ref(x_25); +lean_inc_ref(x_4); +x_44 = l_Lean_IR_EmitC_emitCName(x_2, x_4, x_42); +if (lean_obj_tag(x_44) == 0) +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; +x_45 = lean_ctor_get(x_44, 1); +lean_inc(x_45); +lean_dec_ref(x_44); +x_46 = lean_unsigned_to_nat(0u); +x_47 = lean_array_get_size(x_3); +x_48 = lean_nat_dec_lt(x_46, x_47); +if (x_48 == 0) +{ +lean_dec_ref(x_43); +lean_dec_ref(x_4); +lean_dec_ref(x_3); +x_14 = x_45; +goto block_21; +} +else +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = l_Lean_IR_EmitC_emitFnDeclAux___closed__2; +x_50 = lean_string_append(x_45, x_49); +x_51 = l_Lean_IR_EmitC_emitFullAppArgs(x_43, x_3, x_4, x_50); +lean_dec_ref(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_43); +if (lean_obj_tag(x_51) == 0) +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_51, 1); +lean_inc(x_52); +lean_dec_ref(x_51); +x_53 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__1; +x_54 = lean_string_append(x_52, x_53); +x_14 = x_54; +goto block_21; +} +else +{ +return x_51; +} +} +} +else +{ +lean_dec_ref(x_43); +lean_dec_ref(x_4); +lean_dec_ref(x_3); +return x_44; +} +} +else +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; +lean_inc_ref(x_39); +x_55 = lean_ctor_get(x_24, 1); +lean_inc(x_55); +lean_dec_ref(x_24); +x_56 = lean_ctor_get(x_25, 1); +lean_inc_ref(x_56); +lean_dec_ref(x_25); +x_57 = l_Lean_IR_EmitC_emitExternCall(x_2, x_56, x_39, x_3, x_4, x_55); +lean_dec_ref(x_4); +lean_dec_ref(x_56); +return x_57; +} +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +lean_inc_ref(x_39); +x_58 = lean_ctor_get(x_24, 1); +lean_inc(x_58); +lean_dec_ref(x_24); +x_59 = lean_ctor_get(x_25, 1); +lean_inc_ref(x_59); +lean_dec_ref(x_25); +x_60 = l_Lean_IR_EmitC_emitExternCall(x_2, x_59, x_39, x_3, x_4, x_58); +lean_dec_ref(x_4); +lean_dec_ref(x_59); +return x_60; +} +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; +lean_inc(x_39); +x_61 = lean_ctor_get(x_24, 1); +lean_inc(x_61); +lean_dec_ref(x_24); +x_62 = lean_ctor_get(x_25, 1); +lean_inc_ref(x_62); +lean_dec_ref(x_25); +x_63 = l_Lean_IR_EmitC_emitExternCall(x_2, x_62, x_39, x_3, x_4, x_61); +lean_dec_ref(x_4); +lean_dec_ref(x_62); +return x_63; +} +} +} +else +{ +uint8_t x_64; +lean_dec_ref(x_4); +lean_dec_ref(x_3); +lean_dec(x_2); +x_64 = !lean_is_exclusive(x_24); +if (x_64 == 0) +{ +return x_24; +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_24, 0); +x_66 = lean_ctor_get(x_24, 1); +lean_inc(x_66); +lean_inc(x_65); +lean_dec(x_24); +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_13: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_7 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_8 = lean_string_append(x_6, x_7); +x_9 = l_Lean_IR_EmitC_emitLn___closed__0; +x_10 = lean_box(0); +x_11 = lean_string_append(x_8, x_9); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set(x_12, 1, x_11); +return x_12; +} +block_21: +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_15 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_16 = lean_string_append(x_14, x_15); +x_17 = l_Lean_IR_EmitC_emitLn___closed__0; +x_18 = lean_box(0); +x_19 = lean_string_append(x_16, x_17); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +return x_20; +} +} +} +static lean_object* _init_l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitPartialApp_spec__0___redArg___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_closure_set(", 17, 17); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitPartialApp_spec__0___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; uint8_t x_9; +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_nat_dec_eq(x_6, x_8); +if (x_9 == 1) +{ +lean_object* x_10; lean_object* x_11; +lean_dec(x_6); +lean_dec(x_2); +x_10 = lean_box(0); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_7); +return x_11; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_12 = lean_unsigned_to_nat(1u); +x_13 = lean_nat_sub(x_6, x_12); +lean_dec(x_6); +x_14 = lean_nat_sub(x_5, x_13); +x_15 = lean_nat_sub(x_14, x_12); +lean_dec(x_14); +x_16 = lean_array_fget_borrowed(x_1, x_15); +x_17 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitPartialApp_spec__0___redArg___closed__0; +x_18 = lean_string_append(x_7, x_17); +x_19 = l_Lean_IR_EmitC_argToCString___closed__0; +lean_inc(x_2); +x_20 = l_Nat_reprFast(x_2); +x_21 = lean_string_append(x_19, x_20); +lean_dec_ref(x_20); +x_22 = lean_string_append(x_18, x_21); +lean_dec_ref(x_21); +x_23 = lean_string_append(x_22, x_3); +x_24 = l_Nat_reprFast(x_15); +x_25 = lean_string_append(x_23, x_24); +lean_dec_ref(x_24); +x_26 = lean_string_append(x_25, x_3); +lean_inc(x_16); +x_27 = l_Lean_IR_EmitC_emitArg___redArg(x_16, x_26); +x_28 = lean_ctor_get(x_27, 1); +lean_inc(x_28); +lean_dec_ref(x_27); +x_29 = lean_string_append(x_28, x_4); +x_30 = l_Lean_IR_EmitC_emitLn___closed__0; +x_31 = lean_string_append(x_29, x_30); +x_6 = x_13; +x_7 = x_31; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitPartialApp_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitPartialApp_spec__0___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_9); +return x_10; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitPartialApp___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_alloc_closure((void*)(", 27, 27); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitPartialApp___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("), ", 3, 3); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitPartialApp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +lean_inc_ref(x_4); +lean_inc(x_2); +x_6 = l_Lean_IR_EmitC_getDecl(x_2, x_4, x_5); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +lean_dec_ref(x_6); +lean_inc(x_1); +x_9 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_8); +x_10 = lean_ctor_get(x_9, 1); +lean_inc(x_10); +lean_dec_ref(x_9); +x_11 = l_Lean_IR_EmitC_emitPartialApp___closed__0; +x_12 = lean_string_append(x_10, x_11); +x_13 = l_Lean_IR_EmitC_emitCName(x_2, x_4, x_12); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec_ref(x_13); +x_15 = l_Lean_IR_Decl_params(x_7); +lean_dec(x_7); +x_16 = lean_array_get_size(x_15); +lean_dec_ref(x_15); +x_17 = l_Lean_IR_EmitC_emitPartialApp___closed__1; +x_18 = lean_string_append(x_14, x_17); +x_19 = l_Nat_reprFast(x_16); +x_20 = lean_string_append(x_18, x_19); +lean_dec_ref(x_19); +x_21 = l_Lean_IR_EmitC_emitSpreadArg___closed__1; +x_22 = lean_string_append(x_20, x_21); +x_23 = lean_array_get_size(x_3); +x_24 = l_Nat_reprFast(x_23); +x_25 = lean_string_append(x_22, x_24); +lean_dec_ref(x_24); +x_26 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_27 = lean_string_append(x_25, x_26); +x_28 = l_Lean_IR_EmitC_emitLn___closed__0; +x_29 = lean_string_append(x_27, x_28); +x_30 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitPartialApp_spec__0___redArg(x_3, x_1, x_21, x_26, x_23, x_23, x_29); +return x_30; +} +else +{ +lean_dec(x_7); +lean_dec(x_1); +return x_13; +} +} +else +{ +uint8_t x_31; +lean_dec_ref(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_31 = !lean_is_exclusive(x_6); +if (x_31 == 0) +{ +return x_6; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_6, 0); +x_33 = lean_ctor_get(x_6, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_6); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitPartialApp_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitPartialApp_spec__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec_ref(x_8); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitPartialApp___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_IR_EmitC_emitPartialApp(x_1, x_2, x_3, x_4, x_5); +lean_dec_ref(x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitPartialApp_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitPartialApp_spec__0___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_1); +return x_8; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitApp___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_apply_", 11, 11); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitApp___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("{ lean_object* _aargs[] = {", 27, 27); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitApp___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("};", 2, 2); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitApp___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_apply_m(", 13, 13); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitApp___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(", _aargs); }", 12, 12); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitApp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_6 = l_Lean_IR_EmitC_emitFnDeclAux___closed__1; +x_7 = lean_array_get_size(x_3); +x_8 = lean_nat_dec_lt(x_6, x_7); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_9 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_5); +x_10 = lean_ctor_get(x_9, 1); +lean_inc(x_10); +lean_dec_ref(x_9); +x_11 = l_Lean_IR_EmitC_emitApp___closed__0; +x_12 = lean_string_append(x_10, x_11); +x_13 = l_Nat_reprFast(x_7); +x_14 = lean_string_append(x_12, x_13); +lean_dec_ref(x_13); +x_15 = l_Lean_IR_EmitC_emitFnDeclAux___closed__2; +x_16 = lean_string_append(x_14, x_15); +x_17 = l_Lean_IR_EmitC_argToCString___closed__0; +x_18 = l_Nat_reprFast(x_2); +x_19 = lean_string_append(x_17, x_18); +lean_dec_ref(x_18); +x_20 = lean_string_append(x_16, x_19); +lean_dec_ref(x_19); +x_21 = l_Lean_IR_EmitC_emitSpreadArg___closed__1; +x_22 = lean_string_append(x_20, x_21); +x_23 = l_Lean_IR_EmitC_emitArgs(x_3, x_4, x_22); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_25 = lean_ctor_get(x_23, 1); +x_26 = lean_ctor_get(x_23, 0); +lean_dec(x_26); +x_27 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_28 = lean_string_append(x_25, x_27); +x_29 = l_Lean_IR_EmitC_emitLn___closed__0; +x_30 = lean_box(0); +x_31 = lean_string_append(x_28, x_29); +lean_ctor_set(x_23, 1, x_31); +lean_ctor_set(x_23, 0, x_30); +return x_23; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_32 = lean_ctor_get(x_23, 1); +lean_inc(x_32); +lean_dec(x_23); +x_33 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_34 = lean_string_append(x_32, x_33); +x_35 = l_Lean_IR_EmitC_emitLn___closed__0; +x_36 = lean_box(0); +x_37 = lean_string_append(x_34, x_35); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +return x_38; +} +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; +x_39 = l_Lean_IR_EmitC_emitApp___closed__1; +x_40 = lean_string_append(x_5, x_39); +x_41 = l_Lean_IR_EmitC_emitArgs(x_3, x_4, x_40); +x_42 = lean_ctor_get(x_41, 1); +lean_inc(x_42); +lean_dec_ref(x_41); +x_43 = l_Lean_IR_EmitC_emitApp___closed__2; +x_44 = lean_string_append(x_42, x_43); +x_45 = l_Lean_IR_EmitC_emitLn___closed__0; +x_46 = lean_string_append(x_44, x_45); +x_47 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_46); +x_48 = !lean_is_exclusive(x_47); +if (x_48 == 0) +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_49 = lean_ctor_get(x_47, 1); +x_50 = lean_ctor_get(x_47, 0); +lean_dec(x_50); +x_51 = l_Lean_IR_EmitC_emitApp___closed__3; +x_52 = lean_string_append(x_49, x_51); +x_53 = l_Lean_IR_EmitC_argToCString___closed__0; +x_54 = l_Nat_reprFast(x_2); +x_55 = lean_string_append(x_53, x_54); +lean_dec_ref(x_54); +x_56 = lean_string_append(x_52, x_55); +lean_dec_ref(x_55); +x_57 = l_Lean_IR_EmitC_emitSpreadArg___closed__1; +x_58 = lean_string_append(x_56, x_57); +x_59 = l_Nat_reprFast(x_7); +x_60 = lean_string_append(x_58, x_59); +lean_dec_ref(x_59); +x_61 = l_Lean_IR_EmitC_emitApp___closed__4; +x_62 = lean_string_append(x_60, x_61); +x_63 = lean_box(0); +x_64 = lean_string_append(x_62, x_45); +lean_ctor_set(x_47, 1, x_64); +lean_ctor_set(x_47, 0, x_63); +return x_47; +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_65 = lean_ctor_get(x_47, 1); +lean_inc(x_65); +lean_dec(x_47); +x_66 = l_Lean_IR_EmitC_emitApp___closed__3; +x_67 = lean_string_append(x_65, x_66); +x_68 = l_Lean_IR_EmitC_argToCString___closed__0; +x_69 = l_Nat_reprFast(x_2); +x_70 = lean_string_append(x_68, x_69); +lean_dec_ref(x_69); +x_71 = lean_string_append(x_67, x_70); +lean_dec_ref(x_70); +x_72 = l_Lean_IR_EmitC_emitSpreadArg___closed__1; +x_73 = lean_string_append(x_71, x_72); +x_74 = l_Nat_reprFast(x_7); +x_75 = lean_string_append(x_73, x_74); +lean_dec_ref(x_74); +x_76 = l_Lean_IR_EmitC_emitApp___closed__4; +x_77 = lean_string_append(x_75, x_76); +x_78 = lean_box(0); +x_79 = lean_string_append(x_77, x_45); +x_80 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_80, 0, x_78); +lean_ctor_set(x_80, 1, x_79); +return x_80; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitApp___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_IR_EmitC_emitApp(x_1, x_2, x_3, x_4, x_5); +lean_dec_ref(x_4); +lean_dec_ref(x_3); +return x_6; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitBoxFn___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_box_usize", 14, 14); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitBoxFn___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_box_uint32", 15, 15); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitBoxFn___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_box_uint64", 15, 15); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitBoxFn___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_box_float", 14, 14); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitBoxFn___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_box_float32", 16, 16); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitBoxFn___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_box", 8, 8); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitBoxFn(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 5: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_5 = l_Lean_IR_EmitC_emitBoxFn___closed__0; +x_6 = lean_box(0); +x_7 = lean_string_append(x_4, x_5); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_6); +lean_ctor_set(x_8, 1, x_7); +return x_8; +} +case 3: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = l_Lean_IR_EmitC_emitBoxFn___closed__1; +x_10 = lean_box(0); +x_11 = lean_string_append(x_4, x_9); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set(x_12, 1, x_11); +return x_12; +} +case 4: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = l_Lean_IR_EmitC_emitBoxFn___closed__2; +x_14 = lean_box(0); +x_15 = lean_string_append(x_4, x_13); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +return x_16; +} +case 0: +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = l_Lean_IR_EmitC_emitBoxFn___closed__3; +x_18 = lean_box(0); +x_19 = lean_string_append(x_4, x_17); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +return x_20; +} +case 9: +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_21 = l_Lean_IR_EmitC_emitBoxFn___closed__4; +x_22 = lean_box(0); +x_23 = lean_string_append(x_4, x_21); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +return x_24; +} +case 10: +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_25 = lean_ctor_get(x_3, 6); +x_26 = lean_unsigned_to_nat(0u); +x_27 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0___redArg(x_26, x_25, x_1); +lean_dec_ref(x_1); +x_28 = l_Lean_IR_IRType_isStruct(x_2); +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; +x_29 = l_Lean_IR_EmitC_structBoxFnPrefix___closed__0; +x_30 = lean_string_append(x_4, x_29); +x_31 = lean_box(0); +x_32 = l_Nat_reprFast(x_27); +x_33 = lean_string_append(x_30, x_32); +lean_dec_ref(x_32); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_31); +lean_ctor_set(x_34, 1, x_33); +return x_34; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_35 = l_Lean_IR_EmitC_structReshapeFnPrefix___closed__0; +x_36 = lean_string_append(x_4, x_35); +x_37 = l_Nat_reprFast(x_27); +x_38 = lean_string_append(x_36, x_37); +lean_dec_ref(x_37); +x_39 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadArg_spec__0___redArg___closed__0; +x_40 = lean_string_append(x_38, x_39); +x_41 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0___redArg(x_26, x_25, x_2); +x_42 = lean_box(0); +x_43 = l_Nat_reprFast(x_41); +x_44 = lean_string_append(x_40, x_43); +lean_dec_ref(x_43); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_42); +lean_ctor_set(x_45, 1, x_44); +return x_45; +} +} +case 11: +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; +x_46 = lean_ctor_get(x_3, 6); +x_47 = lean_unsigned_to_nat(0u); +x_48 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0___redArg(x_47, x_46, x_1); +x_49 = !lean_is_exclusive(x_1); +if (x_49 == 0) +{ +lean_object* x_50; lean_object* x_51; uint8_t x_52; +x_50 = lean_ctor_get(x_1, 1); +lean_dec(x_50); +x_51 = lean_ctor_get(x_1, 0); +lean_dec(x_51); +x_52 = l_Lean_IR_IRType_isStruct(x_2); +if (x_52 == 0) +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_53 = l_Lean_IR_EmitC_structBoxFnPrefix___closed__0; +x_54 = lean_string_append(x_4, x_53); +x_55 = lean_box(0); +x_56 = l_Nat_reprFast(x_48); +x_57 = lean_string_append(x_54, x_56); +lean_dec_ref(x_56); +lean_ctor_set_tag(x_1, 0); +lean_ctor_set(x_1, 1, x_57); +lean_ctor_set(x_1, 0, x_55); +return x_1; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_58 = l_Lean_IR_EmitC_structReshapeFnPrefix___closed__0; +x_59 = lean_string_append(x_4, x_58); +x_60 = l_Nat_reprFast(x_48); +x_61 = lean_string_append(x_59, x_60); +lean_dec_ref(x_60); +x_62 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadArg_spec__0___redArg___closed__0; +x_63 = lean_string_append(x_61, x_62); +x_64 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0___redArg(x_47, x_46, x_2); +x_65 = lean_box(0); +x_66 = l_Nat_reprFast(x_64); +x_67 = lean_string_append(x_63, x_66); +lean_dec_ref(x_66); +lean_ctor_set_tag(x_1, 0); +lean_ctor_set(x_1, 1, x_67); +lean_ctor_set(x_1, 0, x_65); +return x_1; +} +} +else +{ +uint8_t x_68; +lean_dec(x_1); +x_68 = l_Lean_IR_IRType_isStruct(x_2); +if (x_68 == 0) +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_69 = l_Lean_IR_EmitC_structBoxFnPrefix___closed__0; +x_70 = lean_string_append(x_4, x_69); +x_71 = lean_box(0); +x_72 = l_Nat_reprFast(x_48); +x_73 = lean_string_append(x_70, x_72); +lean_dec_ref(x_72); +x_74 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_74, 0, x_71); +lean_ctor_set(x_74, 1, x_73); +return x_74; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_75 = l_Lean_IR_EmitC_structReshapeFnPrefix___closed__0; +x_76 = lean_string_append(x_4, x_75); +x_77 = l_Nat_reprFast(x_48); +x_78 = lean_string_append(x_76, x_77); +lean_dec_ref(x_77); +x_79 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadArg_spec__0___redArg___closed__0; +x_80 = lean_string_append(x_78, x_79); +x_81 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0___redArg(x_47, x_46, x_2); +x_82 = lean_box(0); +x_83 = l_Nat_reprFast(x_81); +x_84 = lean_string_append(x_80, x_83); +lean_dec_ref(x_83); +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_82); +lean_ctor_set(x_85, 1, x_84); +return x_85; +} +} +} +default: +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +lean_dec(x_1); +x_86 = l_Lean_IR_EmitC_emitBoxFn___closed__5; +x_87 = lean_box(0); +x_88 = lean_string_append(x_4, x_86); +x_89 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_89, 0, x_87); +lean_ctor_set(x_89, 1, x_88); +return x_89; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitBoxFn___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_IR_EmitC_emitBoxFn(x_1, x_2, x_3, x_4); +lean_dec_ref(x_3); +lean_dec(x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitBox(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_7 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_6); +x_8 = lean_ctor_get(x_7, 1); +lean_inc(x_8); +lean_dec_ref(x_7); +x_9 = l_Lean_IR_EmitC_emitBoxFn(x_3, x_4, x_5, x_8); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_11 = lean_ctor_get(x_9, 1); +x_12 = lean_ctor_get(x_9, 0); +lean_dec(x_12); +x_13 = l_Lean_IR_EmitC_emitFnDeclAux___closed__2; +x_14 = lean_string_append(x_11, x_13); +x_15 = l_Lean_IR_EmitC_argToCString___closed__0; +x_16 = l_Nat_reprFast(x_2); +x_17 = lean_string_append(x_15, x_16); +lean_dec_ref(x_16); +x_18 = lean_string_append(x_14, x_17); +lean_dec_ref(x_17); +x_19 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_20 = lean_string_append(x_18, x_19); +x_21 = l_Lean_IR_EmitC_emitLn___closed__0; +x_22 = lean_box(0); +x_23 = lean_string_append(x_20, x_21); +lean_ctor_set(x_9, 1, x_23); +lean_ctor_set(x_9, 0, x_22); +return x_9; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_24 = lean_ctor_get(x_9, 1); +lean_inc(x_24); +lean_dec(x_9); +x_25 = l_Lean_IR_EmitC_emitFnDeclAux___closed__2; +x_26 = lean_string_append(x_24, x_25); +x_27 = l_Lean_IR_EmitC_argToCString___closed__0; +x_28 = l_Nat_reprFast(x_2); +x_29 = lean_string_append(x_27, x_28); +lean_dec_ref(x_28); +x_30 = lean_string_append(x_26, x_29); +lean_dec_ref(x_29); +x_31 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_32 = lean_string_append(x_30, x_31); +x_33 = l_Lean_IR_EmitC_emitLn___closed__0; +x_34 = lean_box(0); +x_35 = lean_string_append(x_32, x_33); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitBox___boxed(lean_object* x_1, 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_IR_EmitC_emitBox(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUnboxFn(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = l_Lean_IR_IRType_isStruct(x_1); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_5 = l_Lean_IR_getUnboxOpName(x_1); +x_6 = lean_box(0); +x_7 = lean_string_append(x_3, x_5); +lean_dec_ref(x_5); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_6); +lean_ctor_set(x_8, 1, x_7); +return x_8; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_9 = lean_ctor_get(x_2, 6); +x_10 = l_Lean_IR_EmitC_structUnboxFnPrefix___closed__0; +x_11 = lean_string_append(x_3, x_10); +x_12 = lean_unsigned_to_nat(0u); +x_13 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0___redArg(x_12, x_9, x_1); +x_14 = lean_box(0); +x_15 = l_Nat_reprFast(x_13); +x_16 = lean_string_append(x_11, x_15); +lean_dec_ref(x_15); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_14); +lean_ctor_set(x_17, 1, x_16); +return x_17; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUnboxFn___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_IR_EmitC_emitUnboxFn(x_1, x_2, x_3); +lean_dec_ref(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUnbox(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_6 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_5); +x_7 = lean_ctor_get(x_6, 1); +lean_inc(x_7); +lean_dec_ref(x_6); +x_8 = l_Lean_IR_EmitC_emitUnboxFn(x_2, x_4, x_7); +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_10 = lean_ctor_get(x_8, 1); +x_11 = lean_ctor_get(x_8, 0); +lean_dec(x_11); +x_12 = l_Lean_IR_EmitC_emitFnDeclAux___closed__2; +x_13 = lean_string_append(x_10, x_12); +x_14 = l_Lean_IR_EmitC_argToCString___closed__0; +x_15 = l_Nat_reprFast(x_3); +x_16 = lean_string_append(x_14, x_15); +lean_dec_ref(x_15); +x_17 = lean_string_append(x_13, x_16); +lean_dec_ref(x_16); +x_18 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_19 = lean_string_append(x_17, x_18); +x_20 = l_Lean_IR_EmitC_emitLn___closed__0; +x_21 = lean_box(0); +x_22 = lean_string_append(x_19, x_20); +lean_ctor_set(x_8, 1, x_22); +lean_ctor_set(x_8, 0, x_21); +return x_8; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_23 = lean_ctor_get(x_8, 1); +lean_inc(x_23); +lean_dec(x_8); +x_24 = l_Lean_IR_EmitC_emitFnDeclAux___closed__2; +x_25 = lean_string_append(x_23, x_24); +x_26 = l_Lean_IR_EmitC_argToCString___closed__0; +x_27 = l_Nat_reprFast(x_3); +x_28 = lean_string_append(x_26, x_27); +lean_dec_ref(x_27); +x_29 = lean_string_append(x_25, x_28); +lean_dec_ref(x_28); +x_30 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_31 = lean_string_append(x_29, x_30); +x_32 = l_Lean_IR_EmitC_emitLn___closed__0; +x_33 = lean_box(0); +x_34 = lean_string_append(x_31, x_32); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUnbox___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_IR_EmitC_emitUnbox(x_1, x_2, x_3, x_4, x_5); +lean_dec_ref(x_4); +lean_dec(x_2); +return x_6; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitIsShared___redArg___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("!lean_is_exclusive(", 19, 19); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitIsShared___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_3); +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_6 = lean_ctor_get(x_4, 1); +x_7 = lean_ctor_get(x_4, 0); +lean_dec(x_7); +x_8 = l_Lean_IR_EmitC_emitIsShared___redArg___closed__0; +x_9 = lean_string_append(x_6, x_8); +x_10 = l_Lean_IR_EmitC_argToCString___closed__0; +x_11 = l_Nat_reprFast(x_2); +x_12 = lean_string_append(x_10, x_11); +lean_dec_ref(x_11); +x_13 = lean_string_append(x_9, x_12); +lean_dec_ref(x_12); +x_14 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_15 = lean_string_append(x_13, x_14); +x_16 = l_Lean_IR_EmitC_emitLn___closed__0; +x_17 = lean_box(0); +x_18 = lean_string_append(x_15, x_16); +lean_ctor_set(x_4, 1, x_18); +lean_ctor_set(x_4, 0, x_17); +return x_4; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_19 = lean_ctor_get(x_4, 1); +lean_inc(x_19); +lean_dec(x_4); +x_20 = l_Lean_IR_EmitC_emitIsShared___redArg___closed__0; +x_21 = lean_string_append(x_19, x_20); +x_22 = l_Lean_IR_EmitC_argToCString___closed__0; +x_23 = l_Nat_reprFast(x_2); +x_24 = lean_string_append(x_22, x_23); +lean_dec_ref(x_23); +x_25 = lean_string_append(x_21, x_24); +lean_dec_ref(x_24); +x_26 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_27 = lean_string_append(x_25, x_26); +x_28 = l_Lean_IR_EmitC_emitLn___closed__0; +x_29 = lean_box(0); +x_30 = lean_string_append(x_27, x_28); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitIsShared(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_IR_EmitC_emitIsShared___redArg(x_1, x_2, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitIsShared___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_IR_EmitC_emitIsShared(x_1, x_2, x_3, x_4); +lean_dec_ref(x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_toHexDigit(lean_object* x_1) { +_start: +{ +uint32_t x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Nat_digitChar(x_1); +x_3 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__1; +x_4 = lean_string_push(x_3, x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_toHexDigit___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_IR_EmitC_toHexDigit(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("\\x", 2, 2); +return x_1; +} +} +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("\\\?", 2, 2); +return x_1; +} +} +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("\\\"", 2, 2); +return x_1; +} +} +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("\\\\", 2, 2); +return x_1; +} +} +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("\\t", 2, 2); +return x_1; +} +} +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("\\r", 2, 2); +return x_1; +} +} +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("\\n", 2, 2); +return x_1; +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_5 = lean_ctor_get(x_1, 1); +x_6 = lean_ctor_get(x_1, 2); +x_7 = lean_nat_sub(x_6, x_5); +x_8 = lean_nat_dec_eq(x_3, x_7); +lean_dec(x_7); +if (x_8 == 0) +{ +lean_object* x_9; uint32_t x_10; uint32_t x_11; uint8_t x_12; +x_9 = lean_string_utf8_next_fast(x_2, x_3); +x_10 = lean_string_utf8_get_fast(x_2, x_3); +lean_dec(x_3); +x_11 = 10; +x_12 = lean_uint32_dec_eq(x_10, x_11); +if (x_12 == 0) +{ +uint32_t x_13; uint8_t x_14; +x_13 = 13; +x_14 = lean_uint32_dec_eq(x_10, x_13); +if (x_14 == 0) +{ +uint32_t x_15; uint8_t x_16; +x_15 = 9; +x_16 = lean_uint32_dec_eq(x_10, x_15); +if (x_16 == 0) +{ +uint32_t x_17; uint8_t x_18; +x_17 = 92; +x_18 = lean_uint32_dec_eq(x_10, x_17); +if (x_18 == 0) +{ +uint32_t x_19; uint8_t x_20; +x_19 = 34; +x_20 = lean_uint32_dec_eq(x_10, x_19); +if (x_20 == 0) +{ +uint32_t x_21; uint8_t x_22; +x_21 = 63; +x_22 = lean_uint32_dec_eq(x_10, x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_23 = lean_uint32_to_nat(x_10); +x_24 = lean_unsigned_to_nat(31u); +x_25 = lean_nat_dec_le(x_23, x_24); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +lean_dec(x_23); +x_26 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__1; +x_27 = lean_string_push(x_26, x_10); +x_28 = lean_string_append(x_4, x_27); +lean_dec_ref(x_27); +x_3 = x_9; +x_4 = x_28; +goto _start; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_30 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__0; +x_31 = lean_unsigned_to_nat(16u); +x_32 = lean_unsigned_to_nat(4u); +x_33 = lean_nat_shiftr(x_23, x_32); +x_34 = l_Lean_IR_EmitC_toHexDigit(x_33); +lean_dec(x_33); +x_35 = lean_string_append(x_30, x_34); +lean_dec_ref(x_34); +x_36 = lean_nat_mod(x_23, x_31); +lean_dec(x_23); +x_37 = l_Lean_IR_EmitC_toHexDigit(x_36); +lean_dec(x_36); +x_38 = lean_string_append(x_35, x_37); +lean_dec_ref(x_37); +x_39 = lean_string_append(x_4, x_38); +lean_dec_ref(x_38); +x_3 = x_9; +x_4 = x_39; +goto _start; +} +} +else +{ +lean_object* x_41; lean_object* x_42; +x_41 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__1; +x_42 = lean_string_append(x_4, x_41); +x_3 = x_9; +x_4 = x_42; +goto _start; +} +} +else +{ +lean_object* x_44; lean_object* x_45; +x_44 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__2; +x_45 = lean_string_append(x_4, x_44); +x_3 = x_9; +x_4 = x_45; +goto _start; +} +} +else +{ +lean_object* x_47; lean_object* x_48; +x_47 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__3; +x_48 = lean_string_append(x_4, x_47); +x_3 = x_9; +x_4 = x_48; +goto _start; +} +} +else +{ +lean_object* x_50; lean_object* x_51; +x_50 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__4; +x_51 = lean_string_append(x_4, x_50); +x_3 = x_9; +x_4 = x_51; +goto _start; +} +} +else +{ +lean_object* x_53; lean_object* x_54; +x_53 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__5; +x_54 = lean_string_append(x_4, x_53); +x_3 = x_9; +x_4 = x_54; +goto _start; +} +} +else +{ +lean_object* x_56; lean_object* x_57; +x_56 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__6; +x_57 = lean_string_append(x_4, x_56); +x_3 = x_9; +x_4 = x_57; +goto _start; +} +} +else +{ +lean_dec(x_3); +return x_4; +} +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg(x_1, x_2, x_5, x_6); +return x_8; +} +} +static lean_object* _init_l_Lean_IR_EmitC_quoteString___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("\"", 1, 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_quoteString(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_2 = l_Lean_IR_EmitC_quoteString___closed__0; +x_3 = lean_unsigned_to_nat(0u); +x_4 = lean_string_utf8_byte_size(x_1); +lean_inc_ref(x_1); +x_5 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_3); +lean_ctor_set(x_5, 2, x_4); +x_6 = l_String_Slice_positions(x_5); +x_7 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg(x_5, x_1, x_6, x_2); +lean_dec_ref(x_1); +lean_dec_ref(x_5); +x_8 = lean_string_append(x_7, x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___boxed(lean_object* x_1, lean_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_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___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_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg(x_1, x_2, x_3, x_4); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitNumLit___redArg___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ULL", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitNumLit___redArg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("((size_t)", 9, 9); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitNumLit___redArg___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ULL)", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitNumLit___redArg___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_cstr_to_nat(\"", 18, 18); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitNumLit___redArg___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("\")", 2, 2); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitNumLit___redArg___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_unsigned_to_nat(", 21, 21); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitNumLit___redArg___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("u)", 2, 2); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitNumLit___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = l_Lean_IR_IRType_isObj(x_1); +if (x_4 == 0) +{ +lean_object* x_5; uint8_t x_6; +x_5 = lean_cstr_to_nat("4294967296"); +x_6 = lean_nat_dec_lt(x_2, x_5); +if (x_6 == 0) +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_box(5); +x_8 = l_Lean_IR_instBEqIRType_beq(x_1, x_7); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_9 = l_Nat_reprFast(x_2); +x_10 = lean_string_append(x_3, x_9); +lean_dec_ref(x_9); +x_11 = l_Lean_IR_EmitC_emitNumLit___redArg___closed__0; +x_12 = lean_box(0); +x_13 = lean_string_append(x_10, x_11); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_15 = l_Lean_IR_EmitC_emitNumLit___redArg___closed__1; +x_16 = lean_string_append(x_3, x_15); +x_17 = l_Nat_reprFast(x_2); +x_18 = lean_string_append(x_16, x_17); +lean_dec_ref(x_17); +x_19 = l_Lean_IR_EmitC_emitNumLit___redArg___closed__2; +x_20 = lean_box(0); +x_21 = lean_string_append(x_18, x_19); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_dec(x_1); +x_23 = lean_box(0); +x_24 = l_Nat_reprFast(x_2); +x_25 = lean_string_append(x_3, x_24); +lean_dec_ref(x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_23); +lean_ctor_set(x_26, 1, x_25); +return x_26; +} +} +else +{ +lean_object* x_27; uint8_t x_28; +lean_dec(x_1); +x_27 = lean_cstr_to_nat("4294967296"); +x_28 = lean_nat_dec_lt(x_2, x_27); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_29 = l_Lean_IR_EmitC_emitNumLit___redArg___closed__3; +x_30 = lean_string_append(x_3, x_29); +x_31 = l_Nat_reprFast(x_2); +x_32 = lean_string_append(x_30, x_31); +lean_dec_ref(x_31); +x_33 = l_Lean_IR_EmitC_emitNumLit___redArg___closed__4; +x_34 = lean_box(0); +x_35 = lean_string_append(x_32, x_33); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_37 = l_Lean_IR_EmitC_emitNumLit___redArg___closed__5; +x_38 = lean_string_append(x_3, x_37); +x_39 = l_Nat_reprFast(x_2); +x_40 = lean_string_append(x_38, x_39); +lean_dec_ref(x_39); +x_41 = l_Lean_IR_EmitC_emitNumLit___redArg___closed__6; +x_42 = lean_box(0); +x_43 = lean_string_append(x_40, x_41); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_43); +return x_44; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitNumLit(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_IR_EmitC_emitNumLit___redArg(x_1, x_2, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitNumLit___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_IR_EmitC_emitNumLit(x_1, x_2, x_3, x_4); +lean_dec_ref(x_3); +return x_5; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitLit___redArg___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_mk_string_unchecked(", 25, 25); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLit___redArg(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_IR_EmitC_emitLhs___redArg(x_1, x_4); +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_6 = lean_ctor_get(x_5, 1); +lean_inc(x_6); +lean_dec_ref(x_5); +x_7 = lean_ctor_get(x_3, 0); +lean_inc(x_7); +lean_dec_ref(x_3); +x_8 = l_Lean_IR_EmitC_emitNumLit___redArg(x_2, x_7, x_6); +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_10 = lean_ctor_get(x_8, 1); +x_11 = lean_ctor_get(x_8, 0); +lean_dec(x_11); +x_12 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_13 = lean_string_append(x_10, x_12); +x_14 = l_Lean_IR_EmitC_emitLn___closed__0; +x_15 = lean_box(0); +x_16 = lean_string_append(x_13, x_14); +lean_ctor_set(x_8, 1, x_16); +lean_ctor_set(x_8, 0, x_15); +return x_8; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_17 = lean_ctor_get(x_8, 1); +lean_inc(x_17); +lean_dec(x_8); +x_18 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_19 = lean_string_append(x_17, x_18); +x_20 = l_Lean_IR_EmitC_emitLn___closed__0; +x_21 = lean_box(0); +x_22 = lean_string_append(x_19, x_20); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +else +{ +uint8_t x_24; +lean_dec(x_2); +x_24 = !lean_is_exclusive(x_5); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_25 = lean_ctor_get(x_5, 1); +x_26 = lean_ctor_get(x_5, 0); +lean_dec(x_26); +x_27 = lean_ctor_get(x_3, 0); +lean_inc_ref(x_27); +lean_dec_ref(x_3); +x_28 = l_Lean_IR_EmitC_emitLit___redArg___closed__0; +x_29 = lean_string_append(x_25, x_28); +lean_inc_ref(x_27); +x_30 = l_Lean_IR_EmitC_quoteString(x_27); +x_31 = lean_string_append(x_29, x_30); +lean_dec_ref(x_30); +x_32 = l_Lean_IR_EmitC_emitSpreadArg___closed__1; +x_33 = lean_string_append(x_31, x_32); +x_34 = lean_string_utf8_byte_size(x_27); +x_35 = l_Nat_reprFast(x_34); +x_36 = lean_string_append(x_33, x_35); +lean_dec_ref(x_35); +x_37 = lean_string_append(x_36, x_32); +x_38 = lean_string_length(x_27); +lean_dec_ref(x_27); +x_39 = l_Nat_reprFast(x_38); +x_40 = lean_string_append(x_37, x_39); +lean_dec_ref(x_39); +x_41 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_42 = lean_string_append(x_40, x_41); +x_43 = l_Lean_IR_EmitC_emitLn___closed__0; +x_44 = lean_box(0); +x_45 = lean_string_append(x_42, x_43); +lean_ctor_set(x_5, 1, x_45); +lean_ctor_set(x_5, 0, x_44); +return x_5; +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_46 = lean_ctor_get(x_5, 1); +lean_inc(x_46); +lean_dec(x_5); +x_47 = lean_ctor_get(x_3, 0); +lean_inc_ref(x_47); +lean_dec_ref(x_3); +x_48 = l_Lean_IR_EmitC_emitLit___redArg___closed__0; +x_49 = lean_string_append(x_46, x_48); +lean_inc_ref(x_47); +x_50 = l_Lean_IR_EmitC_quoteString(x_47); +x_51 = lean_string_append(x_49, x_50); +lean_dec_ref(x_50); +x_52 = l_Lean_IR_EmitC_emitSpreadArg___closed__1; +x_53 = lean_string_append(x_51, x_52); +x_54 = lean_string_utf8_byte_size(x_47); +x_55 = l_Nat_reprFast(x_54); +x_56 = lean_string_append(x_53, x_55); +lean_dec_ref(x_55); +x_57 = lean_string_append(x_56, x_52); +x_58 = lean_string_length(x_47); +lean_dec_ref(x_47); +x_59 = l_Nat_reprFast(x_58); +x_60 = lean_string_append(x_57, x_59); +lean_dec_ref(x_59); +x_61 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_62 = lean_string_append(x_60, x_61); +x_63 = l_Lean_IR_EmitC_emitLn___closed__0; +x_64 = lean_box(0); +x_65 = lean_string_append(x_62, x_63); +x_66 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +return x_66; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLit(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_IR_EmitC_emitLit___redArg(x_1, x_2, x_3, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLit___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_IR_EmitC_emitLit(x_1, x_2, x_3, x_4, x_5); +lean_dec_ref(x_4); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitVDecl(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +switch (lean_obj_tag(x_3)) { +case 0: +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_6 = lean_ctor_get(x_3, 0); +lean_inc_ref(x_6); +x_7 = lean_ctor_get(x_3, 1); +lean_inc_ref(x_7); +lean_dec_ref(x_3); +x_8 = l_Lean_IR_EmitC_emitCtor(x_1, x_2, x_6, x_7, x_4, x_5); +lean_dec_ref(x_7); +lean_dec(x_2); +return x_8; +} +case 1: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +lean_dec(x_2); +x_9 = lean_ctor_get(x_3, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_3, 1); +lean_inc(x_10); +lean_dec_ref(x_3); +x_11 = l_Lean_IR_EmitC_emitReset(x_1, x_9, x_10, x_4, x_5); +lean_dec_ref(x_4); +return x_11; +} +case 2: +{ +lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_2); +x_12 = lean_ctor_get(x_3, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_3, 1); +lean_inc_ref(x_13); +x_14 = lean_ctor_get_uint8(x_3, sizeof(void*)*3); +x_15 = lean_ctor_get(x_3, 2); +lean_inc_ref(x_15); +lean_dec_ref(x_3); +x_16 = l_Lean_IR_EmitC_emitReuse(x_1, x_12, x_13, x_14, x_15, x_4, x_5); +lean_dec_ref(x_4); +lean_dec_ref(x_15); +return x_16; +} +case 3: +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +lean_dec(x_2); +x_17 = lean_ctor_get(x_3, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_3, 1); +lean_inc(x_18); +x_19 = lean_ctor_get(x_3, 2); +lean_inc(x_19); +lean_dec_ref(x_3); +x_20 = l_Lean_IR_EmitC_emitProj(x_1, x_17, x_18, x_19, x_4, x_5); +lean_dec_ref(x_4); +return x_20; +} +case 4: +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +lean_dec(x_2); +x_21 = lean_ctor_get(x_3, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_3, 1); +lean_inc(x_22); +x_23 = lean_ctor_get(x_3, 2); +lean_inc(x_23); +lean_dec_ref(x_3); +x_24 = l_Lean_IR_EmitC_emitUProj(x_1, x_21, x_22, x_23, x_4, x_5); +return x_24; +} +case 5: +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_25 = lean_ctor_get(x_3, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_3, 1); +lean_inc(x_26); +x_27 = lean_ctor_get(x_3, 2); +lean_inc(x_27); +x_28 = lean_ctor_get(x_3, 3); +lean_inc(x_28); +lean_dec_ref(x_3); +x_29 = l_Lean_IR_EmitC_emitSProj(x_1, x_2, x_25, x_26, x_27, x_28, x_4, x_5); +lean_dec_ref(x_4); +lean_dec(x_2); +return x_29; +} +case 6: +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +lean_dec(x_2); +x_30 = lean_ctor_get(x_3, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_3, 1); +lean_inc_ref(x_31); +lean_dec_ref(x_3); +x_32 = l_Lean_IR_EmitC_emitFullApp(x_1, x_30, x_31, x_4, x_5); +return x_32; +} +case 7: +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +lean_dec(x_2); +x_33 = lean_ctor_get(x_3, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_3, 1); +lean_inc_ref(x_34); +lean_dec_ref(x_3); +x_35 = l_Lean_IR_EmitC_emitPartialApp(x_1, x_33, x_34, x_4, x_5); +lean_dec_ref(x_34); +return x_35; +} +case 8: +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +lean_dec(x_2); +x_36 = lean_ctor_get(x_3, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_3, 1); +lean_inc_ref(x_37); +lean_dec_ref(x_3); +x_38 = l_Lean_IR_EmitC_emitApp(x_1, x_36, x_37, x_4, x_5); +lean_dec_ref(x_4); +lean_dec_ref(x_37); +return x_38; +} +case 9: +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_3, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_3, 1); +lean_inc(x_40); +lean_dec_ref(x_3); +x_41 = l_Lean_IR_EmitC_emitBox(x_1, x_40, x_39, x_2, x_4, x_5); +lean_dec_ref(x_4); +lean_dec(x_2); +return x_41; +} +case 10: +{ +lean_object* x_42; lean_object* x_43; +x_42 = lean_ctor_get(x_3, 0); +lean_inc(x_42); +lean_dec_ref(x_3); +x_43 = l_Lean_IR_EmitC_emitUnbox(x_1, x_2, x_42, x_4, x_5); +lean_dec_ref(x_4); +lean_dec(x_2); +return x_43; +} +case 11: +{ +lean_object* x_44; lean_object* x_45; +lean_dec_ref(x_4); +x_44 = lean_ctor_get(x_3, 0); +lean_inc_ref(x_44); +lean_dec_ref(x_3); +x_45 = l_Lean_IR_EmitC_emitLit___redArg(x_1, x_2, x_44, x_5); +return x_45; +} +default: +{ +lean_object* x_46; lean_object* x_47; +lean_dec_ref(x_4); +lean_dec(x_2); +x_46 = lean_ctor_get(x_3, 0); +lean_inc(x_46); +lean_dec_ref(x_3); +x_47 = l_Lean_IR_EmitC_emitIsShared___redArg(x_1, x_46, x_5); +return x_47; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_isTailCall(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; +if (lean_obj_tag(x_2) == 6) +{ +if (lean_obj_tag(x_3) == 10) +{ +lean_object* x_11; +x_11 = lean_ctor_get(x_3, 0); +if (lean_obj_tag(x_11) == 0) +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_2); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_13 = lean_ctor_get(x_2, 0); +x_14 = lean_ctor_get(x_2, 1); +lean_dec(x_14); +x_15 = lean_ctor_get(x_11, 0); +x_16 = lean_ctor_get(x_4, 3); +x_17 = lean_name_eq(x_13, x_16); +lean_dec(x_13); +if (x_17 == 0) +{ +lean_object* x_18; +x_18 = lean_box(x_17); +lean_ctor_set_tag(x_2, 0); +lean_ctor_set(x_2, 1, x_5); +lean_ctor_set(x_2, 0, x_18); +return x_2; +} +else +{ +uint8_t x_19; lean_object* x_20; +x_19 = l_Lean_IR_instBEqVarId_beq(x_1, x_15); +x_20 = lean_box(x_19); +lean_ctor_set_tag(x_2, 0); +lean_ctor_set(x_2, 1, x_5); +lean_ctor_set(x_2, 0, x_20); +return x_2; +} +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_21 = lean_ctor_get(x_2, 0); +lean_inc(x_21); +lean_dec(x_2); +x_22 = lean_ctor_get(x_11, 0); +x_23 = lean_ctor_get(x_4, 3); +x_24 = lean_name_eq(x_21, x_23); +lean_dec(x_21); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_box(x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_5); +return x_26; +} +else +{ +uint8_t x_27; lean_object* x_28; lean_object* x_29; +x_27 = l_Lean_IR_instBEqVarId_beq(x_1, x_22); +x_28 = lean_box(x_27); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_5); +return x_29; +} +} +} +else +{ +lean_dec_ref(x_2); +x_6 = x_5; +goto block_10; +} +} +else +{ +lean_dec_ref(x_2); +x_6 = x_5; +goto block_10; +} +} +else +{ +lean_dec_ref(x_2); +x_6 = x_5; +goto block_10; +} +block_10: +{ +uint8_t x_7; lean_object* x_8; lean_object* x_9; +x_7 = 0; +x_8 = lean_box(x_7); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_6); +return x_9; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_isTailCall___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_IR_EmitC_isTailCall(x_1, x_2, x_3, x_4, x_5); +lean_dec_ref(x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_6; +} +} +LEAN_EXPORT uint8_t l_Lean_IR_EmitC_paramEqArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +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_2, 0); +x_4 = lean_ctor_get(x_1, 0); +x_5 = l_Lean_IR_instBEqVarId_beq(x_4, x_3); +return x_5; +} +else +{ +uint8_t x_6; +x_6 = 0; +return x_6; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_paramEqArg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_IR_EmitC_paramEqArg(x_1, x_2); +lean_dec(x_2); +lean_dec_ref(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__0___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_unsigned_to_nat(0u); +x_8 = lean_nat_dec_eq(x_6, x_7); +if (x_8 == 1) +{ +uint8_t x_9; +lean_dec(x_6); +lean_dec(x_2); +x_9 = 0; +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_10 = lean_nat_sub(x_5, x_6); +x_11 = lean_nat_add(x_1, x_10); +lean_dec(x_10); +lean_inc(x_2); +x_12 = lean_array_get_borrowed(x_2, x_3, x_11); +lean_dec(x_11); +x_13 = l_Lean_IR_EmitC_paramEqArg(x_4, x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_unsigned_to_nat(1u); +x_15 = lean_nat_sub(x_6, x_14); +lean_dec(x_6); +x_6 = x_15; +goto _start; +} +else +{ +lean_dec(x_6); +lean_dec(x_2); +return x_13; +} +} +} +} +LEAN_EXPORT uint8_t l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +uint8_t x_8; +x_8 = l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__0___redArg(x_1, x_2, x_3, x_4, x_5, x_6); +return x_8; +} +} +LEAN_EXPORT uint8_t l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00__private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1_spec__1___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_unsigned_to_nat(0u); +x_8 = lean_nat_dec_eq(x_6, x_7); +if (x_8 == 1) +{ +uint8_t x_9; +lean_dec(x_6); +lean_dec(x_3); +x_9 = 0; +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_10 = lean_nat_sub(x_5, x_6); +x_11 = lean_array_fget_borrowed(x_1, x_10); +x_12 = lean_unsigned_to_nat(1u); +x_13 = lean_nat_add(x_10, x_12); +lean_dec(x_10); +x_14 = lean_nat_sub(x_2, x_13); +lean_inc(x_14); +lean_inc(x_3); +x_15 = l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__0___redArg(x_13, x_3, x_4, x_11, x_14, x_14); +lean_dec(x_14); +lean_dec(x_13); +if (x_15 == 0) +{ +lean_object* x_16; +x_16 = lean_nat_sub(x_6, x_12); +lean_dec(x_6); +x_6 = x_16; +goto _start; +} +else +{ +lean_dec(x_6); +lean_dec(x_3); +return x_15; +} +} +} +} +LEAN_EXPORT uint8_t l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_unsigned_to_nat(0u); +x_8 = lean_nat_dec_eq(x_6, x_7); +if (x_8 == 1) +{ +uint8_t x_9; +lean_dec(x_2); +x_9 = 0; +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_10 = lean_nat_sub(x_5, x_6); +x_11 = lean_array_fget_borrowed(x_1, x_10); +x_12 = lean_unsigned_to_nat(1u); +x_13 = lean_nat_add(x_10, x_12); +lean_dec(x_10); +x_14 = lean_nat_sub(x_4, x_13); +lean_inc(x_14); +lean_inc(x_2); +x_15 = l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__0___redArg(x_13, x_2, x_3, x_11, x_14, x_14); +lean_dec(x_14); +lean_dec(x_13); +if (x_15 == 0) +{ +lean_object* x_16; uint8_t x_17; +x_16 = lean_nat_sub(x_6, x_12); +x_17 = l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00__private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1_spec__1___redArg(x_1, x_4, x_2, x_3, x_5, x_16); +return x_17; +} +else +{ +lean_dec(x_2); +return x_15; +} +} +} +} +LEAN_EXPORT uint8_t l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_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: +{ +uint8_t x_8; +x_8 = l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1___redArg(x_1, x_2, x_3, x_4, x_5, x_6); +return x_8; +} +} +LEAN_EXPORT uint8_t l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00__private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1_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: +{ +uint8_t x_8; +x_8 = l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00__private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1_spec__1___redArg(x_1, x_2, x_3, x_4, x_5, x_6); +return x_8; +} +} +static lean_object* _init_l_Lean_IR_EmitC_overwriteParam___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_IR_instInhabitedArg_default; +return x_1; +} +} +LEAN_EXPORT uint8_t l_Lean_IR_EmitC_overwriteParam(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = l_Lean_IR_EmitC_overwriteParam___closed__0; +x_4 = lean_array_get_size(x_1); +x_5 = l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1___redArg(x_1, x_3, x_2, x_4, x_4, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__0___boxed(lean_object* x_1, lean_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___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec_ref(x_3); +lean_dec(x_1); +x_9 = lean_box(x_8); +return x_9; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_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___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_1); +x_9 = lean_box(x_8); +return x_9; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00__private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1_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___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00__private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1_spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec(x_2); +lean_dec_ref(x_1); +x_9 = lean_box(x_8); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_overwriteParam___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_IR_EmitC_overwriteParam(x_1, x_2); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; lean_object* x_8; +x_7 = l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__0___redArg(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec_ref(x_3); +lean_dec(x_1); +x_8 = lean_box(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; lean_object* x_8; +x_7 = l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1___redArg(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_1); +x_8 = lean_box(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00__private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1_spec__1___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; lean_object* x_8; +x_7 = l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00__private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1_spec__1___redArg(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec(x_2); +lean_dec_ref(x_1); +x_8 = lean_box(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__0___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_nat_dec_eq(x_4, x_6); +if (x_7 == 1) +{ +lean_object* x_8; lean_object* x_9; +lean_dec(x_4); +x_8 = lean_box(0); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_5); +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_sub(x_4, x_10); +lean_dec(x_4); +x_12 = lean_nat_sub(x_3, x_11); +x_13 = lean_nat_sub(x_12, x_10); +lean_dec(x_12); +x_14 = lean_array_fget_borrowed(x_1, x_13); +x_15 = l_Lean_IR_EmitC_overwriteParam___closed__0; +x_16 = lean_array_get_borrowed(x_15, x_2, x_13); +lean_dec(x_13); +x_17 = l_Lean_IR_EmitC_paramEqArg(x_14, x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_18 = lean_ctor_get(x_14, 0); +x_19 = l_Lean_IR_EmitC_argToCString___closed__0; +lean_inc(x_18); +x_20 = l_Nat_reprFast(x_18); +x_21 = lean_string_append(x_19, x_20); +lean_dec_ref(x_20); +x_22 = lean_string_append(x_5, x_21); +lean_dec_ref(x_21); +x_23 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__1; +x_24 = lean_string_append(x_22, x_23); +lean_inc(x_16); +x_25 = l_Lean_IR_EmitC_emitArg___redArg(x_16, x_24); +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec_ref(x_25); +x_27 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_28 = lean_string_append(x_26, x_27); +x_29 = l_Lean_IR_EmitC_emitLn___closed__0; +x_30 = lean_string_append(x_28, x_29); +x_4 = x_11; +x_5 = x_30; +goto _start; +} +else +{ +x_4 = x_11; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__0___redArg(x_1, x_2, x_3, x_4, x_7); +return x_8; +} +} +static lean_object* _init_l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__1___redArg___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" _tmp_", 6, 6); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__1___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_unsigned_to_nat(0u); +x_8 = lean_nat_dec_eq(x_4, x_7); +if (x_8 == 1) +{ +lean_object* x_9; lean_object* x_10; +lean_dec(x_4); +x_9 = lean_box(0); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_6); +return x_10; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_sub(x_4, x_11); +lean_dec(x_4); +x_13 = lean_nat_sub(x_3, x_12); +x_14 = lean_nat_sub(x_13, x_11); +lean_dec(x_13); +x_15 = lean_array_fget_borrowed(x_1, x_14); +x_16 = l_Lean_IR_EmitC_overwriteParam___closed__0; +x_17 = lean_array_get_borrowed(x_16, x_2, x_14); +x_18 = l_Lean_IR_EmitC_paramEqArg(x_15, x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_19 = lean_ctor_get(x_15, 1); +x_20 = l_Lean_IR_EmitC_toCType(x_19, x_5, x_6); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec_ref(x_20); +x_23 = lean_string_append(x_22, x_21); +lean_dec(x_21); +x_24 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__1___redArg___closed__0; +x_25 = lean_string_append(x_23, x_24); +x_26 = l_Nat_reprFast(x_14); +x_27 = lean_string_append(x_25, x_26); +lean_dec_ref(x_26); +x_28 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__1; +x_29 = lean_string_append(x_27, x_28); +lean_inc(x_17); +x_30 = l_Lean_IR_EmitC_emitArg___redArg(x_17, x_29); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec_ref(x_30); +x_32 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_33 = lean_string_append(x_31, x_32); +x_34 = l_Lean_IR_EmitC_emitLn___closed__0; +x_35 = lean_string_append(x_33, x_34); +x_4 = x_12; +x_6 = x_35; +goto _start; +} +else +{ +lean_dec(x_14); +x_4 = x_12; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_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; +x_8 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__1___redArg(x_1, x_2, x_3, x_4, x_6, x_7); +return x_8; +} +} +static lean_object* _init_l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__2___redArg___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" = _tmp_", 8, 8); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__2___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_nat_dec_eq(x_4, x_6); +if (x_7 == 1) +{ +lean_object* x_8; lean_object* x_9; +lean_dec(x_4); +x_8 = lean_box(0); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_5); +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_sub(x_4, x_10); +lean_dec(x_4); +x_12 = lean_nat_sub(x_3, x_11); +x_13 = lean_nat_sub(x_12, x_10); +lean_dec(x_12); +x_14 = lean_array_fget_borrowed(x_1, x_13); +x_15 = l_Lean_IR_EmitC_overwriteParam___closed__0; +x_16 = lean_array_get_borrowed(x_15, x_2, x_13); +x_17 = l_Lean_IR_EmitC_paramEqArg(x_14, x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_18 = lean_ctor_get(x_14, 0); +x_19 = l_Lean_IR_EmitC_argToCString___closed__0; +lean_inc(x_18); +x_20 = l_Nat_reprFast(x_18); +x_21 = lean_string_append(x_19, x_20); +lean_dec_ref(x_20); +x_22 = lean_string_append(x_5, x_21); +lean_dec_ref(x_21); +x_23 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__2___redArg___closed__0; +x_24 = lean_string_append(x_22, x_23); +x_25 = l_Nat_reprFast(x_13); +x_26 = lean_string_append(x_24, x_25); +lean_dec_ref(x_25); +x_27 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_28 = lean_string_append(x_26, x_27); +x_29 = l_Lean_IR_EmitC_emitLn___closed__0; +x_30 = lean_string_append(x_28, x_29); +x_4 = x_11; +x_5 = x_30; +goto _start; +} +else +{ +lean_dec(x_13); +x_4 = x_11; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_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; +x_8 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__2___redArg(x_1, x_2, x_3, x_4, x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitTailCall_spec__3(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; uint8_t x_10; +x_10 = lean_usize_dec_eq(x_2, x_3); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_array_uget(x_1, x_2); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +lean_dec(x_12); +x_14 = l_Lean_IR_IRType_isVoid(x_13); +lean_dec(x_13); +if (x_14 == 0) +{ +lean_object* x_15; +x_15 = lean_array_push(x_4, x_11); +x_5 = x_15; +goto block_9; +} +else +{ +lean_dec(x_11); +x_5 = x_4; +goto block_9; +} +} +else +{ +return x_4; +} +block_9: +{ +size_t x_6; size_t x_7; +x_6 = 1; +x_7 = lean_usize_add(x_2, x_6); +x_2 = x_7; +x_4 = x_5; +goto _start; +} +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitTailCall___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("goto _start;", 12, 12); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitTailCall___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid tail call", 17, 17); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitTailCall___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitTailCall___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("bug at emitTailCall", 19, 19); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitTailCall(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_12; +if (lean_obj_tag(x_1) == 6) +{ +uint8_t x_33; +x_33 = !lean_is_exclusive(x_1); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; +x_34 = lean_ctor_get(x_1, 1); +x_35 = lean_ctor_get(x_1, 0); +lean_dec(x_35); +x_36 = lean_ctor_get(x_2, 4); +x_37 = lean_array_get_size(x_36); +x_38 = lean_array_get_size(x_34); +x_39 = lean_nat_dec_eq(x_37, x_38); +if (x_39 == 0) +{ +lean_object* x_40; +lean_dec_ref(x_34); +x_40 = l_Lean_IR_EmitC_emitTailCall___closed__1; +lean_ctor_set_tag(x_1, 1); +lean_ctor_set(x_1, 1, x_3); +lean_ctor_set(x_1, 0, x_40); +return x_1; +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; +lean_free_object(x_1); +x_41 = l_Array_zip___redArg(x_36, x_34); +lean_dec_ref(x_34); +x_42 = lean_unsigned_to_nat(0u); +x_43 = lean_array_get_size(x_41); +x_44 = l_Lean_IR_EmitC_emitTailCall___closed__2; +x_45 = lean_nat_dec_lt(x_42, x_43); +if (x_45 == 0) +{ +lean_dec_ref(x_41); +x_12 = x_44; +goto block_32; +} +else +{ +uint8_t x_46; +x_46 = lean_nat_dec_le(x_43, x_43); +if (x_46 == 0) +{ +if (x_45 == 0) +{ +lean_dec_ref(x_41); +x_12 = x_44; +goto block_32; +} +else +{ +size_t x_47; size_t x_48; lean_object* x_49; +x_47 = 0; +x_48 = lean_usize_of_nat(x_43); +x_49 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitTailCall_spec__3(x_41, x_47, x_48, x_44); +lean_dec_ref(x_41); +x_12 = x_49; +goto block_32; +} +} +else +{ +size_t x_50; size_t x_51; lean_object* x_52; +x_50 = 0; +x_51 = lean_usize_of_nat(x_43); +x_52 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitTailCall_spec__3(x_41, x_50, x_51, x_44); +lean_dec_ref(x_41); +x_12 = x_52; +goto block_32; +} +} +} +} +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; +x_53 = lean_ctor_get(x_1, 1); +lean_inc(x_53); +lean_dec(x_1); +x_54 = lean_ctor_get(x_2, 4); +x_55 = lean_array_get_size(x_54); +x_56 = lean_array_get_size(x_53); +x_57 = lean_nat_dec_eq(x_55, x_56); +if (x_57 == 0) +{ +lean_object* x_58; lean_object* x_59; +lean_dec_ref(x_53); +x_58 = l_Lean_IR_EmitC_emitTailCall___closed__1; +x_59 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_3); +return x_59; +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; +x_60 = l_Array_zip___redArg(x_54, x_53); +lean_dec_ref(x_53); +x_61 = lean_unsigned_to_nat(0u); +x_62 = lean_array_get_size(x_60); +x_63 = l_Lean_IR_EmitC_emitTailCall___closed__2; +x_64 = lean_nat_dec_lt(x_61, x_62); +if (x_64 == 0) +{ +lean_dec_ref(x_60); +x_12 = x_63; +goto block_32; +} +else +{ +uint8_t x_65; +x_65 = lean_nat_dec_le(x_62, x_62); +if (x_65 == 0) +{ +if (x_64 == 0) +{ +lean_dec_ref(x_60); +x_12 = x_63; +goto block_32; +} +else +{ +size_t x_66; size_t x_67; lean_object* x_68; +x_66 = 0; +x_67 = lean_usize_of_nat(x_62); +x_68 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitTailCall_spec__3(x_60, x_66, x_67, x_63); +lean_dec_ref(x_60); +x_12 = x_68; +goto block_32; +} +} +else +{ +size_t x_69; size_t x_70; lean_object* x_71; +x_69 = 0; +x_70 = lean_usize_of_nat(x_62); +x_71 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitTailCall_spec__3(x_60, x_69, x_70, x_63); +lean_dec_ref(x_60); +x_12 = x_71; +goto block_32; +} +} +} +} +} +else +{ +lean_object* x_72; lean_object* x_73; +lean_dec_ref(x_1); +x_72 = l_Lean_IR_EmitC_emitTailCall___closed__3; +x_73 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_3); +return x_73; +} +block_11: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_5 = l_Lean_IR_EmitC_emitTailCall___closed__0; +x_6 = lean_string_append(x_4, x_5); +x_7 = l_Lean_IR_EmitC_emitLn___closed__0; +x_8 = lean_box(0); +x_9 = lean_string_append(x_6, x_7); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_8); +lean_ctor_set(x_10, 1, x_9); +return x_10; +} +block_32: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = l_Array_unzip___redArg(x_12); +lean_dec_ref(x_12); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec_ref(x_13); +x_16 = l_Lean_IR_EmitC_overwriteParam(x_14, x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_array_get_size(x_14); +x_18 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__0___redArg(x_14, x_15, x_17, x_17, x_3); +lean_dec(x_15); +lean_dec(x_14); +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +lean_dec_ref(x_18); +x_4 = x_19; +goto block_11; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_20 = l_Lean_IR_EmitC_emitSpreadValue___closed__0; +x_21 = lean_string_append(x_3, x_20); +x_22 = l_Lean_IR_EmitC_emitLn___closed__0; +x_23 = lean_string_append(x_21, x_22); +x_24 = lean_array_get_size(x_14); +x_25 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__1___redArg(x_14, x_15, x_24, x_24, x_2, x_23); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec_ref(x_25); +x_27 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__2___redArg(x_14, x_15, x_24, x_24, x_26); +lean_dec(x_15); +lean_dec(x_14); +x_28 = lean_ctor_get(x_27, 1); +lean_inc(x_28); +lean_dec_ref(x_27); +x_29 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; +x_30 = lean_string_append(x_28, x_29); +x_31 = lean_string_append(x_30, x_22); +x_4 = x_31; +goto block_11; +} +else +{ +lean_dec(x_15); +lean_dec(x_14); +return x_25; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec_ref(x_6); +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec_ref(x_6); +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_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: +{ +lean_object* x_8; +x_8 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec_ref(x_6); +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitTailCall_spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; lean_object* x_7; +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitTailCall_spec__3(x_1, x_5, x_6, x_4); +lean_dec_ref(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__2___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__2___redArg(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__0___redArg(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__1___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__1___redArg(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec_ref(x_5); +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitTailCall___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_IR_EmitC_emitTailCall(x_1, x_2, x_3); +lean_dec_ref(x_2); +return x_4; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitIf___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("if (", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitIf___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" == ", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitIf___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("else", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitCase___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("switch (", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitCase___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(") {", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitCase_spec__4___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("case ", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitJPs___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(":", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitCase_spec__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("default: ", 9, 9); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitCase_spec__4(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; uint8_t x_14; +x_14 = lean_usize_dec_eq(x_2, x_3); +if (x_14 == 0) +{ +lean_object* x_15; +x_15 = lean_array_uget(x_1, x_2); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_16 = lean_ctor_get(x_15, 0); +lean_inc_ref(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec_ref(x_15); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec_ref(x_16); +x_19 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitCase_spec__4___closed__0; +x_20 = lean_string_append(x_6, x_19); +x_21 = l_Nat_reprFast(x_18); +x_22 = lean_string_append(x_20, x_21); +lean_dec_ref(x_21); +x_23 = l_Lean_IR_EmitC_emitJPs___closed__0; +x_24 = lean_string_append(x_22, x_23); +x_25 = l_Lean_IR_EmitC_emitLn___closed__0; +x_26 = lean_string_append(x_24, x_25); +lean_inc_ref(x_5); +x_27 = l_Lean_IR_EmitC_emitFnBody(x_17, x_5, x_26); +x_7 = x_27; +goto block_13; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_28 = lean_ctor_get(x_15, 0); +lean_inc(x_28); +lean_dec_ref(x_15); +x_29 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitCase_spec__4___closed__1; +x_30 = lean_string_append(x_6, x_29); +x_31 = l_Lean_IR_EmitC_emitLn___closed__0; +x_32 = lean_string_append(x_30, x_31); +lean_inc_ref(x_5); +x_33 = l_Lean_IR_EmitC_emitFnBody(x_28, x_5, x_32); +x_7 = x_33; +goto block_13; +} +} +else +{ +lean_object* x_34; +lean_dec_ref(x_5); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_4); +lean_ctor_set(x_34, 1, x_6); +return x_34; +} +block_13: +{ +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_8; lean_object* x_9; size_t x_10; size_t x_11; +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +lean_dec_ref(x_7); +x_10 = 1; +x_11 = lean_usize_add(x_2, x_10); +x_2 = x_11; +x_4 = x_8; +x_6 = x_9; +goto _start; +} +else +{ +lean_dec_ref(x_5); +return x_7; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCase(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; lean_object* x_14; lean_object* x_17; +x_17 = l_Lean_IR_EmitC_isIf(x_3); +if (lean_obj_tag(x_17) == 1) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +lean_dec_ref(x_3); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +lean_dec_ref(x_17); +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 0); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_ctor_get(x_19, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_19, 1); +lean_inc(x_22); +lean_dec(x_19); +x_23 = l_Lean_IR_EmitC_emitIf(x_1, x_2, x_20, x_21, x_22, x_4, x_5); +return x_23; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_dec(x_17); +x_24 = l_Lean_IR_EmitC_emitCase___closed__0; +x_25 = lean_string_append(x_5, x_24); +x_26 = l_Lean_IR_EmitC_emitTag___redArg(x_1, x_2, x_25); +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +lean_dec_ref(x_26); +x_28 = l_Lean_IR_EmitC_emitCase___closed__1; +x_29 = lean_string_append(x_27, x_28); +x_30 = l_Lean_IR_EmitC_emitLn___closed__0; +x_31 = lean_string_append(x_29, x_30); +x_32 = l_Lean_IR_ensureHasDefault(x_3); +x_33 = lean_unsigned_to_nat(0u); +x_34 = lean_array_get_size(x_32); +x_35 = lean_nat_dec_lt(x_33, x_34); +if (x_35 == 0) +{ +lean_dec_ref(x_32); +lean_dec_ref(x_4); +x_6 = x_31; +goto block_13; +} +else +{ +lean_object* x_36; uint8_t x_37; +x_36 = lean_box(0); +x_37 = lean_nat_dec_le(x_34, x_34); +if (x_37 == 0) +{ +if (x_35 == 0) +{ +lean_dec_ref(x_32); +lean_dec_ref(x_4); +x_6 = x_31; +goto block_13; +} +else +{ +size_t x_38; size_t x_39; lean_object* x_40; +x_38 = 0; +x_39 = lean_usize_of_nat(x_34); +x_40 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitCase_spec__4(x_32, x_38, x_39, x_36, x_4, x_31); +lean_dec_ref(x_32); +x_14 = x_40; +goto block_16; +} +} +else +{ +size_t x_41; size_t x_42; lean_object* x_43; +x_41 = 0; +x_42 = lean_usize_of_nat(x_34); +x_43 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitCase_spec__4(x_32, x_41, x_42, x_36, x_4, x_31); +lean_dec_ref(x_32); +x_14 = x_43; +goto block_16; +} +} +} +else +{ +lean_dec_ref(x_4); +lean_dec_ref(x_3); +return x_26; +} +} +block_13: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_7 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; +x_8 = lean_string_append(x_6, x_7); +x_9 = l_Lean_IR_EmitC_emitLn___closed__0; +x_10 = lean_box(0); +x_11 = lean_string_append(x_8, x_9); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set(x_12, 1, x_11); +return x_12; +} +block_16: +{ +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +lean_dec_ref(x_14); +x_6 = x_15; +goto block_13; +} +else +{ +return x_14; +} +} +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitBlock___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("return ", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitBlock___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_internal_panic_unreachable();", 34, 34); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitBlock(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 0: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_1, 1); +lean_inc(x_5); +x_6 = lean_ctor_get(x_1, 2); +lean_inc_ref(x_6); +x_7 = lean_ctor_get(x_1, 3); +lean_inc(x_7); +x_8 = lean_ctor_get(x_2, 3); +x_9 = l_Lean_IR_isTailCallTo(x_8, x_1); +lean_dec_ref(x_1); +if (x_9 == 0) +{ +lean_object* x_10; +lean_inc_ref(x_2); +x_10 = l_Lean_IR_EmitC_emitVDecl(x_4, x_5, x_6, x_2, x_3); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec_ref(x_10); +x_1 = x_7; +x_3 = x_11; +goto _start; +} +else +{ +lean_dec(x_7); +lean_dec_ref(x_2); +return x_10; +} +} +else +{ +lean_object* x_13; +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +x_13 = l_Lean_IR_EmitC_emitTailCall(x_6, x_2, x_3); +lean_dec_ref(x_2); +return x_13; +} +} +case 1: +{ +lean_object* x_14; +x_14 = lean_ctor_get(x_1, 3); +lean_inc(x_14); +lean_dec_ref(x_1); +x_1 = x_14; +goto _start; +} +case 2: +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_16 = lean_ctor_get(x_1, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_1, 1); +lean_inc(x_17); +x_18 = lean_ctor_get(x_1, 2); +lean_inc(x_18); +x_19 = lean_ctor_get(x_1, 3); +lean_inc(x_19); +lean_dec_ref(x_1); +x_20 = l_Lean_IR_EmitC_emitSet___redArg(x_16, x_17, x_18, x_3); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; +x_21 = lean_ctor_get(x_20, 1); +lean_inc(x_21); +lean_dec_ref(x_20); +x_1 = x_19; +x_3 = x_21; +goto _start; +} +else +{ +lean_dec(x_19); +lean_dec_ref(x_2); +return x_20; +} +} +case 3: +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_23 = lean_ctor_get(x_1, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_1, 1); +lean_inc(x_24); +x_25 = lean_ctor_get(x_1, 2); +lean_inc(x_25); +lean_dec_ref(x_1); +x_26 = l_Lean_IR_EmitC_emitSetTag___redArg(x_23, x_24, x_3); +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_27; +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +lean_dec_ref(x_26); +x_1 = x_25; +x_3 = x_27; +goto _start; +} +else +{ +lean_dec(x_25); +lean_dec_ref(x_2); +return x_26; +} +} +case 4: +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_29 = lean_ctor_get(x_1, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_1, 1); +lean_inc(x_30); +x_31 = lean_ctor_get(x_1, 2); +lean_inc(x_31); +x_32 = lean_ctor_get(x_1, 3); +lean_inc(x_32); +x_33 = lean_ctor_get(x_1, 4); +lean_inc(x_33); +lean_dec_ref(x_1); +lean_inc_ref(x_2); +x_34 = l_Lean_IR_EmitC_emitUSet(x_29, x_30, x_31, x_32, x_2, x_3); +if (lean_obj_tag(x_34) == 0) +{ +lean_object* x_35; +x_35 = lean_ctor_get(x_34, 1); +lean_inc(x_35); +lean_dec_ref(x_34); +x_1 = x_33; +x_3 = x_35; +goto _start; +} +else +{ +lean_dec(x_33); +lean_dec_ref(x_2); +return x_34; +} +} +case 5: +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_37 = lean_ctor_get(x_1, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_1, 1); +lean_inc(x_38); +x_39 = lean_ctor_get(x_1, 2); +lean_inc(x_39); +x_40 = lean_ctor_get(x_1, 3); +lean_inc(x_40); +x_41 = lean_ctor_get(x_1, 4); +lean_inc(x_41); +x_42 = lean_ctor_get(x_1, 5); +lean_inc(x_42); +x_43 = lean_ctor_get(x_1, 6); +lean_inc(x_43); +lean_dec_ref(x_1); +x_44 = l_Lean_IR_EmitC_emitSSet(x_37, x_38, x_39, x_40, x_41, x_42, x_2, x_3); +lean_dec(x_42); +if (lean_obj_tag(x_44) == 0) +{ +lean_object* x_45; +x_45 = lean_ctor_get(x_44, 1); +lean_inc(x_45); +lean_dec_ref(x_44); +x_1 = x_43; +x_3 = x_45; +goto _start; +} +else +{ +lean_dec(x_43); +lean_dec_ref(x_2); +return x_44; +} +} +case 6: +{ +uint8_t x_47; +x_47 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 1); +if (x_47 == 0) +{ +lean_object* x_48; lean_object* x_49; uint8_t x_50; lean_object* x_51; lean_object* x_52; +x_48 = lean_ctor_get(x_1, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_1, 1); +lean_inc(x_49); +x_50 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); +x_51 = lean_ctor_get(x_1, 2); +lean_inc(x_51); +lean_dec_ref(x_1); +x_52 = l_Lean_IR_EmitC_emitInc(x_48, x_49, x_50, x_2, x_3); +if (lean_obj_tag(x_52) == 0) +{ +lean_object* x_53; +x_53 = lean_ctor_get(x_52, 1); +lean_inc(x_53); +lean_dec_ref(x_52); +x_1 = x_51; +x_3 = x_53; +goto _start; +} +else +{ +lean_dec(x_51); +lean_dec_ref(x_2); +return x_52; +} +} +else +{ +lean_object* x_55; +x_55 = lean_ctor_get(x_1, 2); +lean_inc(x_55); +lean_dec_ref(x_1); +x_1 = x_55; +goto _start; +} +} +case 7: +{ +uint8_t x_57; +x_57 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 1); +if (x_57 == 0) +{ +lean_object* x_58; lean_object* x_59; uint8_t x_60; lean_object* x_61; lean_object* x_62; +x_58 = lean_ctor_get(x_1, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_1, 1); +lean_inc(x_59); +x_60 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); +x_61 = lean_ctor_get(x_1, 2); +lean_inc(x_61); +lean_dec_ref(x_1); +x_62 = l_Lean_IR_EmitC_emitDec(x_58, x_59, x_60, x_2, x_3); +if (lean_obj_tag(x_62) == 0) +{ +lean_object* x_63; +x_63 = lean_ctor_get(x_62, 1); +lean_inc(x_63); +lean_dec_ref(x_62); +x_1 = x_61; +x_3 = x_63; +goto _start; +} +else +{ +lean_dec(x_61); +lean_dec_ref(x_2); +return x_62; +} +} +else +{ +lean_object* x_65; +x_65 = lean_ctor_get(x_1, 2); +lean_inc(x_65); +lean_dec_ref(x_1); +x_1 = x_65; +goto _start; +} +} +case 8: +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_1, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_1, 1); +lean_inc(x_68); +lean_dec_ref(x_1); +x_69 = l_Lean_IR_EmitC_emitDel___redArg(x_67, x_3); +if (lean_obj_tag(x_69) == 0) +{ +lean_object* x_70; +x_70 = lean_ctor_get(x_69, 1); +lean_inc(x_70); +lean_dec_ref(x_69); +x_1 = x_68; +x_3 = x_70; +goto _start; +} +else +{ +lean_dec(x_68); +lean_dec_ref(x_2); +return x_69; +} +} +case 9: +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_72 = lean_ctor_get(x_1, 1); +lean_inc(x_72); +x_73 = lean_ctor_get(x_1, 2); +lean_inc(x_73); +x_74 = lean_ctor_get(x_1, 3); +lean_inc_ref(x_74); +lean_dec_ref(x_1); +x_75 = l_Lean_IR_EmitC_emitCase(x_72, x_73, x_74, x_2, x_3); +lean_dec(x_73); +return x_75; +} +case 10: +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +lean_dec_ref(x_2); +x_76 = lean_ctor_get(x_1, 0); +lean_inc(x_76); +lean_dec_ref(x_1); +x_77 = l_Lean_IR_EmitC_emitBlock___closed__0; +x_78 = lean_string_append(x_3, x_77); +x_79 = l_Lean_IR_EmitC_emitArg___redArg(x_76, x_78); +if (lean_obj_tag(x_79) == 0) +{ +uint8_t x_80; +x_80 = !lean_is_exclusive(x_79); +if (x_80 == 0) +{ +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_81 = lean_ctor_get(x_79, 1); +x_82 = lean_ctor_get(x_79, 0); +lean_dec(x_82); +x_83 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_84 = lean_string_append(x_81, x_83); +x_85 = l_Lean_IR_EmitC_emitLn___closed__0; +x_86 = lean_box(0); +x_87 = lean_string_append(x_84, x_85); +lean_ctor_set(x_79, 1, x_87); +lean_ctor_set(x_79, 0, x_86); +return x_79; +} +else +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_88 = lean_ctor_get(x_79, 1); +lean_inc(x_88); +lean_dec(x_79); +x_89 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_90 = lean_string_append(x_88, x_89); +x_91 = l_Lean_IR_EmitC_emitLn___closed__0; +x_92 = lean_box(0); +x_93 = lean_string_append(x_90, x_91); +x_94 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_94, 0, x_92); +lean_ctor_set(x_94, 1, x_93); +return x_94; +} +} +else +{ +return x_79; +} +} +case 11: +{ +lean_object* x_95; lean_object* x_96; lean_object* x_97; +x_95 = lean_ctor_get(x_1, 0); +lean_inc(x_95); +x_96 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_96); +lean_dec_ref(x_1); +x_97 = l_Lean_IR_EmitC_emitJmp(x_95, x_96, x_2, x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_96); +return x_97; +} +default: +{ +lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; +lean_dec_ref(x_2); +x_98 = l_Lean_IR_EmitC_emitBlock___closed__1; +x_99 = lean_string_append(x_3, x_98); +x_100 = l_Lean_IR_EmitC_emitLn___closed__0; +x_101 = lean_box(0); +x_102 = lean_string_append(x_99, x_100); +x_103 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_103, 0, x_101); +lean_ctor_set(x_103, 1, x_102); +return x_103; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitJPs(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_1) == 1) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_1, 2); +lean_inc(x_5); +x_6 = lean_ctor_get(x_1, 3); +lean_inc(x_6); +lean_dec_ref(x_1); +x_7 = l_Lean_IR_EmitC_emitJmp___closed__2; +x_8 = l_Nat_reprFast(x_4); +x_9 = lean_string_append(x_7, x_8); +lean_dec_ref(x_8); +x_10 = lean_string_append(x_3, x_9); +lean_dec_ref(x_9); +x_11 = l_Lean_IR_EmitC_emitJPs___closed__0; +x_12 = lean_string_append(x_10, x_11); +x_13 = l_Lean_IR_EmitC_emitLn___closed__0; +x_14 = lean_string_append(x_12, x_13); +lean_inc_ref(x_2); +x_15 = l_Lean_IR_EmitC_emitFnBody(x_5, x_2, x_14); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec_ref(x_15); +x_1 = x_6; +x_3 = x_16; +goto _start; +} +else +{ +lean_dec(x_6); +lean_dec_ref(x_2); +return x_15; +} +} +else +{ +uint8_t x_18; +x_18 = l_Lean_IR_FnBody_isTerminal(x_1); +if (x_18 == 0) +{ +lean_object* x_19; +x_19 = l_Lean_IR_FnBody_body(x_1); +lean_dec(x_1); +x_1 = x_19; +goto _start; +} +else +{ +lean_object* x_21; lean_object* x_22; +lean_dec_ref(x_2); +lean_dec(x_1); +x_21 = lean_box(0); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_3); +return x_22; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFnBody(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; +x_25 = l_Lean_IR_EmitC_emitSpreadValue___closed__0; +x_26 = lean_string_append(x_3, x_25); +x_27 = l_Lean_IR_EmitC_emitLn___closed__0; +x_28 = lean_string_append(x_26, x_27); +x_29 = 0; +lean_inc(x_1); +x_30 = l_Lean_IR_EmitC_declareVars(x_1, x_29, x_2, x_28); +if (lean_obj_tag(x_30) == 0) +{ +lean_object* x_31; uint8_t x_32; +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_unbox(x_31); +lean_dec(x_31); +if (x_32 == 0) +{ +lean_object* x_33; +x_33 = lean_ctor_get(x_30, 1); +lean_inc(x_33); +lean_dec_ref(x_30); +x_4 = x_2; +x_5 = x_33; +goto block_24; +} +else +{ +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_30, 1); +lean_inc(x_34); +lean_dec_ref(x_30); +x_35 = lean_string_append(x_34, x_27); +x_4 = x_2; +x_5 = x_35; +goto block_24; +} +} +else +{ +uint8_t x_36; +lean_dec_ref(x_2); +lean_dec(x_1); +x_36 = !lean_is_exclusive(x_30); +if (x_36 == 0) +{ +return x_30; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_30, 0); +x_38 = lean_ctor_get(x_30, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_30); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; +} +} +block_24: +{ +lean_object* x_6; +lean_inc_ref(x_4); +lean_inc(x_1); +x_6 = l_Lean_IR_EmitC_emitBlock(x_1, x_4, x_5); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_6, 1); +lean_inc(x_7); +lean_dec_ref(x_6); +x_8 = l_Lean_IR_EmitC_emitJPs(x_1, x_4, x_7); +if (lean_obj_tag(x_8) == 0) +{ +uint8_t x_9; +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_10 = lean_ctor_get(x_8, 1); +x_11 = lean_ctor_get(x_8, 0); +lean_dec(x_11); +x_12 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; +x_13 = lean_string_append(x_10, x_12); +x_14 = l_Lean_IR_EmitC_emitLn___closed__0; +x_15 = lean_box(0); +x_16 = lean_string_append(x_13, x_14); +lean_ctor_set(x_8, 1, x_16); +lean_ctor_set(x_8, 0, x_15); +return x_8; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_17 = lean_ctor_get(x_8, 1); +lean_inc(x_17); +lean_dec(x_8); +x_18 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; +x_19 = lean_string_append(x_17, x_18); +x_20 = l_Lean_IR_EmitC_emitLn___closed__0; +x_21 = lean_box(0); +x_22 = lean_string_append(x_19, x_20); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +else +{ +return x_8; +} +} +else +{ +lean_dec_ref(x_4); +lean_dec(x_1); +return x_6; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitIf(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = l_Lean_IR_EmitC_emitIf___closed__0; +x_9 = lean_string_append(x_7, x_8); +x_10 = l_Lean_IR_EmitC_emitTag___redArg(x_1, x_2, x_9); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec_ref(x_10); +x_12 = l_Lean_IR_EmitC_emitIf___closed__1; +x_13 = lean_string_append(x_11, x_12); +x_14 = l_Nat_reprFast(x_3); +x_15 = lean_string_append(x_13, x_14); +lean_dec_ref(x_14); +x_16 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__1; +x_17 = lean_string_append(x_15, x_16); +x_18 = l_Lean_IR_EmitC_emitLn___closed__0; +x_19 = lean_string_append(x_17, x_18); +lean_inc_ref(x_6); +x_20 = l_Lean_IR_EmitC_emitFnBody(x_4, x_6, x_19); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_21 = lean_ctor_get(x_20, 1); +lean_inc(x_21); +lean_dec_ref(x_20); +x_22 = l_Lean_IR_EmitC_emitIf___closed__2; +x_23 = lean_string_append(x_21, x_22); +x_24 = lean_string_append(x_23, x_18); +x_25 = l_Lean_IR_EmitC_emitFnBody(x_5, x_6, x_24); +return x_25; +} +else +{ +lean_dec_ref(x_6); +lean_dec(x_5); +return x_20; +} +} +else +{ +lean_dec_ref(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_10; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitIf___boxed(lean_object* x_1, lean_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_IR_EmitC_emitIf(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitCase_spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_8 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_9 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitCase_spec__4(x_1, x_7, x_8, x_4, x_5, x_6); +lean_dec_ref(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCase___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_IR_EmitC_emitCase(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_2); +return x_6; +} +} +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__0___redArg___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" c", 2, 2); +return x_1; +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__0___redArg(lean_object* x_1, lean_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_nat_dec_lt(x_4, x_1); +if (x_8 == 0) +{ +lean_object* x_9; +lean_dec(x_4); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_5); +lean_ctor_set(x_9, 1, x_7); +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_array_fget_borrowed(x_2, x_4); +x_11 = l_Lean_IR_EmitC_toCType(x_10, x_6, x_7); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; 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_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_ref(x_11); +x_14 = lean_string_append(x_13, x_12); +lean_dec(x_12); +x_15 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__0___redArg___closed__0; +x_16 = lean_string_append(x_14, x_15); +lean_inc(x_4); +x_17 = l_Nat_reprFast(x_4); +x_18 = lean_string_append(x_16, x_17); +lean_dec_ref(x_17); +x_19 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_20 = lean_string_append(x_18, x_19); +x_21 = l_Lean_IR_EmitC_emitLn___closed__0; +x_22 = lean_string_append(x_20, x_21); +x_23 = lean_unsigned_to_nat(1u); +x_24 = lean_nat_add(x_4, x_23); +lean_dec(x_4); +{ +lean_object* _tmp_3 = x_24; +lean_object* _tmp_4 = x_3; +lean_object* _tmp_6 = x_22; +x_4 = _tmp_3; +x_5 = _tmp_4; +x_7 = _tmp_6; +} +goto _start; +} +else +{ +uint8_t x_26; +lean_dec(x_4); +x_26 = !lean_is_exclusive(x_11); +if (x_26 == 0) +{ +return x_11; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_11, 0); +x_28 = lean_ctor_get(x_11, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_11); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; +} +} +} +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__0___redArg(x_1, x_2, x_3, x_6, x_7, x_9, x_10); +return x_11; +} +} +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__1___redArg___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" i", 2, 2); +return x_1; +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__1___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; uint8_t x_14; +x_14 = lean_nat_dec_lt(x_4, x_1); +if (x_14 == 0) +{ +lean_object* x_15; +lean_dec(x_4); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_5); +lean_ctor_set(x_15, 1, x_7); +return x_15; +} +else +{ +lean_object* x_16; +x_16 = lean_array_fget_borrowed(x_2, x_4); +switch (lean_obj_tag(x_16)) { +case 6: +{ +x_8 = x_3; +x_9 = x_7; +goto block_13; +} +case 13: +{ +x_8 = x_3; +x_9 = x_7; +goto block_13; +} +default: +{ +lean_object* x_17; +x_17 = l_Lean_IR_EmitC_toCType(x_16, x_6, x_7); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +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_ref(x_17); +x_20 = lean_string_append(x_19, x_18); +lean_dec(x_18); +x_21 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__1___redArg___closed__0; +x_22 = lean_string_append(x_20, x_21); +lean_inc(x_4); +x_23 = l_Nat_reprFast(x_4); +x_24 = lean_string_append(x_22, x_23); +lean_dec_ref(x_23); +x_25 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_26 = lean_string_append(x_24, x_25); +x_27 = l_Lean_IR_EmitC_emitLn___closed__0; +x_28 = lean_string_append(x_26, x_27); +x_8 = x_3; +x_9 = x_28; +goto block_13; +} +else +{ +uint8_t x_29; +lean_dec(x_4); +x_29 = !lean_is_exclusive(x_17); +if (x_29 == 0) +{ +return x_17; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_17, 0); +x_31 = lean_ctor_get(x_17, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_17); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; +} +} +} +} +} +block_13: +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_add(x_4, x_10); +lean_dec(x_4); +x_4 = x_11; +x_5 = x_8; +x_7 = x_9; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__1___redArg(x_1, x_2, x_3, x_6, x_7, x_9, x_10); +return x_11; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitStructDefn___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" {", 2, 2); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitStructDefn___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("union {", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitStructDefn___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.IR.EmitC.emitStructDefn", 28, 28); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitStructDefn___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("assertion violation: tys.size ≤ 256\n ", 42, 40); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitStructDefn___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_IR_EmitC_emitStructDefn___closed__3; +x_2 = lean_unsigned_to_nat(4u); +x_3 = lean_unsigned_to_nat(868u); +x_4 = l_Lean_IR_EmitC_emitStructDefn___closed__2; +x_5 = l_Lean_IR_EmitC_emitUSet___closed__3; +x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); +return x_6; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitStructDefn___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("} cs;", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitStructDefn___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("uint8_t tag;", 12, 12); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitStructDefn___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("size_t u[", 9, 9); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitStructDefn___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("uint8_t s[", 10, 10); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitStructDefn___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_IR_EmitC_emitUSet___closed__5; +x_2 = lean_unsigned_to_nat(9u); +x_3 = lean_unsigned_to_nat(880u); +x_4 = l_Lean_IR_EmitC_emitStructDefn___closed__2; +x_5 = l_Lean_IR_EmitC_emitUSet___closed__3; +x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructDefn(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_13 = l_Lean_IR_EmitC_emitStructDefn___closed__0; +x_14 = lean_string_append(x_2, x_13); +x_15 = lean_string_append(x_4, x_14); +lean_dec_ref(x_14); +x_16 = l_Lean_IR_EmitC_emitLn___closed__0; +x_17 = lean_string_append(x_15, x_16); +switch (lean_obj_tag(x_1)) { +case 11: +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_18 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_18); +lean_dec_ref(x_1); +x_19 = lean_array_get_size(x_18); +x_20 = lean_unsigned_to_nat(0u); +x_21 = l_Lean_IR_EmitC_emitStructDefn___closed__1; +x_22 = lean_string_append(x_17, x_21); +x_23 = lean_string_append(x_22, x_16); +x_24 = lean_box(0); +x_25 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__0___redArg(x_19, x_18, x_24, x_20, x_24, x_3, x_23); +lean_dec_ref(x_18); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec_ref(x_25); +x_27 = lean_unsigned_to_nat(256u); +x_28 = lean_nat_dec_le(x_19, x_27); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; +x_29 = l_Lean_IR_EmitC_emitStructDefn___closed__4; +x_30 = l_panic___at___00Lean_IR_EmitC_emitUSet_spec__0(x_29, x_3, x_26); +return x_30; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +lean_dec_ref(x_3); +x_31 = l_Lean_IR_EmitC_emitStructDefn___closed__5; +x_32 = lean_string_append(x_26, x_31); +x_33 = lean_string_append(x_32, x_16); +x_34 = l_Lean_IR_EmitC_emitStructDefn___closed__6; +x_35 = lean_string_append(x_33, x_34); +x_36 = lean_string_append(x_35, x_16); +x_5 = x_36; +goto block_12; +} +} +else +{ +lean_dec_ref(x_3); +return x_25; +} +} +case 10: +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_37 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_37); +x_38 = lean_ctor_get(x_1, 2); +lean_inc(x_38); +x_39 = lean_ctor_get(x_1, 3); +lean_inc(x_39); +lean_dec_ref(x_1); +x_40 = lean_array_get_size(x_37); +x_41 = lean_unsigned_to_nat(0u); +x_42 = lean_box(0); +x_43 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__1___redArg(x_40, x_37, x_42, x_41, x_42, x_3, x_17); +lean_dec_ref(x_3); +lean_dec_ref(x_37); +if (lean_obj_tag(x_43) == 0) +{ +lean_object* x_44; uint8_t x_59; +x_44 = lean_ctor_get(x_43, 1); +lean_inc(x_44); +lean_dec_ref(x_43); +x_59 = lean_nat_dec_eq(x_38, x_41); +if (x_59 == 0) +{ +goto block_58; +} +else +{ +uint8_t x_60; +x_60 = lean_nat_dec_eq(x_39, x_41); +if (x_60 == 0) +{ +goto block_58; +} +else +{ +lean_dec(x_39); +lean_dec(x_38); +x_5 = x_44; +goto block_12; +} +} +block_58: +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_45 = l_Lean_IR_EmitC_emitStructDefn___closed__7; +x_46 = l_Nat_reprFast(x_38); +x_47 = lean_string_append(x_45, x_46); +lean_dec_ref(x_46); +x_48 = l_Lean_IR_EmitC_emitUProj___closed__0; +x_49 = lean_string_append(x_47, x_48); +x_50 = lean_string_append(x_44, x_49); +lean_dec_ref(x_49); +x_51 = lean_string_append(x_50, x_16); +x_52 = l_Lean_IR_EmitC_emitStructDefn___closed__8; +x_53 = l_Nat_reprFast(x_39); +x_54 = lean_string_append(x_52, x_53); +lean_dec_ref(x_53); +x_55 = lean_string_append(x_54, x_48); +x_56 = lean_string_append(x_51, x_55); +lean_dec_ref(x_55); +x_57 = lean_string_append(x_56, x_16); +x_5 = x_57; +goto block_12; +} +} +else +{ +lean_dec(x_39); +lean_dec(x_38); +return x_43; +} +} +default: +{ +lean_object* x_61; lean_object* x_62; +lean_dec(x_1); +x_61 = l_Lean_IR_EmitC_emitStructDefn___closed__9; +x_62 = l_panic___at___00Lean_IR_EmitC_emitUSet_spec__0(x_61, x_3, x_17); +if (lean_obj_tag(x_62) == 0) +{ +lean_object* x_63; +x_63 = lean_ctor_get(x_62, 1); +lean_inc(x_63); +lean_dec_ref(x_62); +x_5 = x_63; +goto block_12; +} +else +{ +return x_62; +} +} +} +block_12: +{ +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 = l_Lean_IR_EmitC_emitApp___closed__2; +x_7 = lean_string_append(x_5, x_6); +x_8 = l_Lean_IR_EmitC_emitLn___closed__0; +x_9 = lean_box(0); +x_10 = lean_string_append(x_7, x_8); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_9); +lean_ctor_set(x_11, 1, x_10); +return x_11; +} +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec_ref(x_9); +lean_dec_ref(x_2); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec_ref(x_9); +lean_dec_ref(x_2); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__0___redArg___boxed(lean_object* x_1, lean_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_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__0___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec_ref(x_6); +lean_dec_ref(x_2); +lean_dec(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__1___redArg___boxed(lean_object* x_1, lean_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_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__1___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec_ref(x_6); +lean_dec_ref(x_2); +lean_dec(x_1); +return x_8; +} +} +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitUnionSwitch_spec__0___redArg___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("break;", 6, 6); +return x_1; +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitUnionSwitch_spec__0___redArg(lean_object* x_1, lean_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_nat_dec_lt(x_4, x_1); +if (x_8 == 0) +{ +lean_object* x_9; +lean_dec_ref(x_6); +lean_dec(x_4); +lean_dec_ref(x_2); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_5); +lean_ctor_set(x_9, 1, x_7); +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_10 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitCase_spec__4___closed__0; +x_11 = lean_string_append(x_7, x_10); +lean_inc(x_4); +x_12 = l_Nat_reprFast(x_4); +x_13 = lean_string_append(x_11, x_12); +lean_dec_ref(x_12); +x_14 = l_Lean_IR_EmitC_emitJPs___closed__0; +x_15 = lean_string_append(x_13, x_14); +x_16 = l_Lean_IR_EmitC_emitLn___closed__0; +x_17 = lean_string_append(x_15, x_16); +lean_inc_ref(x_2); +lean_inc_ref(x_6); +lean_inc(x_4); +x_18 = lean_apply_4(x_2, x_4, lean_box(0), x_6, x_17); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +lean_dec_ref(x_18); +x_20 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitUnionSwitch_spec__0___redArg___closed__0; +x_21 = lean_string_append(x_19, x_20); +x_22 = lean_string_append(x_21, x_16); +x_23 = lean_unsigned_to_nat(1u); +x_24 = lean_nat_add(x_4, x_23); +lean_dec(x_4); +{ +lean_object* _tmp_3 = x_24; +lean_object* _tmp_4 = x_3; +lean_object* _tmp_6 = x_22; +x_4 = _tmp_3; +x_5 = _tmp_4; +x_7 = _tmp_6; +} +goto _start; +} +else +{ +lean_dec_ref(x_6); +lean_dec(x_4); +lean_dec_ref(x_2); +return x_18; +} +} +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitUnionSwitch_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitUnionSwitch_spec__0___redArg(x_1, x_2, x_3, x_6, x_7, x_9, x_10); +return x_11; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitUnionSwitch___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" == 0) {", 8, 8); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUnionSwitch(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_unsigned_to_nat(1u); +x_7 = lean_nat_dec_eq(x_1, x_6); +if (x_7 == 0) +{ +lean_object* x_8; uint8_t x_9; +x_8 = lean_unsigned_to_nat(2u); +x_9 = lean_nat_dec_eq(x_1, x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_10 = lean_unsigned_to_nat(0u); +x_11 = l_Lean_IR_EmitC_emitCase___closed__0; +x_12 = lean_string_append(x_5, x_11); +x_13 = lean_string_append(x_12, x_2); +x_14 = l_Lean_IR_EmitC_emitCase___closed__1; +x_15 = lean_string_append(x_13, x_14); +x_16 = l_Lean_IR_EmitC_emitLn___closed__0; +x_17 = lean_string_append(x_15, x_16); +x_18 = lean_box(0); +x_19 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitUnionSwitch_spec__0___redArg(x_1, x_3, x_18, x_10, x_18, x_4, x_17); +if (lean_obj_tag(x_19) == 0) +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_21 = lean_ctor_get(x_19, 1); +x_22 = lean_ctor_get(x_19, 0); +lean_dec(x_22); +x_23 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; +x_24 = lean_string_append(x_21, x_23); +x_25 = lean_string_append(x_24, x_16); +lean_ctor_set(x_19, 1, x_25); +lean_ctor_set(x_19, 0, x_18); +return x_19; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_26 = lean_ctor_get(x_19, 1); +lean_inc(x_26); +lean_dec(x_19); +x_27 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; +x_28 = lean_string_append(x_26, x_27); +x_29 = lean_string_append(x_28, x_16); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_18); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} +} +else +{ +return x_19; +} +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_31 = l_Lean_IR_EmitC_emitIf___closed__0; +x_32 = lean_string_append(x_5, x_31); +x_33 = lean_string_append(x_32, x_2); +x_34 = l_Lean_IR_EmitC_emitUnionSwitch___closed__0; +x_35 = lean_string_append(x_33, x_34); +x_36 = l_Lean_IR_EmitC_emitLn___closed__0; +x_37 = lean_string_append(x_35, x_36); +x_38 = lean_unsigned_to_nat(0u); +lean_inc_ref(x_3); +lean_inc_ref(x_4); +x_39 = lean_apply_4(x_3, x_38, lean_box(0), x_4, x_37); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +lean_dec_ref(x_39); +x_41 = l_Lean_IR_EmitC_emitMainFn___closed__2; +x_42 = lean_string_append(x_40, x_41); +x_43 = lean_string_append(x_42, x_36); +x_44 = lean_apply_4(x_3, x_6, lean_box(0), x_4, x_43); +if (lean_obj_tag(x_44) == 0) +{ +uint8_t x_45; +x_45 = !lean_is_exclusive(x_44); +if (x_45 == 0) +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_46 = lean_ctor_get(x_44, 1); +x_47 = lean_ctor_get(x_44, 0); +lean_dec(x_47); +x_48 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; +x_49 = lean_string_append(x_46, x_48); +x_50 = lean_box(0); +x_51 = lean_string_append(x_49, x_36); +lean_ctor_set(x_44, 1, x_51); +lean_ctor_set(x_44, 0, x_50); +return x_44; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_52 = lean_ctor_get(x_44, 1); +lean_inc(x_52); +lean_dec(x_44); +x_53 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; +x_54 = lean_string_append(x_52, x_53); +x_55 = lean_box(0); +x_56 = lean_string_append(x_54, x_36); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +return x_57; +} +} +else +{ +return x_44; +} +} +else +{ +lean_dec_ref(x_4); +lean_dec_ref(x_3); +return x_39; +} +} +} +else +{ +lean_object* x_58; lean_object* x_59; +x_58 = lean_unsigned_to_nat(0u); +x_59 = lean_apply_4(x_3, x_58, lean_box(0), x_4, x_5); +return x_59; +} +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitUnionSwitch_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitUnionSwitch_spec__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitUnionSwitch_spec__0___redArg___boxed(lean_object* x_1, lean_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_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitUnionSwitch_spec__0___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUnionSwitch___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_IR_EmitC_emitUnionSwitch(x_1, x_2, x_3, x_4, x_5); +lean_dec_ref(x_2); +lean_dec(x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUnionSwitchWithImpossible___lam__0(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitUnionSwitchWithImpossible_spec__1(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; uint8_t x_11; +x_11 = lean_usize_dec_eq(x_3, x_4); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = lean_array_uget(x_2, x_3); +lean_inc_ref(x_1); +lean_inc(x_12); +x_13 = lean_apply_1(x_1, x_12); +x_14 = lean_unbox(x_13); +if (x_14 == 0) +{ +lean_dec(x_12); +x_6 = x_5; +goto block_10; +} +else +{ +lean_object* x_15; +x_15 = lean_array_push(x_5, x_12); +x_6 = x_15; +goto block_10; +} +} +else +{ +lean_dec_ref(x_1); +return x_5; +} +block_10: +{ +size_t x_7; size_t x_8; +x_7 = 1; +x_8 = lean_usize_add(x_3, x_7); +x_3 = x_8; +x_5 = x_6; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitUnionSwitchWithImpossible_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +uint8_t x_9; +x_9 = lean_usize_dec_lt(x_5, x_4); +if (x_9 == 0) +{ +lean_object* x_10; +lean_dec_ref(x_7); +lean_dec_ref(x_1); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_6); +lean_ctor_set(x_10, 1, x_8); +return x_10; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_11 = lean_array_uget(x_3, x_5); +x_12 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitCase_spec__4___closed__0; +x_13 = lean_string_append(x_8, x_12); +lean_inc(x_11); +x_14 = l_Nat_reprFast(x_11); +x_15 = lean_string_append(x_13, x_14); +lean_dec_ref(x_14); +x_16 = l_Lean_IR_EmitC_emitJPs___closed__0; +x_17 = lean_string_append(x_15, x_16); +x_18 = l_Lean_IR_EmitC_emitLn___closed__0; +x_19 = lean_string_append(x_17, x_18); +lean_inc_ref(x_1); +lean_inc_ref(x_7); +x_20 = lean_apply_4(x_1, x_11, lean_box(0), x_7, x_19); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; size_t x_25; size_t x_26; +x_21 = lean_ctor_get(x_20, 1); +lean_inc(x_21); +lean_dec_ref(x_20); +x_22 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitUnionSwitch_spec__0___redArg___closed__0; +x_23 = lean_string_append(x_21, x_22); +x_24 = lean_string_append(x_23, x_18); +x_25 = 1; +x_26 = lean_usize_add(x_5, x_25); +{ +size_t _tmp_4 = x_26; +lean_object* _tmp_5 = x_2; +lean_object* _tmp_7 = x_24; +x_5 = _tmp_4; +x_6 = _tmp_5; +x_8 = _tmp_7; +} +goto _start; +} +else +{ +lean_dec_ref(x_7); +lean_dec_ref(x_1); +return x_20; +} +} +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitUnionSwitchWithImpossible___closed__0() { +_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_IR_EmitC_emitUnionSwitchWithImpossible___lam__0___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_IR_EmitC_emitUnionSwitchWithImpossible___lam__0(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUnionSwitchWithImpossible(lean_object* x_1, 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_69; lean_object* x_70; uint8_t x_71; +x_7 = lean_alloc_closure((void*)(l_Lean_IR_EmitC_emitUnionSwitchWithImpossible___lam__0___boxed), 1, 0); +x_8 = l_Array_ofFn___redArg(x_1, x_7); +x_9 = lean_unsigned_to_nat(0u); +x_69 = lean_array_get_size(x_8); +x_70 = l_Lean_IR_EmitC_emitUnionSwitchWithImpossible___closed__0; +x_71 = lean_nat_dec_lt(x_9, x_69); +if (x_71 == 0) +{ +lean_dec_ref(x_8); +lean_dec_ref(x_3); +x_10 = x_70; +goto block_68; +} +else +{ +uint8_t x_72; +x_72 = lean_nat_dec_le(x_69, x_69); +if (x_72 == 0) +{ +if (x_71 == 0) +{ +lean_dec_ref(x_8); +lean_dec_ref(x_3); +x_10 = x_70; +goto block_68; +} +else { -lean_object* x_6; -x_6 = l_Lean_IR_EmitC_emitBox(x_1, x_2, x_3, x_4, x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -return x_6; +size_t x_73; size_t x_74; lean_object* x_75; +x_73 = 0; +x_74 = lean_usize_of_nat(x_69); +x_75 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitUnionSwitchWithImpossible_spec__1(x_3, x_8, x_73, x_74, x_70); +lean_dec_ref(x_8); +x_10 = x_75; +goto block_68; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitBox___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +else { -lean_object* x_5; -x_5 = l_Lean_IR_EmitC_emitBox___redArg(x_1, x_2, x_3, x_4); -lean_dec(x_3); -return x_5; +size_t x_76; size_t x_77; lean_object* x_78; +x_76 = 0; +x_77 = lean_usize_of_nat(x_69); +x_78 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitUnionSwitchWithImpossible_spec__1(x_3, x_8, x_76, x_77, x_70); +lean_dec_ref(x_8); +x_10 = x_78; +goto block_68; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUnbox___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +block_68: { -lean_object* x_5; uint8_t x_6; -x_5 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_4); -x_6 = !lean_is_exclusive(x_5); -if (x_6 == 0) +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = lean_array_get_size(x_10); +x_12 = lean_unsigned_to_nat(1u); +x_13 = lean_nat_dec_eq(x_11, x_12); +if (x_13 == 0) { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_7 = lean_ctor_get(x_5, 1); -x_8 = lean_ctor_get(x_5, 0); -lean_dec(x_8); -x_9 = l_Lean_IR_getUnboxOpName(x_2); -x_10 = lean_string_append(x_7, x_9); -lean_dec_ref(x_9); -x_11 = l_Lean_IR_EmitC_emitFnDeclAux___closed__5; -x_12 = lean_string_append(x_10, x_11); -x_13 = l_Lean_IR_EmitC_argToCString___closed__0; -x_14 = l_Nat_reprFast(x_3); -x_15 = lean_string_append(x_13, x_14); -lean_dec_ref(x_14); -x_16 = lean_string_append(x_12, x_15); -lean_dec_ref(x_15); -x_17 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_18 = lean_string_append(x_16, x_17); -x_19 = l_Lean_IR_EmitC_emitLn___closed__0; -x_20 = lean_box(0); -x_21 = lean_string_append(x_18, x_19); -lean_ctor_set(x_5, 1, x_21); -lean_ctor_set(x_5, 0, x_20); -return x_5; +lean_object* x_14; uint8_t x_15; +x_14 = lean_unsigned_to_nat(2u); +x_15 = lean_nat_dec_eq(x_11, x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; size_t x_24; size_t x_25; lean_object* x_26; +x_16 = l_Lean_IR_EmitC_emitCase___closed__0; +x_17 = lean_string_append(x_6, x_16); +x_18 = lean_string_append(x_17, x_2); +x_19 = l_Lean_IR_EmitC_emitCase___closed__1; +x_20 = lean_string_append(x_18, x_19); +x_21 = l_Lean_IR_EmitC_emitLn___closed__0; +x_22 = lean_string_append(x_20, x_21); +x_23 = lean_box(0); +x_24 = lean_array_size(x_10); +x_25 = 0; +x_26 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitUnionSwitchWithImpossible_spec__0(x_4, x_23, x_10, x_24, x_25, x_23, x_5, x_22); +lean_dec_ref(x_10); +if (lean_obj_tag(x_26) == 0) +{ +uint8_t x_27; +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_28 = lean_ctor_get(x_26, 1); +x_29 = lean_ctor_get(x_26, 0); +lean_dec(x_29); +x_30 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; +x_31 = lean_string_append(x_28, x_30); +x_32 = lean_string_append(x_31, x_21); +lean_ctor_set(x_26, 1, x_32); +lean_ctor_set(x_26, 0, x_23); +return x_26; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_22 = lean_ctor_get(x_5, 1); -lean_inc(x_22); -lean_dec(x_5); -x_23 = l_Lean_IR_getUnboxOpName(x_2); -x_24 = lean_string_append(x_22, x_23); -lean_dec_ref(x_23); -x_25 = l_Lean_IR_EmitC_emitFnDeclAux___closed__5; -x_26 = lean_string_append(x_24, x_25); -x_27 = l_Lean_IR_EmitC_argToCString___closed__0; -x_28 = l_Nat_reprFast(x_3); -x_29 = lean_string_append(x_27, x_28); -lean_dec_ref(x_28); -x_30 = lean_string_append(x_26, x_29); -lean_dec_ref(x_29); -x_31 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_32 = lean_string_append(x_30, x_31); -x_33 = l_Lean_IR_EmitC_emitLn___closed__0; -x_34 = lean_box(0); -x_35 = lean_string_append(x_32, x_33); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; -} +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_33 = lean_ctor_get(x_26, 1); +lean_inc(x_33); +lean_dec(x_26); +x_34 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; +x_35 = lean_string_append(x_33, x_34); +x_36 = lean_string_append(x_35, x_21); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_23); +lean_ctor_set(x_37, 1, x_36); +return x_37; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUnbox(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +else { -lean_object* x_6; -x_6 = l_Lean_IR_EmitC_emitUnbox___redArg(x_1, x_2, x_3, x_5); -return x_6; +return x_26; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUnbox___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +else { -lean_object* x_6; -x_6 = l_Lean_IR_EmitC_emitUnbox(x_1, x_2, x_3, x_4, x_5); -lean_dec_ref(x_4); -lean_dec(x_2); -return x_6; -} +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_38 = l_Lean_IR_EmitC_emitIf___closed__0; +x_39 = lean_string_append(x_6, x_38); +x_40 = lean_string_append(x_39, x_2); +x_41 = l_Lean_IR_EmitC_emitUnionSwitch___closed__0; +x_42 = lean_string_append(x_40, x_41); +x_43 = l_Lean_IR_EmitC_emitLn___closed__0; +x_44 = lean_string_append(x_42, x_43); +x_45 = lean_array_fget_borrowed(x_10, x_9); +lean_inc_ref(x_4); +lean_inc_ref(x_5); +lean_inc(x_45); +x_46 = lean_apply_4(x_4, x_45, lean_box(0), x_5, x_44); +if (lean_obj_tag(x_46) == 0) +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_47 = lean_ctor_get(x_46, 1); +lean_inc(x_47); +lean_dec_ref(x_46); +x_48 = l_Lean_IR_EmitC_emitMainFn___closed__2; +x_49 = lean_string_append(x_47, x_48); +x_50 = lean_string_append(x_49, x_43); +x_51 = lean_array_fget(x_10, x_12); +lean_dec_ref(x_10); +x_52 = lean_apply_4(x_4, x_51, lean_box(0), x_5, x_50); +if (lean_obj_tag(x_52) == 0) +{ +uint8_t x_53; +x_53 = !lean_is_exclusive(x_52); +if (x_53 == 0) +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_54 = lean_ctor_get(x_52, 1); +x_55 = lean_ctor_get(x_52, 0); +lean_dec(x_55); +x_56 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; +x_57 = lean_string_append(x_54, x_56); +x_58 = lean_box(0); +x_59 = lean_string_append(x_57, x_43); +lean_ctor_set(x_52, 1, x_59); +lean_ctor_set(x_52, 0, x_58); +return x_52; } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUnbox___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +else { -lean_object* x_5; -x_5 = l_Lean_IR_EmitC_emitUnbox___redArg(x_1, x_2, x_3, x_4); -lean_dec(x_2); -return x_5; +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_60 = lean_ctor_get(x_52, 1); +lean_inc(x_60); +lean_dec(x_52); +x_61 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; +x_62 = lean_string_append(x_60, x_61); +x_63 = lean_box(0); +x_64 = lean_string_append(x_62, x_43); +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_64); +return x_65; } } -static lean_object* _init_l_Lean_IR_EmitC_emitIsShared___redArg___closed__0() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("!lean_is_exclusive(", 19, 19); -return x_1; +return x_52; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitIsShared___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; uint8_t x_5; -x_4 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_3); -x_5 = !lean_is_exclusive(x_4); -if (x_5 == 0) +else { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_6 = lean_ctor_get(x_4, 1); -x_7 = lean_ctor_get(x_4, 0); -lean_dec(x_7); -x_8 = l_Lean_IR_EmitC_emitIsShared___redArg___closed__0; -x_9 = lean_string_append(x_6, x_8); -x_10 = l_Lean_IR_EmitC_argToCString___closed__0; -x_11 = l_Nat_reprFast(x_2); -x_12 = lean_string_append(x_10, x_11); -lean_dec_ref(x_11); -x_13 = lean_string_append(x_9, x_12); -lean_dec_ref(x_12); -x_14 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_15 = lean_string_append(x_13, x_14); -x_16 = l_Lean_IR_EmitC_emitLn___closed__0; -x_17 = lean_box(0); -x_18 = lean_string_append(x_15, x_16); -lean_ctor_set(x_4, 1, x_18); -lean_ctor_set(x_4, 0, x_17); -return x_4; +lean_dec_ref(x_10); +lean_dec_ref(x_5); +lean_dec_ref(x_4); +return x_46; +} +} } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_19 = lean_ctor_get(x_4, 1); -lean_inc(x_19); -lean_dec(x_4); -x_20 = l_Lean_IR_EmitC_emitIsShared___redArg___closed__0; -x_21 = lean_string_append(x_19, x_20); -x_22 = l_Lean_IR_EmitC_argToCString___closed__0; -x_23 = l_Nat_reprFast(x_2); -x_24 = lean_string_append(x_22, x_23); -lean_dec_ref(x_23); -x_25 = lean_string_append(x_21, x_24); -lean_dec_ref(x_24); -x_26 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_27 = lean_string_append(x_25, x_26); -x_28 = l_Lean_IR_EmitC_emitLn___closed__0; -x_29 = lean_box(0); -x_30 = lean_string_append(x_27, x_28); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_30); -return x_31; +lean_object* x_66; lean_object* x_67; +x_66 = lean_array_fget(x_10, x_9); +lean_dec_ref(x_10); +x_67 = lean_apply_4(x_4, x_66, lean_box(0), x_5, x_6); +return x_67; } } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitIsShared(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitUnionSwitchWithImpossible_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_5; -x_5 = l_Lean_IR_EmitC_emitIsShared___redArg(x_1, x_2, x_4); -return x_5; +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitUnionSwitchWithImpossible_spec__1(x_1, x_2, x_6, x_7, x_5); +lean_dec_ref(x_2); +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitIsShared___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitUnionSwitchWithImpossible_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_5; -x_5 = l_Lean_IR_EmitC_emitIsShared(x_1, x_2, x_3, x_4); +size_t x_9; size_t x_10; lean_object* x_11; +x_9 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_10 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_11 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitUnionSwitchWithImpossible_spec__0(x_1, x_2, x_3, x_9, x_10, x_6, x_7, x_8); lean_dec_ref(x_3); -return x_5; +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_toHexDigit(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUnionSwitchWithImpossible___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -uint32_t x_2; lean_object* x_3; lean_object* x_4; -x_2 = l_Nat_digitChar(x_1); -x_3 = l_panic___at___00Lean_IR_EmitC_toCType_spec__0___closed__0; -x_4 = lean_string_push(x_3, x_2); -return x_4; +lean_object* x_7; +x_7 = l_Lean_IR_EmitC_emitUnionSwitchWithImpossible(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec_ref(x_2); +return x_7; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_toHexDigit___boxed(lean_object* x_1) { +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructIncDecFn_spec__0___redArg___closed__0() { _start: { -lean_object* x_2; -x_2 = l_Lean_IR_EmitC_toHexDigit(x_1); -lean_dec(x_1); -return x_2; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("x.i", 3, 3); +return x_1; } } -static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__0() { +static lean_object* _init_l_Lean_IR_EmitC_emitStructIncDecFn___lam__0___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("\\x", 2, 2); +x_1 = lean_mk_string_unchecked("n", 1, 1); return x_1; } } -static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__1() { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructIncDecFn_spec__0_spec__0___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("\\\?", 2, 2); -return x_1; +lean_object* x_10; lean_object* x_15; uint8_t x_18; +x_18 = lean_nat_dec_lt(x_6, x_1); +if (x_18 == 0) +{ +lean_object* x_19; +lean_dec(x_6); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_7); +lean_ctor_set(x_19, 1, x_9); +return x_19; +} +else +{ +lean_object* x_20; +x_20 = lean_array_fget_borrowed(x_2, x_6); +switch (lean_obj_tag(x_20)) { +case 13: +{ +x_10 = x_9; +goto block_14; } +case 6: +{ +x_10 = x_9; +goto block_14; } -static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__2() { -_start: +default: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("\\\"", 2, 2); -return x_1; +uint8_t x_21; +x_21 = l_Lean_IR_IRType_isScalar(x_20); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructIncDecFn_spec__0___redArg___closed__0; +lean_inc(x_6); +x_23 = l_Nat_reprFast(x_6); +x_24 = lean_string_append(x_22, x_23); +lean_dec_ref(x_23); +if (x_4 == 0) +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_unsigned_to_nat(1u); +x_26 = l_Lean_IR_EmitC_emitDecOfType(x_24, x_20, x_25, x_5, x_8, x_9); +lean_dec_ref(x_24); +x_15 = x_26; +goto block_17; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_unsigned_to_nat(0u); +x_28 = l_Lean_IR_EmitC_emitStructIncDecFn___lam__0___closed__1; +x_29 = l_Lean_IR_EmitC_emitIncOfType(x_24, x_20, x_27, x_5, x_28, x_8, x_9); +lean_dec_ref(x_24); +x_15 = x_29; +goto block_17; +} +} +else +{ +x_10 = x_9; +goto block_14; } } -static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("\\\\", 2, 2); -return x_1; } } -static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__4() { -_start: +block_14: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("\\t", 2, 2); -return x_1; +lean_object* x_11; lean_object* x_12; +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_add(x_6, x_11); +lean_dec(x_6); +{ +lean_object* _tmp_5 = x_12; +lean_object* _tmp_6 = x_3; +lean_object* _tmp_8 = x_10; +x_6 = _tmp_5; +x_7 = _tmp_6; +x_9 = _tmp_8; } +goto _start; } -static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__5() { -_start: +block_17: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("\\r", 2, 2); -return x_1; -} +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec_ref(x_15); +x_10 = x_16; +goto block_14; } -static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__6() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("\\n", 2, 2); -return x_1; +lean_dec(x_6); +return x_15; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructIncDecFn_spec__0___redArg(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_5 = lean_ctor_get(x_1, 1); -x_6 = lean_ctor_get(x_1, 2); -x_7 = lean_nat_sub(x_6, x_5); -x_8 = lean_nat_dec_eq(x_3, x_7); -lean_dec(x_7); -if (x_8 == 0) -{ -lean_object* x_9; uint32_t x_10; uint32_t x_11; uint8_t x_12; -x_9 = lean_string_utf8_next_fast(x_2, x_3); -x_10 = lean_string_utf8_get_fast(x_2, x_3); -lean_dec(x_3); -x_11 = 10; -x_12 = lean_uint32_dec_eq(x_10, x_11); -if (x_12 == 0) +lean_object* x_10; lean_object* x_15; uint8_t x_18; +x_18 = lean_nat_dec_lt(x_6, x_1); +if (x_18 == 0) { -uint32_t x_13; uint8_t x_14; -x_13 = 13; -x_14 = lean_uint32_dec_eq(x_10, x_13); -if (x_14 == 0) +lean_object* x_19; +lean_dec(x_6); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_7); +lean_ctor_set(x_19, 1, x_9); +return x_19; +} +else { -uint32_t x_15; uint8_t x_16; -x_15 = 9; -x_16 = lean_uint32_dec_eq(x_10, x_15); -if (x_16 == 0) +lean_object* x_20; +x_20 = lean_array_fget_borrowed(x_2, x_6); +switch (lean_obj_tag(x_20)) { +case 13: { -uint32_t x_17; uint8_t x_18; -x_17 = 92; -x_18 = lean_uint32_dec_eq(x_10, x_17); -if (x_18 == 0) +x_10 = x_9; +goto block_14; +} +case 6: { -uint32_t x_19; uint8_t x_20; -x_19 = 34; -x_20 = lean_uint32_dec_eq(x_10, x_19); -if (x_20 == 0) +x_10 = x_9; +goto block_14; +} +default: { -uint32_t x_21; uint8_t x_22; -x_21 = 63; -x_22 = lean_uint32_dec_eq(x_10, x_21); -if (x_22 == 0) +uint8_t x_21; +x_21 = l_Lean_IR_IRType_isScalar(x_20); +if (x_21 == 0) { -lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_23 = lean_uint32_to_nat(x_10); -x_24 = lean_unsigned_to_nat(31u); -x_25 = lean_nat_dec_le(x_23, x_24); -if (x_25 == 0) +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructIncDecFn_spec__0___redArg___closed__0; +lean_inc(x_6); +x_23 = l_Nat_reprFast(x_6); +x_24 = lean_string_append(x_22, x_23); +lean_dec_ref(x_23); +if (x_3 == 0) { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -lean_dec(x_23); -x_26 = l_panic___at___00Lean_IR_EmitC_toCType_spec__0___closed__0; -x_27 = lean_string_push(x_26, x_10); -x_28 = lean_string_append(x_4, x_27); -lean_dec_ref(x_27); -x_3 = x_9; -x_4 = x_28; -goto _start; +lean_object* x_25; lean_object* x_26; +x_25 = lean_unsigned_to_nat(1u); +x_26 = l_Lean_IR_EmitC_emitDecOfType(x_24, x_20, x_25, x_4, x_8, x_9); +lean_dec_ref(x_24); +x_15 = x_26; +goto block_17; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_30 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__0; -x_31 = lean_unsigned_to_nat(16u); -x_32 = lean_unsigned_to_nat(4u); -x_33 = lean_nat_shiftr(x_23, x_32); -x_34 = l_Lean_IR_EmitC_toHexDigit(x_33); -lean_dec(x_33); -x_35 = lean_string_append(x_30, x_34); -lean_dec_ref(x_34); -x_36 = lean_nat_mod(x_23, x_31); -lean_dec(x_23); -x_37 = l_Lean_IR_EmitC_toHexDigit(x_36); -lean_dec(x_36); -x_38 = lean_string_append(x_35, x_37); -lean_dec_ref(x_37); -x_39 = lean_string_append(x_4, x_38); -lean_dec_ref(x_38); -x_3 = x_9; -x_4 = x_39; -goto _start; +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_unsigned_to_nat(0u); +x_28 = l_Lean_IR_EmitC_emitStructIncDecFn___lam__0___closed__1; +x_29 = l_Lean_IR_EmitC_emitIncOfType(x_24, x_20, x_27, x_4, x_28, x_8, x_9); +lean_dec_ref(x_24); +x_15 = x_29; +goto block_17; } } else { -lean_object* x_41; lean_object* x_42; -x_41 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__1; -x_42 = lean_string_append(x_4, x_41); -x_3 = x_9; -x_4 = x_42; -goto _start; +x_10 = x_9; +goto block_14; } } -else -{ -lean_object* x_44; lean_object* x_45; -x_44 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__2; -x_45 = lean_string_append(x_4, x_44); -x_3 = x_9; -x_4 = x_45; -goto _start; } } -else +block_14: { -lean_object* x_47; lean_object* x_48; -x_47 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__3; -x_48 = lean_string_append(x_4, x_47); -x_3 = x_9; -x_4 = x_48; -goto _start; -} +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_add(x_6, x_11); +lean_dec(x_6); +x_13 = l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructIncDecFn_spec__0_spec__0___redArg(x_1, x_2, x_5, x_3, x_4, x_12, x_5, x_8, x_10); +return x_13; } -else +block_17: { -lean_object* x_50; lean_object* x_51; -x_50 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__4; -x_51 = lean_string_append(x_4, x_50); -x_3 = x_9; -x_4 = x_51; -goto _start; -} +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec_ref(x_15); +x_10 = x_16; +goto block_14; } else { -lean_object* x_53; lean_object* x_54; -x_53 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__5; -x_54 = lean_string_append(x_4, x_53); -x_3 = x_9; -x_4 = x_54; -goto _start; +lean_dec(x_6); +return x_15; } } -else -{ -lean_object* x_56; lean_object* x_57; -x_56 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__6; -x_57 = lean_string_append(x_4, x_56); -x_3 = x_9; -x_4 = x_57; -goto _start; } } -else +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructIncDecFn_spec__0(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -lean_dec(x_3); -return x_4; -} +lean_object* x_13; +x_13 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructIncDecFn_spec__0___redArg(x_1, x_2, x_3, x_4, x_5, x_8, x_9, x_11, x_12); +return x_13; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructIncDecFn_spec__0_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_8; -x_8 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg(x_1, x_2, x_5, x_6); -return x_8; +lean_object* x_13; +x_13 = l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructIncDecFn_spec__0_spec__0___redArg(x_1, x_2, x_3, x_4, x_5, x_8, x_9, x_11, x_12); +return x_13; } } -static lean_object* _init_l_Lean_IR_EmitC_quoteString___closed__0() { +static lean_object* _init_l_Lean_IR_EmitC_emitStructIncDecFn___lam__0___closed__0() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("\"", 1, 1); +x_1 = lean_mk_string_unchecked("x.cs.c", 6, 6); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_quoteString(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructIncDecFn___lam__0(lean_object* x_1, uint8_t x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_2 = l_Lean_IR_EmitC_quoteString___closed__0; -x_3 = lean_unsigned_to_nat(0u); -x_4 = lean_string_utf8_byte_size(x_1); -lean_inc_ref(x_1); -x_5 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_5, 0, x_1); -lean_ctor_set(x_5, 1, x_3); -lean_ctor_set(x_5, 2, x_4); -x_6 = l_String_Slice_positions(x_5); -x_7 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg(x_5, x_1, x_6, x_2); -lean_dec_ref(x_1); -lean_dec_ref(x_5); -x_8 = lean_string_append(x_7, x_2); -return x_8; +lean_object* x_11; +x_11 = lean_array_fget_borrowed(x_1, x_4); +switch (lean_obj_tag(x_11)) { +case 13: +{ +lean_dec(x_4); +goto block_10; +} +case 6: +{ +lean_dec(x_4); +goto block_10; +} +default: +{ +uint8_t x_12; +x_12 = l_Lean_IR_IRType_isScalar(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = l_Lean_IR_EmitC_emitStructIncDecFn___lam__0___closed__0; +x_14 = l_Nat_reprFast(x_4); +x_15 = lean_string_append(x_13, x_14); +lean_dec_ref(x_14); +if (x_2 == 0) +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_unsigned_to_nat(1u); +x_17 = l_Lean_IR_EmitC_emitDecOfType(x_15, x_11, x_16, x_3, x_6, x_7); +lean_dec_ref(x_15); +return x_17; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_unsigned_to_nat(0u); +x_19 = l_Lean_IR_EmitC_emitStructIncDecFn___lam__0___closed__1; +x_20 = l_Lean_IR_EmitC_emitIncOfType(x_15, x_11, x_18, x_3, x_19, x_6, x_7); +lean_dec_ref(x_15); +return x_20; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: +else { -lean_object* x_8; -x_8 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -return x_8; -} +lean_dec(x_4); +goto block_10; } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___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_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg(x_1, x_2, x_3, x_4); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -return x_5; } } -static lean_object* _init_l_Lean_IR_EmitC_emitNumLit___redArg___closed__0() { -_start: +block_10: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("ULL", 3, 3); -return x_1; +lean_object* x_8; lean_object* x_9; +x_8 = lean_box(0); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_7); +return x_9; } } -static lean_object* _init_l_Lean_IR_EmitC_emitNumLit___redArg___closed__1() { +} +static lean_object* _init_l_Lean_IR_EmitC_emitStructIncDecFn___closed__0() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("((size_t)", 9, 9); +x_1 = lean_mk_string_unchecked("x.tag", 5, 5); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitNumLit___redArg___closed__2() { +static lean_object* _init_l_Lean_IR_EmitC_emitStructIncDecFn___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("ULL)", 4, 4); +x_1 = lean_mk_string_unchecked("static void ", 12, 12); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitNumLit___redArg___closed__3() { +static lean_object* _init_l_Lean_IR_EmitC_emitStructIncDecFn___closed__2() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_cstr_to_nat(\"", 18, 18); +x_1 = lean_mk_string_unchecked(" x) {", 5, 5); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitNumLit___redArg___closed__4() { +static lean_object* _init_l_Lean_IR_EmitC_emitStructIncDecFn___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("\")", 2, 2); +x_1 = lean_mk_string_unchecked(" x, size_t n) {", 15, 15); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitNumLit___redArg___closed__5() { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructIncDecFn___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_unsigned_to_nat(", 21, 21); -return x_1; +uint8_t x_8; uint8_t x_9; lean_object* x_10; +x_8 = lean_unbox(x_2); +x_9 = lean_unbox(x_3); +x_10 = l_Lean_IR_EmitC_emitStructIncDecFn___lam__0(x_1, x_8, x_9, x_4, x_5, x_6, x_7); +lean_dec_ref(x_6); +lean_dec_ref(x_1); +return x_10; } } -static lean_object* _init_l_Lean_IR_EmitC_emitNumLit___redArg___closed__6() { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructIncDecFn(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("u)", 2, 2); -return x_1; +lean_object* x_6; uint8_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_32; +x_14 = 1; +if (x_3 == 0) +{ +lean_object* x_51; +x_51 = l_Lean_IR_EmitC_structDecFnPrefix___closed__0; +x_32 = x_51; +goto block_50; } +else +{ +lean_object* x_52; +x_52 = l_Lean_IR_EmitC_structIncFnPrefix___closed__0; +x_32 = x_52; +goto block_50; } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitNumLit___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +block_13: { -uint8_t x_4; -x_4 = l_Lean_IR_IRType_isObj(x_1); -if (x_4 == 0) +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_7 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; +x_8 = lean_string_append(x_6, x_7); +x_9 = l_Lean_IR_EmitC_emitLn___closed__0; +x_10 = lean_box(0); +x_11 = lean_string_append(x_8, x_9); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set(x_12, 1, x_11); +return x_12; +} +block_31: { -lean_object* x_5; uint8_t x_6; -x_5 = lean_cstr_to_nat("4294967296"); -x_6 = lean_nat_dec_lt(x_2, x_5); -if (x_6 == 0) +if (lean_obj_tag(x_1) == 11) { -lean_object* x_7; uint8_t x_8; -x_7 = lean_box(5); -x_8 = l_Lean_IR_instBEqIRType_beq(x_1, x_7); -if (x_8 == 0) +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_17 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_17); +lean_dec_ref(x_1); +x_18 = lean_box(x_3); +x_19 = lean_box(x_14); +lean_inc_ref(x_17); +x_20 = lean_alloc_closure((void*)(l_Lean_IR_EmitC_emitStructIncDecFn___lam__0___boxed), 7, 3); +lean_closure_set(x_20, 0, x_17); +lean_closure_set(x_20, 1, x_18); +lean_closure_set(x_20, 2, x_19); +x_21 = lean_array_get_size(x_17); +lean_dec_ref(x_17); +x_22 = l_Lean_IR_EmitC_emitStructIncDecFn___closed__0; +x_23 = l_Lean_IR_EmitC_emitUnionSwitch(x_21, x_22, x_20, x_15, x_16); +if (lean_obj_tag(x_23) == 0) { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_9 = l_Nat_reprFast(x_2); -x_10 = lean_string_append(x_3, x_9); -lean_dec_ref(x_9); -x_11 = l_Lean_IR_EmitC_emitNumLit___redArg___closed__0; -x_12 = lean_box(0); -x_13 = lean_string_append(x_10, x_11); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_12); -lean_ctor_set(x_14, 1, x_13); -return x_14; +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec_ref(x_23); +x_6 = x_24; +goto block_13; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_15 = l_Lean_IR_EmitC_emitNumLit___redArg___closed__1; -x_16 = lean_string_append(x_3, x_15); -x_17 = l_Nat_reprFast(x_2); -x_18 = lean_string_append(x_16, x_17); -lean_dec_ref(x_17); -x_19 = l_Lean_IR_EmitC_emitNumLit___redArg___closed__2; -x_20 = lean_box(0); -x_21 = lean_string_append(x_18, x_19); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -return x_22; +return x_23; } } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -lean_dec(x_1); -x_23 = lean_box(0); -x_24 = l_Nat_reprFast(x_2); -x_25 = lean_string_append(x_3, x_24); -lean_dec_ref(x_24); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_23); -lean_ctor_set(x_26, 1, x_25); -return x_26; -} +if (lean_obj_tag(x_1) == 10) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_25 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_25); +lean_dec_ref(x_1); +x_26 = lean_array_get_size(x_25); +x_27 = lean_unsigned_to_nat(0u); +x_28 = lean_box(0); +x_29 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructIncDecFn_spec__0___redArg(x_26, x_25, x_3, x_14, x_28, x_27, x_28, x_15, x_16); +lean_dec_ref(x_15); +lean_dec_ref(x_25); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +lean_dec_ref(x_29); +x_6 = x_30; +goto block_13; } else { -lean_object* x_27; uint8_t x_28; -lean_dec(x_1); -x_27 = lean_cstr_to_nat("4294967296"); -x_28 = lean_nat_dec_lt(x_2, x_27); -if (x_28 == 0) -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_29 = l_Lean_IR_EmitC_emitNumLit___redArg___closed__3; -x_30 = lean_string_append(x_3, x_29); -x_31 = l_Nat_reprFast(x_2); -x_32 = lean_string_append(x_30, x_31); -lean_dec_ref(x_31); -x_33 = l_Lean_IR_EmitC_emitNumLit___redArg___closed__4; -x_34 = lean_box(0); -x_35 = lean_string_append(x_32, x_33); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; +return x_29; +} } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_37 = l_Lean_IR_EmitC_emitNumLit___redArg___closed__5; -x_38 = lean_string_append(x_3, x_37); -x_39 = l_Nat_reprFast(x_2); -x_40 = lean_string_append(x_38, x_39); -lean_dec_ref(x_39); -x_41 = l_Lean_IR_EmitC_emitNumLit___redArg___closed__6; -x_42 = lean_box(0); -x_43 = lean_string_append(x_40, x_41); -x_44 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -return x_44; +lean_dec_ref(x_15); +lean_dec(x_1); +x_6 = x_16; +goto block_13; } } } +block_50: +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_33 = l_Lean_IR_EmitC_emitStructIncDecFn___closed__1; +x_34 = lean_string_append(x_5, x_33); +x_35 = lean_string_append(x_34, x_32); +lean_dec_ref(x_32); +lean_inc(x_2); +x_36 = l_Nat_reprFast(x_2); +x_37 = lean_string_append(x_35, x_36); +lean_dec_ref(x_36); +x_38 = l_Lean_IR_EmitC_emitFnDeclAux___closed__2; +x_39 = lean_string_append(x_37, x_38); +x_40 = l_Lean_IR_EmitC_structType(x_2); +x_41 = lean_string_append(x_39, x_40); +lean_dec_ref(x_40); +if (x_3 == 0) +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_42 = l_Lean_IR_EmitC_emitStructIncDecFn___closed__2; +x_43 = lean_string_append(x_41, x_42); +x_44 = l_Lean_IR_EmitC_emitLn___closed__0; +x_45 = lean_string_append(x_43, x_44); +x_15 = x_4; +x_16 = x_45; +goto block_31; } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitNumLit(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +else { -lean_object* x_5; -x_5 = l_Lean_IR_EmitC_emitNumLit___redArg(x_1, x_2, x_4); -return x_5; +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_46 = l_Lean_IR_EmitC_emitStructIncDecFn___closed__3; +x_47 = lean_string_append(x_41, x_46); +x_48 = l_Lean_IR_EmitC_emitLn___closed__0; +x_49 = lean_string_append(x_47, x_48); +x_15 = x_4; +x_16 = x_49; +goto block_31; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitNumLit___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_IR_EmitC_emitNumLit(x_1, x_2, x_3, x_4); -lean_dec_ref(x_3); -return x_5; } } -static lean_object* _init_l_Lean_IR_EmitC_emitLit___redArg___closed__0() { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructIncDecFn_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_mk_string_unchecked(", 25, 25); -return x_1; +uint8_t x_13; uint8_t x_14; lean_object* x_15; +x_13 = lean_unbox(x_3); +x_14 = lean_unbox(x_4); +x_15 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructIncDecFn_spec__0(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec_ref(x_11); +lean_dec_ref(x_2); +lean_dec(x_1); +return x_15; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLit___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructIncDecFn_spec__0_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_5; -x_5 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_4); -if (lean_obj_tag(x_3) == 0) -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_6 = lean_ctor_get(x_5, 1); -lean_inc(x_6); -lean_dec_ref(x_5); -x_7 = lean_ctor_get(x_3, 0); -lean_inc(x_7); -lean_dec_ref(x_3); -x_8 = l_Lean_IR_EmitC_emitNumLit___redArg(x_2, x_7, x_6); -x_9 = !lean_is_exclusive(x_8); -if (x_9 == 0) -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_10 = lean_ctor_get(x_8, 1); -x_11 = lean_ctor_get(x_8, 0); -lean_dec(x_11); -x_12 = l_Lean_IR_EmitC_emitFnDeclAux___closed__0; -x_13 = lean_string_append(x_10, x_12); -x_14 = l_Lean_IR_EmitC_emitLn___closed__0; -x_15 = lean_box(0); -x_16 = lean_string_append(x_13, x_14); -lean_ctor_set(x_8, 1, x_16); -lean_ctor_set(x_8, 0, x_15); -return x_8; -} -else -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_17 = lean_ctor_get(x_8, 1); -lean_inc(x_17); -lean_dec(x_8); -x_18 = l_Lean_IR_EmitC_emitFnDeclAux___closed__0; -x_19 = lean_string_append(x_17, x_18); -x_20 = l_Lean_IR_EmitC_emitLn___closed__0; -x_21 = lean_box(0); -x_22 = lean_string_append(x_19, x_20); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -return x_23; +uint8_t x_13; uint8_t x_14; lean_object* x_15; +x_13 = lean_unbox(x_4); +x_14 = lean_unbox(x_5); +x_15 = l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructIncDecFn_spec__0_spec__0(x_1, x_2, x_3, x_13, x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec_ref(x_11); +lean_dec_ref(x_2); +lean_dec(x_1); +return x_15; } } -else -{ -uint8_t x_24; -lean_dec(x_2); -x_24 = !lean_is_exclusive(x_5); -if (x_24 == 0) +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructIncDecFn_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_25 = lean_ctor_get(x_5, 1); -x_26 = lean_ctor_get(x_5, 0); -lean_dec(x_26); -x_27 = lean_ctor_get(x_3, 0); -lean_inc_ref(x_27); -lean_dec_ref(x_3); -x_28 = l_Lean_IR_EmitC_emitLit___redArg___closed__0; -x_29 = lean_string_append(x_25, x_28); -lean_inc_ref(x_27); -x_30 = l_Lean_IR_EmitC_quoteString(x_27); -x_31 = lean_string_append(x_29, x_30); -lean_dec_ref(x_30); -x_32 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___closed__0; -x_33 = lean_string_append(x_31, x_32); -x_34 = lean_string_utf8_byte_size(x_27); -x_35 = l_Nat_reprFast(x_34); -x_36 = lean_string_append(x_33, x_35); -lean_dec_ref(x_35); -x_37 = lean_string_append(x_36, x_32); -x_38 = lean_string_length(x_27); -lean_dec_ref(x_27); -x_39 = l_Nat_reprFast(x_38); -x_40 = lean_string_append(x_37, x_39); -lean_dec_ref(x_39); -x_41 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_42 = lean_string_append(x_40, x_41); -x_43 = l_Lean_IR_EmitC_emitLn___closed__0; -x_44 = lean_box(0); -x_45 = lean_string_append(x_42, x_43); -lean_ctor_set(x_5, 1, x_45); -lean_ctor_set(x_5, 0, x_44); -return x_5; +uint8_t x_10; uint8_t x_11; lean_object* x_12; +x_10 = lean_unbox(x_3); +x_11 = lean_unbox(x_4); +x_12 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructIncDecFn_spec__0___redArg(x_1, x_2, x_10, x_11, x_5, x_6, x_7, x_8, x_9); +lean_dec_ref(x_8); +lean_dec_ref(x_2); +lean_dec(x_1); +return x_12; } -else +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructIncDecFn_spec__0_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_46 = lean_ctor_get(x_5, 1); -lean_inc(x_46); -lean_dec(x_5); -x_47 = lean_ctor_get(x_3, 0); -lean_inc_ref(x_47); -lean_dec_ref(x_3); -x_48 = l_Lean_IR_EmitC_emitLit___redArg___closed__0; -x_49 = lean_string_append(x_46, x_48); -lean_inc_ref(x_47); -x_50 = l_Lean_IR_EmitC_quoteString(x_47); -x_51 = lean_string_append(x_49, x_50); -lean_dec_ref(x_50); -x_52 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___closed__0; -x_53 = lean_string_append(x_51, x_52); -x_54 = lean_string_utf8_byte_size(x_47); -x_55 = l_Nat_reprFast(x_54); -x_56 = lean_string_append(x_53, x_55); -lean_dec_ref(x_55); -x_57 = lean_string_append(x_56, x_52); -x_58 = lean_string_length(x_47); -lean_dec_ref(x_47); -x_59 = l_Nat_reprFast(x_58); -x_60 = lean_string_append(x_57, x_59); -lean_dec_ref(x_59); -x_61 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; -x_62 = lean_string_append(x_60, x_61); -x_63 = l_Lean_IR_EmitC_emitLn___closed__0; -x_64 = lean_box(0); -x_65 = lean_string_append(x_62, x_63); -x_66 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_66, 0, x_64); -lean_ctor_set(x_66, 1, x_65); -return x_66; +uint8_t x_10; uint8_t x_11; lean_object* x_12; +x_10 = lean_unbox(x_4); +x_11 = lean_unbox(x_5); +x_12 = l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructIncDecFn_spec__0_spec__0___redArg(x_1, x_2, x_3, x_10, x_11, x_6, x_7, x_8, x_9); +lean_dec_ref(x_8); +lean_dec_ref(x_2); +lean_dec(x_1); +return x_12; } } +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructIncDecFn___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; lean_object* x_7; +x_6 = lean_unbox(x_3); +x_7 = l_Lean_IR_EmitC_emitStructIncDecFn(x_1, x_2, x_6, x_4, x_5); +return x_7; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLit(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_IR_EmitC_emitReshapeFn(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { +uint8_t x_5; +x_5 = l_Lean_IR_IRType_isObj(x_1); +if (x_5 == 0) +{ lean_object* x_6; -x_6 = l_Lean_IR_EmitC_emitLit___redArg(x_1, x_2, x_3, x_5); +x_6 = l_Lean_IR_EmitC_emitBoxFn(x_1, x_2, x_3, x_4); return x_6; } +else +{ +lean_object* x_7; +lean_dec(x_1); +x_7 = l_Lean_IR_EmitC_emitUnboxFn(x_2, x_3, x_4); +return x_7; +} } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLit___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_IR_EmitC_emitReshapeFn___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_6; -x_6 = l_Lean_IR_EmitC_emitLit(x_1, x_2, x_3, x_4, x_5); -lean_dec_ref(x_4); -return x_6; +lean_object* x_5; +x_5 = l_Lean_IR_EmitC_emitReshapeFn(x_1, x_2, x_3, x_4); +lean_dec_ref(x_3); +lean_dec(x_2); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitVDecl(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT uint8_t l_Lean_IR_EmitC_compatibleReshape(lean_object* x_1, lean_object* x_2) { _start: { -switch (lean_obj_tag(x_3)) { -case 0: +switch (lean_obj_tag(x_1)) { +case 10: { -lean_object* x_6; lean_object* x_7; lean_object* x_8; -lean_dec(x_2); -x_6 = lean_ctor_get(x_3, 0); -lean_inc_ref(x_6); -x_7 = lean_ctor_get(x_3, 1); -lean_inc_ref(x_7); -lean_dec_ref(x_3); -x_8 = l_Lean_IR_EmitC_emitCtor(x_1, x_6, x_7, x_4, x_5); -lean_dec_ref(x_4); -lean_dec_ref(x_7); -return x_8; +switch (lean_obj_tag(x_2)) { +case 10: +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_12 = lean_ctor_get(x_1, 1); +x_13 = lean_ctor_get(x_2, 1); +x_14 = lean_array_get_size(x_12); +x_15 = lean_array_get_size(x_13); +x_16 = lean_nat_dec_eq(x_14, x_15); +if (x_16 == 0) +{ +return x_16; } -case 1: +else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; -lean_dec(x_2); -x_9 = lean_ctor_get(x_3, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_3, 1); -lean_inc(x_10); -lean_dec_ref(x_3); -x_11 = l_Lean_IR_EmitC_emitReset(x_1, x_9, x_10, x_4, x_5); -lean_dec_ref(x_4); -return x_11; +uint8_t x_17; +x_17 = l_Array_isEqvAux___at___00Lean_IR_EmitC_compatibleReshape_spec__0___redArg(x_12, x_13, x_14); +return x_17; } -case 2: +} +case 11: { -lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; -lean_dec(x_2); -x_12 = lean_ctor_get(x_3, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_3, 1); -lean_inc_ref(x_13); -x_14 = lean_ctor_get_uint8(x_3, sizeof(void*)*3); -x_15 = lean_ctor_get(x_3, 2); -lean_inc_ref(x_15); -lean_dec_ref(x_3); -x_16 = l_Lean_IR_EmitC_emitReuse(x_1, x_12, x_13, x_14, x_15, x_4, x_5); -lean_dec_ref(x_4); -lean_dec_ref(x_15); -return x_16; +uint8_t x_18; +x_18 = 0; +return x_18; } -case 3: +default: { -lean_object* x_17; lean_object* x_18; lean_object* x_19; -lean_dec_ref(x_4); -lean_dec(x_2); -x_17 = lean_ctor_get(x_3, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_3, 1); -lean_inc(x_18); -lean_dec_ref(x_3); -x_19 = l_Lean_IR_EmitC_emitProj___redArg(x_1, x_17, x_18, x_5); +goto block_11; +} +} +} +case 11: +{ +if (lean_obj_tag(x_2) == 10) +{ +uint8_t x_19; +x_19 = 0; return x_19; } -case 4: +else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -lean_dec_ref(x_4); -lean_dec(x_2); -x_20 = lean_ctor_get(x_3, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_3, 1); -lean_inc(x_21); -lean_dec_ref(x_3); -x_22 = l_Lean_IR_EmitC_emitUProj___redArg(x_1, x_20, x_21, x_5); -return x_22; +goto block_11; } -case 5: +} +default: { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -lean_dec_ref(x_4); -x_23 = lean_ctor_get(x_3, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_3, 1); -lean_inc(x_24); -x_25 = lean_ctor_get(x_3, 2); -lean_inc(x_25); -lean_dec_ref(x_3); -x_26 = l_Lean_IR_EmitC_emitSProj___redArg(x_1, x_2, x_23, x_24, x_25, x_5); -lean_dec(x_2); -return x_26; +goto block_11; } -case 6: +} +block_7: { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -lean_dec(x_2); -x_27 = lean_ctor_get(x_3, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_3, 1); -lean_inc_ref(x_28); -lean_dec_ref(x_3); -x_29 = l_Lean_IR_EmitC_emitFullApp(x_1, x_27, x_28, x_4, x_5); -return x_29; +uint8_t x_3; +x_3 = l_Lean_IR_IRType_isStruct(x_1); +if (x_3 == 0) +{ +uint8_t x_4; +x_4 = 1; +return x_4; } -case 7: +else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -lean_dec(x_2); -x_30 = lean_ctor_get(x_3, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_3, 1); -lean_inc_ref(x_31); -lean_dec_ref(x_3); -x_32 = l_Lean_IR_EmitC_emitPartialApp(x_1, x_30, x_31, x_4, x_5); -lean_dec_ref(x_31); -return x_32; +uint8_t x_5; +x_5 = l_Lean_IR_IRType_isScalar(x_2); +if (x_5 == 0) +{ +return x_3; } -case 8: +else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -lean_dec(x_2); -x_33 = lean_ctor_get(x_3, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_3, 1); -lean_inc_ref(x_34); -lean_dec_ref(x_3); -x_35 = l_Lean_IR_EmitC_emitApp(x_1, x_33, x_34, x_4, x_5); -lean_dec_ref(x_4); -lean_dec_ref(x_34); -return x_35; +uint8_t x_6; +x_6 = 0; +return x_6; +} +} +} +block_11: +{ +uint8_t x_8; +x_8 = l_Lean_IR_IRType_isScalar(x_1); +if (x_8 == 0) +{ +goto block_7; +} +else +{ +uint8_t x_9; +x_9 = l_Lean_IR_IRType_isStruct(x_2); +if (x_9 == 0) +{ +goto block_7; } -case 9: +else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; -lean_dec_ref(x_4); -lean_dec(x_2); -x_36 = lean_ctor_get(x_3, 0); -lean_inc(x_36); -x_37 = lean_ctor_get(x_3, 1); -lean_inc(x_37); -lean_dec_ref(x_3); -x_38 = l_Lean_IR_EmitC_emitBox___redArg(x_1, x_37, x_36, x_5); -lean_dec(x_36); -return x_38; +uint8_t x_10; +x_10 = 0; +return x_10; } -case 10: +} +} +} +} +LEAN_EXPORT uint8_t l_Array_isEqvAux___at___00Lean_IR_EmitC_compatibleReshape_spec__0___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_39; lean_object* x_40; -lean_dec_ref(x_4); -x_39 = lean_ctor_get(x_3, 0); -lean_inc(x_39); -lean_dec_ref(x_3); -x_40 = l_Lean_IR_EmitC_emitUnbox___redArg(x_1, x_2, x_39, x_5); -lean_dec(x_2); -return x_40; +lean_object* x_4; uint8_t x_5; +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_nat_dec_eq(x_3, x_4); +if (x_5 == 1) +{ +lean_dec(x_3); +return x_5; } -case 11: +else { -lean_object* x_41; lean_object* x_42; -lean_dec_ref(x_4); -x_41 = lean_ctor_get(x_3, 0); -lean_inc_ref(x_41); -lean_dec_ref(x_3); -x_42 = l_Lean_IR_EmitC_emitLit___redArg(x_1, x_2, x_41, x_5); -return x_42; +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_6 = lean_unsigned_to_nat(1u); +x_7 = lean_nat_sub(x_3, x_6); +lean_dec(x_3); +x_8 = lean_array_fget_borrowed(x_1, x_7); +x_9 = lean_array_fget_borrowed(x_2, x_7); +x_10 = l_Lean_IR_EmitC_compatibleReshape(x_8, x_9); +if (x_10 == 0) +{ +lean_dec(x_7); +return x_10; } -default: +else { -lean_object* x_43; lean_object* x_44; -lean_dec_ref(x_4); -lean_dec(x_2); -x_43 = lean_ctor_get(x_3, 0); -lean_inc(x_43); -lean_dec_ref(x_3); -x_44 = l_Lean_IR_EmitC_emitIsShared___redArg(x_1, x_43, x_5); -return x_44; +x_3 = x_7; +goto _start; } } } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_isTailCall(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT uint8_t l_Array_isEqvAux___at___00Lean_IR_EmitC_compatibleReshape_spec__0(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; -if (lean_obj_tag(x_2) == 6) +uint8_t x_6; +x_6 = l_Array_isEqvAux___at___00Lean_IR_EmitC_compatibleReshape_spec__0___redArg(x_1, x_2, x_4); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Array_isEqvAux___at___00Lean_IR_EmitC_compatibleReshape_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -if (lean_obj_tag(x_3) == 10) +uint8_t x_6; lean_object* x_7; +x_6 = l_Array_isEqvAux___at___00Lean_IR_EmitC_compatibleReshape_spec__0(x_1, x_2, x_3, x_4, x_5); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_7 = lean_box(x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Array_isEqvAux___at___00Lean_IR_EmitC_compatibleReshape_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_11; -x_11 = lean_ctor_get(x_3, 0); -if (lean_obj_tag(x_11) == 0) +uint8_t x_4; lean_object* x_5; +x_4 = l_Array_isEqvAux___at___00Lean_IR_EmitC_compatibleReshape_spec__0___redArg(x_1, x_2, x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_5 = lean_box(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_compatibleReshape___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -uint8_t x_12; -x_12 = !lean_is_exclusive(x_2); -if (x_12 == 0) +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_IR_EmitC_compatibleReshape(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__0() { +_start: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_13 = lean_ctor_get(x_2, 0); -x_14 = lean_ctor_get(x_2, 1); -lean_dec(x_14); -x_15 = lean_ctor_get(x_11, 0); -x_16 = lean_ctor_get(x_4, 3); -x_17 = lean_name_eq(x_13, x_16); -lean_dec(x_13); -if (x_17 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("y.i", 3, 3); +return x_1; +} +} +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__1() { +_start: { -lean_object* x_18; -x_18 = lean_box(x_17); -lean_ctor_set_tag(x_2, 0); -lean_ctor_set(x_2, 1, x_5); -lean_ctor_set(x_2, 0, x_18); -return x_2; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("(x.i", 4, 4); +return x_1; } -else +} +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___lam__0___closed__0() { +_start: { -uint8_t x_19; lean_object* x_20; -x_19 = l_Lean_IR_instBEqVarId_beq(x_1, x_15); -x_20 = lean_box(x_19); -lean_ctor_set_tag(x_2, 0); -lean_ctor_set(x_2, 1, x_5); -lean_ctor_set(x_2, 0, x_20); -return x_2; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_dec(x.i", 12, 12); +return x_1; } } -else +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_21 = lean_ctor_get(x_2, 0); -lean_inc(x_21); -lean_dec(x_2); -x_22 = lean_ctor_get(x_11, 0); -x_23 = lean_ctor_get(x_4, 3); -x_24 = lean_name_eq(x_21, x_23); -lean_dec(x_21); -if (x_24 == 0) +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_7 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___lam__0___closed__0; +x_8 = lean_string_append(x_6, x_7); +x_9 = lean_string_append(x_8, x_1); +x_10 = lean_string_append(x_9, x_2); +x_11 = l_Lean_IR_EmitC_emitLn___closed__0; +x_12 = lean_string_append(x_10, x_11); +x_13 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_13, 0, x_3); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_12); +return x_14; +} +} +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__2() { +_start: { -lean_object* x_25; lean_object* x_26; -x_25 = lean_box(x_24); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_5); -return x_26; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("1", 1, 1); +return x_1; } -else +} +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__3() { +_start: { -uint8_t x_27; lean_object* x_28; lean_object* x_29; -x_27 = l_Lean_IR_instBEqVarId_beq(x_1, x_22); -x_28 = lean_box(x_27); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_5); -return x_29; +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" = x.i", 6, 6); +return x_1; +} +} +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" = lean_box(0);", 15, 15); +return x_1; } } +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; lean_object* x_18; lean_object* x_33; uint8_t x_34; lean_object* x_74; uint8_t x_75; uint8_t x_85; uint8_t x_88; +x_88 = lean_nat_dec_lt(x_8, x_6); +if (x_88 == 0) +{ +lean_object* x_89; +lean_dec(x_8); +lean_dec(x_2); +x_89 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_89, 0, x_9); +lean_ctor_set(x_89, 1, x_11); +return x_89; } else { -lean_dec_ref(x_2); -x_6 = x_5; -goto block_10; +lean_object* x_90; +lean_inc(x_2); +x_90 = lean_array_get_borrowed(x_2, x_3, x_8); +switch (lean_obj_tag(x_90)) { +case 6: +{ +x_85 = x_5; +goto block_87; } +case 13: +{ +x_85 = x_5; +goto block_87; } -else +default: { -lean_dec_ref(x_2); -x_6 = x_5; -goto block_10; +x_85 = x_7; +goto block_87; } } -else +} +block_17: { -lean_dec_ref(x_2); -x_6 = x_5; -goto block_10; +lean_object* x_14; lean_object* x_15; +x_14 = lean_unsigned_to_nat(1u); +x_15 = lean_nat_add(x_8, x_14); +lean_dec(x_8); +x_8 = x_15; +x_9 = x_12; +x_11 = x_13; +goto _start; } -block_10: +block_32: { -uint8_t x_7; lean_object* x_8; lean_object* x_9; -x_7 = 0; -x_8 = lean_box(x_7); -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_8); -lean_ctor_set(x_9, 1, x_6); -return x_9; +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +if (lean_obj_tag(x_19) == 0) +{ +uint8_t x_20; +lean_dec(x_8); +lean_dec(x_2); +x_20 = !lean_is_exclusive(x_18); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_18, 0); +lean_dec(x_21); +x_22 = lean_ctor_get(x_19, 0); +lean_inc(x_22); +lean_dec_ref(x_19); +lean_ctor_set(x_18, 0, x_22); +return x_18; } +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_18, 1); +lean_inc(x_23); +lean_dec(x_18); +x_24 = lean_ctor_get(x_19, 0); +lean_inc(x_24); +lean_dec_ref(x_19); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_23); +return x_25; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_isTailCall___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +else { -lean_object* x_6; -x_6 = l_Lean_IR_EmitC_isTailCall(x_1, x_2, x_3, x_4, x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec(x_1); -return x_6; +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_18, 1); +lean_inc(x_26); +lean_dec_ref(x_18); +x_27 = lean_ctor_get(x_19, 0); +lean_inc(x_27); +lean_dec_ref(x_19); +x_12 = x_27; +x_13 = x_26; +goto block_17; } } -LEAN_EXPORT uint8_t l_Lean_IR_EmitC_paramEqArg(lean_object* x_1, lean_object* x_2) { -_start: +else { -if (lean_obj_tag(x_2) == 0) +uint8_t x_28; +lean_dec(x_8); +lean_dec(x_2); +x_28 = !lean_is_exclusive(x_18); +if (x_28 == 0) { -lean_object* x_3; lean_object* x_4; uint8_t x_5; -x_3 = lean_ctor_get(x_2, 0); -x_4 = lean_ctor_get(x_1, 0); -x_5 = l_Lean_IR_instBEqVarId_beq(x_4, x_3); -return x_5; +return x_18; } else { -uint8_t x_6; -x_6 = 0; -return x_6; +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_18, 0); +x_30 = lean_ctor_get(x_18, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_18); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; } } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_paramEqArg___boxed(lean_object* x_1, lean_object* x_2) { -_start: +block_73: { -uint8_t x_3; lean_object* x_4; -x_3 = l_Lean_IR_EmitC_paramEqArg(x_1, x_2); -lean_dec(x_2); -lean_dec_ref(x_1); -x_4 = lean_box(x_3); -return x_4; -} +lean_object* x_35; uint8_t x_36; +lean_inc(x_2); +x_35 = lean_array_get_borrowed(x_2, x_3, x_8); +x_36 = l_Lean_IR_IRType_compatibleWith(x_33, x_35, x_34); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_37 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__0; +x_38 = lean_string_append(x_11, x_37); +lean_inc(x_8); +x_39 = l_Nat_reprFast(x_8); +x_40 = lean_string_append(x_38, x_39); +x_41 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__1; +x_42 = lean_string_append(x_40, x_41); +lean_inc(x_33); +x_43 = l_Lean_IR_EmitC_emitReshapeFn(x_33, x_35, x_10, x_42); +if (lean_obj_tag(x_43) == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; +x_44 = lean_ctor_get(x_43, 1); +lean_inc(x_44); +lean_dec_ref(x_43); +x_45 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__1; +x_46 = lean_string_append(x_44, x_45); +x_47 = lean_string_append(x_46, x_39); +x_48 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_49 = lean_string_append(x_47, x_48); +x_50 = l_Lean_IR_EmitC_emitLn___closed__0; +x_51 = lean_string_append(x_49, x_50); +x_52 = l_Lean_IR_IRType_isObj(x_33); +lean_dec(x_33); +if (x_52 == 0) +{ +lean_dec_ref(x_39); +x_12 = x_4; +x_13 = x_51; +goto block_17; } -LEAN_EXPORT uint8_t l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__0___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +else { -lean_object* x_7; uint8_t x_8; -x_7 = lean_unsigned_to_nat(0u); -x_8 = lean_nat_dec_eq(x_6, x_7); -if (x_8 == 1) +uint8_t x_53; +x_53 = l_Lean_IR_EmitC_needsRC(x_35); +if (x_53 == 0) { -uint8_t x_9; -lean_dec(x_6); -lean_dec(x_2); -x_9 = 0; -return x_9; +lean_object* x_54; +x_54 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___lam__0(x_39, x_48, x_4, x_4, x_10, x_51); +lean_dec_ref(x_39); +x_18 = x_54; +goto block_32; } else { -lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_10 = lean_nat_sub(x_5, x_6); -x_11 = lean_nat_add(x_1, x_10); -lean_dec(x_10); -lean_inc(x_2); -x_12 = lean_array_get_borrowed(x_2, x_3, x_11); -lean_dec(x_11); -x_13 = l_Lean_IR_EmitC_paramEqArg(x_4, x_12); -if (x_13 == 0) +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_55 = lean_string_append(x_37, x_39); +x_56 = lean_unsigned_to_nat(1u); +x_57 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__2; +x_58 = l_Lean_IR_EmitC_emitIncOfType(x_55, x_35, x_56, x_5, x_57, x_10, x_51); +lean_dec_ref(x_55); +if (lean_obj_tag(x_58) == 0) { -lean_object* x_14; lean_object* x_15; -x_14 = lean_unsigned_to_nat(1u); -x_15 = lean_nat_sub(x_6, x_14); -lean_dec(x_6); -x_6 = x_15; -goto _start; +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_58, 0); +lean_inc(x_59); +x_60 = lean_ctor_get(x_58, 1); +lean_inc(x_60); +lean_dec_ref(x_58); +x_61 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___lam__0(x_39, x_48, x_4, x_59, x_10, x_60); +lean_dec_ref(x_39); +x_18 = x_61; +goto block_32; } else { -lean_dec(x_6); +lean_dec_ref(x_39); +lean_dec(x_8); lean_dec(x_2); -return x_13; +return x_58; } } } } -LEAN_EXPORT uint8_t l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: +else { -uint8_t x_8; -x_8 = l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__0___redArg(x_1, x_2, x_3, x_4, x_5, x_6); -return x_8; +lean_dec_ref(x_39); +lean_dec(x_33); +lean_dec(x_8); +lean_dec(x_2); +return x_43; } } -LEAN_EXPORT uint8_t l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00__private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1_spec__1___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; uint8_t x_8; -x_7 = lean_unsigned_to_nat(0u); -x_8 = lean_nat_dec_eq(x_6, x_7); -if (x_8 == 1) +else { -uint8_t x_9; -lean_dec(x_6); -lean_dec(x_3); -x_9 = 0; -return x_9; +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +lean_dec(x_33); +x_62 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__0; +x_63 = lean_string_append(x_11, x_62); +lean_inc(x_8); +x_64 = l_Nat_reprFast(x_8); +x_65 = lean_string_append(x_63, x_64); +x_66 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__3; +x_67 = lean_string_append(x_65, x_66); +x_68 = lean_string_append(x_67, x_64); +lean_dec_ref(x_64); +x_69 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_70 = lean_string_append(x_68, x_69); +x_71 = l_Lean_IR_EmitC_emitLn___closed__0; +x_72 = lean_string_append(x_70, x_71); +x_12 = x_4; +x_13 = x_72; +goto block_17; } -else +} +block_84: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_10 = lean_nat_sub(x_5, x_6); -x_11 = lean_array_fget_borrowed(x_1, x_10); -x_12 = lean_unsigned_to_nat(1u); -x_13 = lean_nat_add(x_10, x_12); -lean_dec(x_10); -x_14 = lean_nat_sub(x_2, x_13); -lean_inc(x_14); -lean_inc(x_3); -x_15 = l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__0___redArg(x_13, x_3, x_4, x_11, x_14, x_14); -lean_dec(x_14); -lean_dec(x_13); -if (x_15 == 0) +if (x_75 == 0) { -lean_object* x_16; -x_16 = lean_nat_sub(x_6, x_12); -lean_dec(x_6); -x_6 = x_16; -goto _start; +x_33 = x_74; +x_34 = x_75; +goto block_73; } else { -lean_dec(x_6); -lean_dec(x_3); -return x_15; -} -} +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +lean_dec(x_74); +x_76 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__0; +x_77 = lean_string_append(x_11, x_76); +lean_inc(x_8); +x_78 = l_Nat_reprFast(x_8); +x_79 = lean_string_append(x_77, x_78); +lean_dec_ref(x_78); +x_80 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__4; +x_81 = lean_string_append(x_79, x_80); +x_82 = l_Lean_IR_EmitC_emitLn___closed__0; +x_83 = lean_string_append(x_81, x_82); +x_12 = x_4; +x_13 = x_83; +goto block_17; } } -LEAN_EXPORT uint8_t l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +block_87: { -lean_object* x_7; uint8_t x_8; -x_7 = lean_unsigned_to_nat(0u); -x_8 = lean_nat_dec_eq(x_6, x_7); -if (x_8 == 1) +if (x_85 == 0) { -uint8_t x_9; -lean_dec(x_2); -x_9 = 0; -return x_9; +lean_object* x_86; +x_86 = lean_array_fget_borrowed(x_1, x_8); +switch (lean_obj_tag(x_86)) { +case 6: +{ +x_74 = x_86; +x_75 = x_5; +goto block_84; } -else +case 13: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_10 = lean_nat_sub(x_5, x_6); -x_11 = lean_array_fget_borrowed(x_1, x_10); -x_12 = lean_unsigned_to_nat(1u); -x_13 = lean_nat_add(x_10, x_12); -lean_dec(x_10); -x_14 = lean_nat_sub(x_4, x_13); -lean_inc(x_14); -lean_inc(x_2); -x_15 = l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__0___redArg(x_13, x_2, x_3, x_11, x_14, x_14); -lean_dec(x_14); -lean_dec(x_13); -if (x_15 == 0) +x_74 = x_86; +x_75 = x_5; +goto block_84; +} +default: { -lean_object* x_16; uint8_t x_17; -x_16 = lean_nat_sub(x_6, x_12); -x_17 = l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00__private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1_spec__1___redArg(x_1, x_4, x_2, x_3, x_5, x_16); -return x_17; +lean_inc(x_86); +x_33 = x_86; +x_34 = x_85; +goto block_73; +} +} } else { -lean_dec(x_2); -return x_15; +x_12 = x_4; +x_13 = x_11; +goto block_17; } } } } -LEAN_EXPORT uint8_t l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_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 lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -uint8_t x_8; -x_8 = l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1___redArg(x_1, x_2, x_3, x_4, x_5, x_6); -return x_8; +lean_object* x_15; +x_15 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_10, x_11, x_13, x_14); +return x_15; } } -LEAN_EXPORT uint8_t l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00__private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1_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_Lean_IR_EmitC_emitStructReshapeFn___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -uint8_t x_8; -x_8 = l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00__private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1_spec__1___redArg(x_1, x_2, x_3, x_4, x_5, x_6); -return x_8; +lean_object* x_5; lean_object* x_6; uint8_t x_7; +lean_inc(x_1); +x_5 = lean_array_get_borrowed(x_1, x_2, x_4); +x_6 = lean_array_get_borrowed(x_1, x_3, x_4); +x_7 = l_Lean_IR_EmitC_compatibleReshape(x_5, x_6); +return x_7; } } -static lean_object* _init_l_Lean_IR_EmitC_overwriteParam___closed__0() { +static lean_object* _init_l_Lean_IR_EmitC_emitStructReshapeFn___lam__1___closed__0() { _start: { lean_object* x_1; -x_1 = l_Lean_IR_instInhabitedArg_default; +x_1 = lean_mk_string_unchecked("y.cs.c", 6, 6); return x_1; } } -LEAN_EXPORT uint8_t l_Lean_IR_EmitC_overwriteParam(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_IR_EmitC_emitStructReshapeFn___lam__1___closed__1() { _start: { -lean_object* x_3; lean_object* x_4; uint8_t x_5; -x_3 = l_Lean_IR_EmitC_overwriteParam___closed__0; -x_4 = lean_array_get_size(x_1); -x_5 = l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1___redArg(x_1, x_3, x_2, x_4, x_4, x_4); -return x_5; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("(x.cs.c", 7, 7); +return x_1; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__0___boxed(lean_object* x_1, lean_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_IR_EmitC_emitStructReshapeFn___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -uint8_t x_8; lean_object* x_9; -x_8 = l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -lean_dec(x_1); -x_9 = lean_box(x_8); -return x_9; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_10 = l_Lean_IR_EmitC_emitStructReshapeFn___lam__1___closed__0; +x_11 = lean_string_append(x_9, x_10); +lean_inc(x_6); +x_12 = l_Nat_reprFast(x_6); +x_13 = lean_string_append(x_11, x_12); +x_14 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__1; +x_15 = lean_string_append(x_13, x_14); +x_16 = lean_array_fget_borrowed(x_1, x_6); +x_17 = lean_array_get_borrowed(x_2, x_3, x_6); +lean_dec(x_6); +x_18 = l_Lean_IR_IRType_compatibleWith(x_16, x_17, x_4); +if (x_18 == 0) +{ +lean_object* x_19; uint8_t x_20; +lean_inc(x_16); +x_19 = l_Lean_IR_EmitC_emitReshapeFn(x_16, x_17, x_8, x_15); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_21 = lean_ctor_get(x_19, 1); +x_22 = lean_ctor_get(x_19, 0); +lean_dec(x_22); +x_23 = l_Lean_IR_EmitC_emitStructReshapeFn___lam__1___closed__1; +x_24 = lean_string_append(x_21, x_23); +x_25 = lean_string_append(x_24, x_12); +lean_dec_ref(x_12); +x_26 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_27 = lean_string_append(x_25, x_26); +x_28 = lean_box(0); +x_29 = lean_string_append(x_27, x_5); +lean_ctor_set(x_19, 1, x_29); +lean_ctor_set(x_19, 0, x_28); +return x_19; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_30 = lean_ctor_get(x_19, 1); +lean_inc(x_30); +lean_dec(x_19); +x_31 = l_Lean_IR_EmitC_emitStructReshapeFn___lam__1___closed__1; +x_32 = lean_string_append(x_30, x_31); +x_33 = lean_string_append(x_32, x_12); +lean_dec_ref(x_12); +x_34 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_35 = lean_string_append(x_33, x_34); +x_36 = lean_box(0); +x_37 = lean_string_append(x_35, x_5); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +return x_38; +} +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_39 = l_Lean_IR_EmitC_emitStructIncDecFn___lam__0___closed__0; +x_40 = lean_string_append(x_15, x_39); +x_41 = lean_string_append(x_40, x_12); +lean_dec_ref(x_12); +x_42 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2; +x_43 = lean_string_append(x_41, x_42); +x_44 = lean_box(0); +x_45 = lean_string_append(x_43, x_5); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitStructReshapeFn___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("return y;", 9, 9); +return x_1; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +static lean_object* _init_l_Lean_IR_EmitC_emitStructReshapeFn___closed__1() { _start: { -uint8_t x_8; lean_object* x_9; -x_8 = l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -lean_dec_ref(x_1); -x_9 = lean_box(x_8); -return x_9; +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" y;", 3, 3); +return x_1; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00__private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +static lean_object* _init_l_Lean_IR_EmitC_emitStructReshapeFn___closed__2() { _start: { -uint8_t x_8; lean_object* x_9; -x_8 = l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00__private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1_spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_2); -lean_dec_ref(x_1); -x_9 = lean_box(x_8); -return x_9; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("y.tag = x.tag;", 14, 14); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_overwriteParam___boxed(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_IR_EmitC_emitStructReshapeFn___closed__3() { _start: { -uint8_t x_3; lean_object* x_4; -x_3 = l_Lean_IR_EmitC_overwriteParam(x_1, x_2); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_4 = lean_box(x_3); -return x_4; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("memcpy(y.u, x.u, sizeof(size_t)*", 32, 32); +return x_1; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +static lean_object* _init_l_Lean_IR_EmitC_emitStructReshapeFn___closed__4() { _start: { -uint8_t x_7; lean_object* x_8; -x_7 = l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__0___redArg(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -lean_dec(x_1); -x_8 = lean_box(x_7); -return x_8; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("+", 1, 1); +return x_1; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1___redArg___boxed(lean_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_IR_EmitC_emitStructReshapeFn___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -uint8_t x_7; lean_object* x_8; -x_7 = l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1___redArg(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_6); -lean_dec(x_5); +uint8_t x_5; lean_object* x_6; +x_5 = l_Lean_IR_EmitC_emitStructReshapeFn___lam__0(x_1, x_2, x_3, x_4); lean_dec(x_4); lean_dec_ref(x_3); -lean_dec_ref(x_1); -x_8 = lean_box(x_7); -return x_8; +lean_dec_ref(x_2); +x_6 = lean_box(x_5); +return x_6; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00__private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1_spec__1___redArg___boxed(lean_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_IR_EmitC_emitStructReshapeFn___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -uint8_t x_7; lean_object* x_8; -x_7 = l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00__private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at___00Lean_IR_EmitC_overwriteParam_spec__1_spec__1___redArg(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_2); +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_4); +x_11 = l_Lean_IR_EmitC_emitStructReshapeFn___lam__1(x_1, x_2, x_3, x_10, x_5, x_6, x_7, x_8, x_9); +lean_dec_ref(x_8); +lean_dec_ref(x_5); +lean_dec_ref(x_3); lean_dec_ref(x_1); -x_8 = lean_box(x_7); -return x_8; +return x_11; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__0___redArg(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_IR_EmitC_emitStructReshapeFn(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_6; uint8_t x_7; -x_6 = lean_unsigned_to_nat(0u); -x_7 = lean_nat_dec_eq(x_4, x_6); -if (x_7 == 1) +lean_object* x_7; uint8_t x_18; +x_18 = lean_nat_dec_eq(x_3, x_4); +if (x_18 == 0) { -lean_object* x_8; lean_object* x_9; +uint8_t x_19; +x_19 = l_Lean_IR_EmitC_compatibleReshape(x_1, x_2); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +lean_dec_ref(x_5); lean_dec(x_4); -x_8 = lean_box(0); -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_8); -lean_ctor_set(x_9, 1, x_5); -return x_9; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_20 = lean_box(0); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_6); +return x_21; } else { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_10 = lean_unsigned_to_nat(1u); -x_11 = lean_nat_sub(x_4, x_10); -lean_dec(x_4); -x_12 = lean_nat_sub(x_3, x_11); -x_13 = lean_nat_sub(x_12, x_10); -lean_dec(x_12); -x_14 = lean_array_fget_borrowed(x_1, x_13); -x_15 = l_Lean_IR_EmitC_overwriteParam___closed__0; -x_16 = lean_array_get_borrowed(x_15, x_2, x_13); -lean_dec(x_13); -x_17 = l_Lean_IR_EmitC_paramEqArg(x_14, x_16); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_18 = lean_ctor_get(x_14, 0); -x_19 = l_Lean_IR_EmitC_argToCString___closed__0; -lean_inc(x_18); -x_20 = l_Nat_reprFast(x_18); -x_21 = lean_string_append(x_19, x_20); -lean_dec_ref(x_20); -x_22 = lean_string_append(x_5, x_21); -lean_dec_ref(x_21); -x_23 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0___redArg___closed__0; -x_24 = lean_string_append(x_22, x_23); -lean_inc(x_16); -x_25 = l_Lean_IR_EmitC_emitArg___redArg(x_16, x_24); -x_26 = lean_ctor_get(x_25, 1); -lean_inc(x_26); -lean_dec_ref(x_25); -x_27 = l_Lean_IR_EmitC_emitFnDeclAux___closed__0; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_22 = lean_box(0); +x_23 = l_Lean_IR_EmitC_emitFnDeclAux___closed__5; +x_24 = lean_string_append(x_6, x_23); +lean_inc(x_4); +x_25 = l_Lean_IR_EmitC_structType(x_4); +x_26 = lean_string_append(x_24, x_25); +x_27 = l_Lean_IR_EmitC_emitSpreadArg___closed__0; x_28 = lean_string_append(x_26, x_27); -x_29 = l_Lean_IR_EmitC_emitLn___closed__0; +x_29 = l_Lean_IR_EmitC_structReshapeFnPrefix___closed__0; x_30 = lean_string_append(x_28, x_29); -x_4 = x_11; -x_5 = x_30; -goto _start; +lean_inc(x_3); +x_31 = l_Nat_reprFast(x_3); +x_32 = lean_string_append(x_30, x_31); +lean_dec_ref(x_31); +x_33 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadArg_spec__0___redArg___closed__0; +x_34 = lean_string_append(x_32, x_33); +x_35 = l_Nat_reprFast(x_4); +x_36 = lean_string_append(x_34, x_35); +lean_dec_ref(x_35); +x_37 = l_Lean_IR_EmitC_emitFnDeclAux___closed__2; +x_38 = lean_string_append(x_36, x_37); +x_39 = l_Lean_IR_EmitC_structType(x_3); +x_40 = lean_string_append(x_38, x_39); +lean_dec_ref(x_39); +x_41 = l_Lean_IR_EmitC_emitStructIncDecFn___closed__2; +x_42 = lean_string_append(x_40, x_41); +x_43 = l_Lean_IR_EmitC_emitLn___closed__0; +x_44 = lean_string_append(x_42, x_43); +x_45 = lean_string_append(x_44, x_25); +lean_dec_ref(x_25); +x_46 = l_Lean_IR_EmitC_emitStructReshapeFn___closed__1; +x_47 = lean_string_append(x_45, x_46); +x_48 = lean_string_append(x_47, x_43); +switch (lean_obj_tag(x_1)) { +case 11: +{ +if (lean_obj_tag(x_2) == 11) +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_49 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_49); +lean_dec_ref(x_1); +x_50 = lean_ctor_get(x_2, 1); +lean_inc_ref(x_50); +lean_dec_ref(x_2); +lean_inc_ref(x_50); +lean_inc_ref(x_49); +x_51 = lean_alloc_closure((void*)(l_Lean_IR_EmitC_emitStructReshapeFn___lam__0___boxed), 4, 3); +lean_closure_set(x_51, 0, x_22); +lean_closure_set(x_51, 1, x_49); +lean_closure_set(x_51, 2, x_50); +x_52 = lean_box(x_18); +lean_inc_ref(x_49); +x_53 = lean_alloc_closure((void*)(l_Lean_IR_EmitC_emitStructReshapeFn___lam__1___boxed), 9, 5); +lean_closure_set(x_53, 0, x_49); +lean_closure_set(x_53, 1, x_22); +lean_closure_set(x_53, 2, x_50); +lean_closure_set(x_53, 3, x_52); +lean_closure_set(x_53, 4, x_43); +x_54 = l_Lean_IR_EmitC_emitStructReshapeFn___closed__2; +x_55 = lean_string_append(x_48, x_54); +x_56 = lean_string_append(x_55, x_43); +x_57 = lean_array_get_size(x_49); +lean_dec_ref(x_49); +x_58 = l_Lean_IR_EmitC_emitStructIncDecFn___closed__0; +x_59 = l_Lean_IR_EmitC_emitUnionSwitchWithImpossible(x_57, x_58, x_51, x_53, x_5, x_56); +if (lean_obj_tag(x_59) == 0) +{ +lean_object* x_60; +x_60 = lean_ctor_get(x_59, 1); +lean_inc(x_60); +lean_dec_ref(x_59); +x_7 = x_60; +goto block_17; } else { -x_4 = x_11; -goto _start; -} -} +return x_59; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: +else { -lean_object* x_8; -x_8 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__0___redArg(x_1, x_2, x_3, x_4, x_7); -return x_8; +lean_dec_ref(x_5); +lean_dec(x_2); +lean_dec_ref(x_1); +x_7 = x_48; +goto block_17; } } -static lean_object* _init_l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__1___redArg___closed__0() { -_start: +case 10: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked(" _tmp_", 6, 6); -return x_1; -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__1___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +if (lean_obj_tag(x_2) == 10) { -lean_object* x_6; uint8_t x_7; -x_6 = lean_unsigned_to_nat(0u); -x_7 = lean_nat_dec_eq(x_4, x_6); -if (x_7 == 1) +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_85; uint8_t x_86; uint8_t x_89; +x_61 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_61); +x_62 = lean_ctor_get(x_1, 2); +lean_inc(x_62); +x_63 = lean_ctor_get(x_1, 3); +lean_inc(x_63); +lean_dec_ref(x_1); +x_64 = lean_ctor_get(x_2, 1); +lean_inc_ref(x_64); +lean_dec_ref(x_2); +x_85 = lean_unsigned_to_nat(0u); +x_89 = lean_nat_dec_eq(x_62, x_85); +if (x_89 == 0) { -lean_object* x_8; lean_object* x_9; -lean_dec(x_4); -x_8 = lean_box(0); -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_8); -lean_ctor_set(x_9, 1, x_5); -return x_9; +x_86 = x_19; +goto block_88; } else { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_10 = lean_unsigned_to_nat(1u); -x_11 = lean_nat_sub(x_4, x_10); -lean_dec(x_4); -x_12 = lean_nat_sub(x_3, x_11); -x_13 = lean_nat_sub(x_12, x_10); -lean_dec(x_12); -x_14 = lean_array_fget_borrowed(x_1, x_13); -x_15 = l_Lean_IR_EmitC_overwriteParam___closed__0; -x_16 = lean_array_get_borrowed(x_15, x_2, x_13); -x_17 = l_Lean_IR_EmitC_paramEqArg(x_14, x_16); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_18 = lean_ctor_get(x_14, 1); -x_19 = l_Lean_IR_EmitC_toCType(x_18); -x_20 = lean_string_append(x_5, x_19); -lean_dec_ref(x_19); -x_21 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__1___redArg___closed__0; -x_22 = lean_string_append(x_20, x_21); -x_23 = l_Nat_reprFast(x_13); -x_24 = lean_string_append(x_22, x_23); -lean_dec_ref(x_23); -x_25 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0___redArg___closed__0; -x_26 = lean_string_append(x_24, x_25); -lean_inc(x_16); -x_27 = l_Lean_IR_EmitC_emitArg___redArg(x_16, x_26); -x_28 = lean_ctor_get(x_27, 1); -lean_inc(x_28); -lean_dec_ref(x_27); -x_29 = l_Lean_IR_EmitC_emitFnDeclAux___closed__0; -x_30 = lean_string_append(x_28, x_29); -x_31 = l_Lean_IR_EmitC_emitLn___closed__0; -x_32 = lean_string_append(x_30, x_31); -x_4 = x_11; -x_5 = x_32; -goto _start; +x_86 = x_18; +goto block_88; } -else +block_72: { -lean_dec(x_13); -x_4 = x_11; -goto _start; -} -} -} +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_67 = lean_array_get_size(x_61); +x_68 = lean_unsigned_to_nat(0u); +x_69 = lean_box(0); +x_70 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg(x_61, x_22, x_64, x_69, x_19, x_67, x_18, x_68, x_69, x_65, x_66); +lean_dec_ref(x_65); +lean_dec_ref(x_64); +lean_dec_ref(x_61); +if (lean_obj_tag(x_70) == 0) +{ +lean_object* x_71; +x_71 = lean_ctor_get(x_70, 1); +lean_inc(x_71); +lean_dec_ref(x_70); +x_7 = x_71; +goto block_17; } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: +else { -lean_object* x_8; -x_8 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__1___redArg(x_1, x_2, x_3, x_4, x_7); -return x_8; +return x_70; } } -static lean_object* _init_l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__2___redArg___closed__0() { -_start: +block_84: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked(" = _tmp_", 8, 8); -return x_1; -} +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_73 = l_Lean_IR_EmitC_emitStructReshapeFn___closed__3; +x_74 = lean_string_append(x_48, x_73); +x_75 = l_Nat_reprFast(x_62); +x_76 = lean_string_append(x_74, x_75); +lean_dec_ref(x_75); +x_77 = l_Lean_IR_EmitC_emitStructReshapeFn___closed__4; +x_78 = lean_string_append(x_76, x_77); +x_79 = l_Nat_reprFast(x_63); +x_80 = lean_string_append(x_78, x_79); +lean_dec_ref(x_79); +x_81 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_82 = lean_string_append(x_80, x_81); +x_83 = lean_string_append(x_82, x_43); +x_65 = x_5; +x_66 = x_83; +goto block_72; } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__2___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +block_88: { -lean_object* x_6; uint8_t x_7; -x_6 = lean_unsigned_to_nat(0u); -x_7 = lean_nat_dec_eq(x_4, x_6); -if (x_7 == 1) +if (x_86 == 0) { -lean_object* x_8; lean_object* x_9; -lean_dec(x_4); -x_8 = lean_box(0); -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_8); -lean_ctor_set(x_9, 1, x_5); -return x_9; -} -else +uint8_t x_87; +x_87 = lean_nat_dec_eq(x_63, x_85); +if (x_87 == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_10 = lean_unsigned_to_nat(1u); -x_11 = lean_nat_sub(x_4, x_10); -lean_dec(x_4); -x_12 = lean_nat_sub(x_3, x_11); -x_13 = lean_nat_sub(x_12, x_10); -lean_dec(x_12); -x_14 = lean_array_fget_borrowed(x_1, x_13); -x_15 = l_Lean_IR_EmitC_overwriteParam___closed__0; -x_16 = lean_array_get_borrowed(x_15, x_2, x_13); -x_17 = l_Lean_IR_EmitC_paramEqArg(x_14, x_16); -if (x_17 == 0) +if (x_19 == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_18 = lean_ctor_get(x_14, 0); -x_19 = l_Lean_IR_EmitC_argToCString___closed__0; -lean_inc(x_18); -x_20 = l_Nat_reprFast(x_18); -x_21 = lean_string_append(x_19, x_20); -lean_dec_ref(x_20); -x_22 = lean_string_append(x_5, x_21); -lean_dec_ref(x_21); -x_23 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__2___redArg___closed__0; -x_24 = lean_string_append(x_22, x_23); -x_25 = l_Nat_reprFast(x_13); -x_26 = lean_string_append(x_24, x_25); -lean_dec_ref(x_25); -x_27 = l_Lean_IR_EmitC_emitFnDeclAux___closed__0; -x_28 = lean_string_append(x_26, x_27); -x_29 = l_Lean_IR_EmitC_emitLn___closed__0; -x_30 = lean_string_append(x_28, x_29); -x_4 = x_11; -x_5 = x_30; -goto _start; +lean_dec(x_63); +lean_dec(x_62); +x_65 = x_5; +x_66 = x_48; +goto block_72; } else { -lean_dec(x_13); -x_4 = x_11; -goto _start; +goto block_84; } } +else +{ +lean_dec(x_63); +lean_dec(x_62); +x_65 = x_5; +x_66 = x_48; +goto block_72; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_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: +else { -lean_object* x_8; -x_8 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__2___redArg(x_1, x_2, x_3, x_4, x_7); -return x_8; +goto block_84; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitTailCall_spec__3(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; uint8_t x_10; -x_10 = lean_usize_dec_eq(x_2, x_3); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_11 = lean_array_uget(x_1, x_2); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_12, 1); -lean_inc(x_13); -lean_dec(x_12); -x_14 = l_Lean_IR_IRType_isVoid(x_13); -lean_dec(x_13); -if (x_14 == 0) -{ -lean_object* x_15; -x_15 = lean_array_push(x_4, x_11); -x_5 = x_15; -goto block_9; } else { -lean_dec(x_11); -x_5 = x_4; -goto block_9; +lean_dec_ref(x_5); +lean_dec(x_2); +lean_dec_ref(x_1); +x_7 = x_48; +goto block_17; +} +} +default: +{ +lean_dec_ref(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_7 = x_48; +goto block_17; +} +} } } else { -return x_4; +lean_object* x_90; lean_object* x_91; +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_90 = lean_box(0); +x_91 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_91, 0, x_90); +lean_ctor_set(x_91, 1, x_6); +return x_91; } -block_9: +block_17: { -size_t x_6; size_t x_7; -x_6 = 1; -x_7 = lean_usize_add(x_2, x_6); -x_2 = x_7; -x_4 = x_5; -goto _start; +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_8 = l_Lean_IR_EmitC_emitStructReshapeFn___closed__0; +x_9 = lean_string_append(x_7, x_8); +x_10 = l_Lean_IR_EmitC_emitLn___closed__0; +x_11 = lean_string_append(x_9, x_10); +x_12 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; +x_13 = lean_string_append(x_11, x_12); +x_14 = lean_box(0); +x_15 = lean_string_append(x_13, x_10); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +return x_16; } } } -static lean_object* _init_l_Lean_IR_EmitC_emitTailCall___closed__0() { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("goto _start;", 12, 12); -return x_1; +uint8_t x_15; uint8_t x_16; lean_object* x_17; +x_15 = lean_unbox(x_5); +x_16 = lean_unbox(x_7); +x_17 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0(x_1, x_2, x_3, x_4, x_15, x_6, x_16, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec_ref(x_13); +lean_dec(x_6); +lean_dec_ref(x_3); +lean_dec_ref(x_1); +return x_17; } } -static lean_object* _init_l_Lean_IR_EmitC_emitTailCall___closed__1() { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("{", 1, 1); -return x_1; +lean_object* x_7; +x_7 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___lam__0(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec_ref(x_5); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_7; } } -static lean_object* _init_l_Lean_IR_EmitC_emitTailCall___closed__2() { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("invalid tail call", 17, 17); -return x_1; +uint8_t x_12; uint8_t x_13; lean_object* x_14; +x_12 = lean_unbox(x_5); +x_13 = lean_unbox(x_7); +x_14 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg(x_1, x_2, x_3, x_4, x_12, x_6, x_13, x_8, x_9, x_10, x_11); +lean_dec_ref(x_10); +lean_dec(x_6); +lean_dec_ref(x_3); +lean_dec_ref(x_1); +return x_14; } } -static lean_object* _init_l_Lean_IR_EmitC_emitTailCall___closed__3() { +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructBox_spec__0___redArg___closed__0() { _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_object* x_1; +x_1 = lean_mk_string_unchecked("lean_ctor_set(y, ", 17, 17); +return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitTailCall___closed__4() { +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructBox_spec__0___redArg___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("bug at emitTailCall", 19, 19); +x_1 = lean_mk_string_unchecked("lean_box(0));", 13, 13); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitTailCall(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructBox_spec__0___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_4; lean_object* x_12; -if (lean_obj_tag(x_1) == 6) -{ -uint8_t x_33; -x_33 = !lean_is_exclusive(x_1); -if (x_33 == 0) -{ -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; -x_34 = lean_ctor_get(x_1, 1); -x_35 = lean_ctor_get(x_1, 0); -lean_dec(x_35); -x_36 = lean_ctor_get(x_2, 4); -x_37 = lean_array_get_size(x_36); -x_38 = lean_array_get_size(x_34); -x_39 = lean_nat_dec_eq(x_37, x_38); -if (x_39 == 0) +lean_object* x_11; lean_object* x_12; uint8_t x_17; +x_17 = lean_nat_dec_lt(x_7, x_1); +if (x_17 == 0) { -lean_object* x_40; -lean_dec_ref(x_34); -x_40 = l_Lean_IR_EmitC_emitTailCall___closed__2; -lean_ctor_set_tag(x_1, 1); -lean_ctor_set(x_1, 1, x_3); -lean_ctor_set(x_1, 0, x_40); -return x_1; +lean_object* x_18; +lean_dec(x_7); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_8); +lean_ctor_set(x_18, 1, x_10); +return x_18; } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; -lean_free_object(x_1); -x_41 = l_Array_zip___redArg(x_36, x_34); -lean_dec_ref(x_34); -x_42 = lean_unsigned_to_nat(0u); -x_43 = lean_array_get_size(x_41); -x_44 = l_Lean_IR_EmitC_emitTailCall___closed__3; -x_45 = lean_nat_dec_lt(x_42, x_43); -if (x_45 == 0) +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_37; +x_19 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructBox_spec__0___redArg___closed__0; +x_20 = lean_string_append(x_10, x_19); +lean_inc(x_7); +x_21 = l_Nat_reprFast(x_7); +x_22 = lean_string_append(x_20, x_21); +x_23 = lean_string_append(x_22, x_2); +x_37 = lean_array_fget_borrowed(x_6, x_7); +switch (lean_obj_tag(x_37)) { +case 6: { -lean_dec_ref(x_41); -x_12 = x_44; -goto block_32; +lean_dec_ref(x_21); +goto block_36; } -else +case 13: { -uint8_t x_46; -x_46 = lean_nat_dec_le(x_43, x_43); -if (x_46 == 0) +lean_dec_ref(x_21); +goto block_36; +} +default: { -if (x_45 == 0) +switch (lean_obj_tag(x_37)) { +case 7: { -lean_dec_ref(x_41); -x_12 = x_44; -goto block_32; +goto block_31; } -else +case 8: { -size_t x_47; size_t x_48; lean_object* x_49; -x_47 = 0; -x_48 = lean_usize_of_nat(x_43); -x_49 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitTailCall_spec__3(x_41, x_47, x_48, x_44); -lean_dec_ref(x_41); -x_12 = x_49; -goto block_32; +goto block_31; +} +case 12: +{ +goto block_31; } +default: +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_box(8); +lean_inc(x_37); +x_39 = l_Lean_IR_EmitC_emitBoxFn(x_37, x_38, x_9, x_23); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +lean_dec_ref(x_39); +x_41 = l_Lean_IR_EmitC_emitFnDeclAux___closed__2; +x_42 = lean_string_append(x_40, x_41); +x_43 = lean_string_append(x_42, x_3); +x_44 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__0; +x_45 = lean_string_append(x_43, x_44); +x_46 = lean_string_append(x_45, x_21); +lean_dec_ref(x_21); +x_47 = l_Lean_IR_EmitC_emitSProj___closed__0; +x_48 = lean_string_append(x_46, x_47); +x_49 = l_Lean_IR_EmitC_emitLn___closed__0; +x_50 = lean_string_append(x_48, x_49); +x_11 = x_5; +x_12 = x_50; +goto block_16; } else { -size_t x_50; size_t x_51; lean_object* x_52; -x_50 = 0; -x_51 = lean_usize_of_nat(x_43); -x_52 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitTailCall_spec__3(x_41, x_50, x_51, x_44); -lean_dec_ref(x_41); -x_12 = x_52; -goto block_32; +lean_dec_ref(x_21); +lean_dec(x_7); +return x_39; +} } } } } -else +block_31: { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; -x_53 = lean_ctor_get(x_1, 1); -lean_inc(x_53); -lean_dec(x_1); -x_54 = lean_ctor_get(x_2, 4); -x_55 = lean_array_get_size(x_54); -x_56 = lean_array_get_size(x_53); -x_57 = lean_nat_dec_eq(x_55, x_56); -if (x_57 == 0) +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_24 = lean_string_append(x_23, x_3); +x_25 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__0; +x_26 = lean_string_append(x_24, x_25); +x_27 = lean_string_append(x_26, x_21); +lean_dec_ref(x_21); +x_28 = lean_string_append(x_27, x_4); +x_29 = l_Lean_IR_EmitC_emitLn___closed__0; +x_30 = lean_string_append(x_28, x_29); +x_11 = x_5; +x_12 = x_30; +goto block_16; +} +block_36: { -lean_object* x_58; lean_object* x_59; -lean_dec_ref(x_53); -x_58 = l_Lean_IR_EmitC_emitTailCall___closed__2; -x_59 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_59, 0, x_58); -lean_ctor_set(x_59, 1, x_3); -return x_59; +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_32 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructBox_spec__0___redArg___closed__1; +x_33 = lean_string_append(x_23, x_32); +x_34 = l_Lean_IR_EmitC_emitLn___closed__0; +x_35 = lean_string_append(x_33, x_34); +x_11 = x_5; +x_12 = x_35; +goto block_16; } -else +} +block_16: { -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; -x_60 = l_Array_zip___redArg(x_54, x_53); -lean_dec_ref(x_53); -x_61 = lean_unsigned_to_nat(0u); -x_62 = lean_array_get_size(x_60); -x_63 = l_Lean_IR_EmitC_emitTailCall___closed__3; -x_64 = lean_nat_dec_lt(x_61, x_62); -if (x_64 == 0) +lean_object* x_13; lean_object* x_14; +x_13 = lean_unsigned_to_nat(1u); +x_14 = lean_nat_add(x_7, x_13); +lean_dec(x_7); +x_7 = x_14; +x_8 = x_11; +x_10 = x_12; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructBox_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: { -lean_dec_ref(x_60); -x_12 = x_63; -goto block_32; +lean_object* x_14; +x_14 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructBox_spec__0___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_9, x_10, x_12, x_13); +return x_14; } -else +} +static lean_object* _init_l_Lean_IR_EmitC_emitStructBox___closed__0() { +_start: { -uint8_t x_65; -x_65 = lean_nat_dec_le(x_62, x_62); -if (x_65 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("memcpy(lean_ctor_scalar_cptr(y), ", 33, 33); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitStructBox___closed__1() { +_start: { -if (x_64 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked(".u, sizeof(size_t)*", 19, 19); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitStructBox___closed__2() { +_start: { -lean_dec_ref(x_60); -x_12 = x_63; -goto block_32; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("y = lean_alloc_ctor(", 20, 20); +return x_1; } -else +} +static lean_object* _init_l_Lean_IR_EmitC_emitStructBox___closed__3() { +_start: { -size_t x_66; size_t x_67; lean_object* x_68; -x_66 = 0; -x_67 = lean_usize_of_nat(x_62); -x_68 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitTailCall_spec__3(x_60, x_66, x_67, x_63); -lean_dec_ref(x_60); -x_12 = x_68; -goto block_32; +lean_object* x_1; +x_1 = lean_mk_string_unchecked(", sizeof(size_t)*", 17, 17); +return x_1; } } -else +static lean_object* _init_l_Lean_IR_EmitC_emitStructBox___closed__4() { +_start: { -size_t x_69; size_t x_70; lean_object* x_71; -x_69 = 0; -x_70 = lean_usize_of_nat(x_62); -x_71 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitTailCall_spec__3(x_60, x_69, x_70, x_63); -lean_dec_ref(x_60); -x_12 = x_71; -goto block_32; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("y = lean_box(", 13, 13); +return x_1; +} } +static lean_object* _init_l_Lean_IR_EmitC_emitStructBox___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.IR.EmitC.emitStructBox", 27, 27); +return x_1; } } +static lean_object* _init_l_Lean_IR_EmitC_emitStructBox___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_IR_EmitC_emitUSet___closed__5; +x_2 = lean_unsigned_to_nat(34u); +x_3 = lean_unsigned_to_nat(1005u); +x_4 = l_Lean_IR_EmitC_emitStructBox___closed__5; +x_5 = l_Lean_IR_EmitC_emitUSet___closed__3; +x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); +return x_6; } } -else +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructBox(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -lean_object* x_72; lean_object* x_73; +if (lean_obj_tag(x_1) == 10) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_66; uint8_t x_80; +x_6 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_6); +x_7 = lean_ctor_get(x_1, 2); +lean_inc(x_7); +x_8 = lean_ctor_get(x_1, 3); +lean_inc(x_8); lean_dec_ref(x_1); -x_72 = l_Lean_IR_EmitC_emitTailCall___closed__4; -x_73 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_3); -return x_73; +x_80 = l_Array_isEmpty___redArg(x_6); +if (x_80 == 0) +{ +x_66 = x_80; +goto block_79; } -block_11: +else { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_5 = l_Lean_IR_EmitC_emitTailCall___closed__0; -x_6 = lean_string_append(x_4, x_5); -x_7 = l_Lean_IR_EmitC_emitLn___closed__0; -x_8 = lean_box(0); -x_9 = lean_string_append(x_6, x_7); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_8); -lean_ctor_set(x_10, 1, x_9); -return x_10; +lean_object* x_81; uint8_t x_82; +x_81 = lean_unsigned_to_nat(0u); +x_82 = lean_nat_dec_eq(x_7, x_81); +x_66 = x_82; +goto block_79; } -block_32: +block_21: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_13 = l_Array_unzip___redArg(x_12); +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_box(0); +x_16 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructBox_spec__0___redArg(x_9, x_11, x_3, x_10, x_15, x_6, x_14, x_15, x_12, x_13); lean_dec_ref(x_12); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec_ref(x_13); -x_16 = l_Lean_IR_EmitC_overwriteParam(x_14, x_15); -if (x_16 == 0) +lean_dec_ref(x_6); +lean_dec_ref(x_10); +lean_dec_ref(x_11); +lean_dec(x_9); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_array_get_size(x_14); -x_18 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__0___redArg(x_14, x_15, x_17, x_17, x_3); -lean_dec(x_15); -lean_dec(x_14); -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -lean_dec_ref(x_18); -x_4 = x_19; -goto block_11; +uint8_t x_17; +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +lean_object* x_18; +x_18 = lean_ctor_get(x_16, 0); +lean_dec(x_18); +lean_ctor_set(x_16, 0, x_15); +return x_16; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_20 = l_Lean_IR_EmitC_emitTailCall___closed__1; -x_21 = lean_string_append(x_3, x_20); -x_22 = l_Lean_IR_EmitC_emitLn___closed__0; -x_23 = lean_string_append(x_21, x_22); -x_24 = lean_array_get_size(x_14); -x_25 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__1___redArg(x_14, x_15, x_24, x_24, x_23); -x_26 = lean_ctor_get(x_25, 1); -lean_inc(x_26); -lean_dec_ref(x_25); -x_27 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__2___redArg(x_14, x_15, x_24, x_24, x_26); -lean_dec(x_15); -lean_dec(x_14); -x_28 = lean_ctor_get(x_27, 1); -lean_inc(x_28); -lean_dec_ref(x_27); -x_29 = l_Lean_IR_EmitC_emitMainFn___closed__12; -x_30 = lean_string_append(x_28, x_29); -x_31 = lean_string_append(x_30, x_22); -x_4 = x_31; -goto block_11; -} -} +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_16, 1); +lean_inc(x_19); +lean_dec(x_16); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_15); +lean_ctor_set(x_20, 1, x_19); +return x_20; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: +else { -lean_object* x_8; -x_8 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec_ref(x_6); -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -return x_8; +return x_16; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_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: +block_40: { -lean_object* x_8; -x_8 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec_ref(x_6); -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -return x_8; -} +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_30 = l_Lean_IR_EmitC_emitStructBox___closed__0; +x_31 = lean_string_append(x_23, x_30); +x_32 = lean_string_append(x_31, x_3); +x_33 = l_Lean_IR_EmitC_emitStructBox___closed__1; +x_34 = lean_string_append(x_32, x_33); +x_35 = lean_string_append(x_34, x_24); +lean_dec_ref(x_24); +x_36 = lean_string_append(x_35, x_22); +lean_dec_ref(x_22); +x_37 = lean_string_append(x_36, x_29); +lean_dec_ref(x_29); +x_38 = lean_string_append(x_37, x_27); +x_39 = lean_string_append(x_38, x_26); +lean_dec_ref(x_26); +x_9 = x_25; +x_10 = x_27; +x_11 = x_28; +x_12 = x_4; +x_13 = x_39; +goto block_21; } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_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: +block_65: { -lean_object* x_8; -x_8 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec_ref(x_6); -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -return x_8; -} +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; +x_41 = l_Lean_IR_EmitC_emitStructBox___closed__2; +x_42 = lean_string_append(x_5, x_41); +x_43 = l_Nat_reprFast(x_2); +x_44 = lean_string_append(x_42, x_43); +lean_dec_ref(x_43); +x_45 = l_Lean_IR_EmitC_emitSpreadArg___closed__1; +x_46 = lean_string_append(x_44, x_45); +x_47 = lean_array_get_size(x_6); +x_48 = l_Nat_reprFast(x_47); +x_49 = lean_string_append(x_46, x_48); +lean_dec_ref(x_48); +x_50 = l_Lean_IR_EmitC_emitStructBox___closed__3; +x_51 = lean_string_append(x_49, x_50); +lean_inc(x_7); +x_52 = l_Nat_reprFast(x_7); +x_53 = lean_string_append(x_51, x_52); +x_54 = l_Lean_IR_EmitC_emitStructReshapeFn___closed__4; +x_55 = lean_string_append(x_53, x_54); +lean_inc(x_8); +x_56 = l_Nat_reprFast(x_8); +x_57 = lean_string_append(x_55, x_56); +x_58 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_59 = lean_string_append(x_57, x_58); +x_60 = l_Lean_IR_EmitC_emitLn___closed__0; +x_61 = lean_string_append(x_59, x_60); +x_62 = lean_unsigned_to_nat(0u); +x_63 = lean_nat_dec_eq(x_7, x_62); +lean_dec(x_7); +if (x_63 == 0) +{ +lean_dec(x_8); +x_22 = x_54; +x_23 = x_61; +x_24 = x_52; +x_25 = x_47; +x_26 = x_60; +x_27 = x_58; +x_28 = x_45; +x_29 = x_56; +goto block_40; } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitTailCall_spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +else { -size_t x_5; size_t x_6; lean_object* x_7; -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitTailCall_spec__3(x_1, x_5, x_6, x_4); -lean_dec_ref(x_1); -return x_7; +uint8_t x_64; +x_64 = lean_nat_dec_eq(x_8, x_62); +lean_dec(x_8); +if (x_64 == 0) +{ +x_22 = x_54; +x_23 = x_61; +x_24 = x_52; +x_25 = x_47; +x_26 = x_60; +x_27 = x_58; +x_28 = x_45; +x_29 = x_56; +goto block_40; +} +else +{ +lean_dec_ref(x_56); +lean_dec_ref(x_52); +x_9 = x_47; +x_10 = x_58; +x_11 = x_45; +x_12 = x_4; +x_13 = x_61; +goto block_21; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__2___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +} +block_79: { -lean_object* x_6; -x_6 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__2___redArg(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -return x_6; +if (x_66 == 0) +{ +goto block_65; } +else +{ +lean_object* x_67; uint8_t x_68; +x_67 = lean_unsigned_to_nat(0u); +x_68 = lean_nat_dec_eq(x_8, x_67); +if (x_68 == 0) +{ +goto block_65; } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +else { -lean_object* x_6; -x_6 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__0___redArg(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -return x_6; +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec_ref(x_4); +x_69 = l_Lean_IR_EmitC_emitStructBox___closed__4; +x_70 = lean_string_append(x_5, x_69); +x_71 = l_Nat_reprFast(x_2); +x_72 = lean_string_append(x_70, x_71); +lean_dec_ref(x_71); +x_73 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_74 = lean_string_append(x_72, x_73); +x_75 = l_Lean_IR_EmitC_emitLn___closed__0; +x_76 = lean_string_append(x_74, x_75); +x_77 = lean_box(0); +x_78 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_78, 0, x_77); +lean_ctor_set(x_78, 1, x_76); +return x_78; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__1___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_6; -x_6 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitTailCall_spec__1___redArg(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -return x_6; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitTailCall___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -lean_object* x_4; -x_4 = l_Lean_IR_EmitC_emitTailCall(x_1, x_2, x_3); -lean_dec_ref(x_2); -return x_4; +lean_object* x_83; lean_object* x_84; +lean_dec(x_2); +lean_dec(x_1); +x_83 = l_Lean_IR_EmitC_emitStructBox___closed__6; +x_84 = l_panic___at___00Lean_IR_EmitC_emitUSet_spec__0(x_83, x_4, x_5); +return x_84; } } -static lean_object* _init_l_Lean_IR_EmitC_emitIf___closed__0() { +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructBox_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("if (", 4, 4); -return x_1; +lean_object* x_14; +x_14 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructBox_spec__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec_ref(x_12); +lean_dec_ref(x_6); +lean_dec_ref(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +lean_dec(x_1); +return x_14; } } -static lean_object* _init_l_Lean_IR_EmitC_emitIf___closed__1() { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructBox_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked(" == ", 4, 4); -return x_1; +lean_object* x_11; +x_11 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructBox_spec__0___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec_ref(x_9); +lean_dec_ref(x_6); +lean_dec_ref(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +lean_dec(x_1); +return x_11; } } -static lean_object* _init_l_Lean_IR_EmitC_emitIf___closed__2() { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructBox___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("else", 4, 4); -return x_1; +lean_object* x_6; +x_6 = l_Lean_IR_EmitC_emitStructBox(x_1, x_2, x_3, x_4, x_5); +lean_dec_ref(x_3); +return x_6; } } -static lean_object* _init_l_Lean_IR_EmitC_emitCase___closed__0() { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructBoxFn___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("switch (", 8, 8); -return x_1; +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_6 = lean_array_fget_borrowed(x_1, x_2); +x_7 = l_Lean_IR_EmitC_emitStructIncDecFn___lam__0___closed__0; +lean_inc(x_2); +x_8 = l_Nat_reprFast(x_2); +x_9 = lean_string_append(x_7, x_8); +lean_dec_ref(x_8); +lean_inc(x_6); +x_10 = l_Lean_IR_EmitC_emitStructBox(x_6, x_2, x_9, x_4, x_5); +lean_dec_ref(x_9); +return x_10; } } -static lean_object* _init_l_Lean_IR_EmitC_emitCase___closed__1() { +static lean_object* _init_l_Lean_IR_EmitC_emitStructBoxFn___closed__0() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked(") {", 3, 3); +x_1 = lean_mk_string_unchecked("static lean_object* ", 20, 20); return x_1; } } -static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitCase_spec__4___closed__0() { +static lean_object* _init_l_Lean_IR_EmitC_emitStructBoxFn___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("case ", 5, 5); +x_1 = lean_mk_string_unchecked("lean_object* y;", 15, 15); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitJPs___closed__0() { +static lean_object* _init_l_Lean_IR_EmitC_emitStructBoxFn___closed__2() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked(":", 1, 1); +x_1 = lean_mk_string_unchecked("x", 1, 1); return x_1; } } -static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitCase_spec__4___closed__1() { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructBoxFn___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("default: ", 9, 9); -return x_1; +lean_object* x_6; +x_6 = l_Lean_IR_EmitC_emitStructBoxFn___lam__0(x_1, x_2, x_3, x_4, x_5); +lean_dec_ref(x_1); +return x_6; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitCase_spec__4(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructBoxFn(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_7; uint8_t x_14; -x_14 = lean_usize_dec_eq(x_2, x_3); -if (x_14 == 0) +lean_object* x_5; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_16 = l_Lean_IR_EmitC_emitStructBoxFn___closed__0; +x_17 = lean_string_append(x_4, x_16); +x_18 = l_Lean_IR_EmitC_structBoxFnPrefix___closed__0; +x_19 = lean_string_append(x_17, x_18); +lean_inc(x_2); +x_20 = l_Nat_reprFast(x_2); +x_21 = lean_string_append(x_19, x_20); +lean_dec_ref(x_20); +x_22 = l_Lean_IR_EmitC_emitFnDeclAux___closed__2; +x_23 = lean_string_append(x_21, x_22); +x_24 = l_Lean_IR_EmitC_structType(x_2); +x_25 = lean_string_append(x_23, x_24); +lean_dec_ref(x_24); +x_26 = l_Lean_IR_EmitC_emitStructIncDecFn___closed__2; +x_27 = lean_string_append(x_25, x_26); +x_28 = l_Lean_IR_EmitC_emitLn___closed__0; +x_29 = lean_string_append(x_27, x_28); +x_30 = l_Lean_IR_EmitC_emitStructBoxFn___closed__1; +x_31 = lean_string_append(x_29, x_30); +x_32 = lean_string_append(x_31, x_28); +if (lean_obj_tag(x_1) == 11) { -lean_object* x_15; -x_15 = lean_array_uget(x_1, x_2); -if (lean_obj_tag(x_15) == 0) +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_33 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_33); +lean_dec_ref(x_1); +lean_inc_ref(x_33); +x_34 = lean_alloc_closure((void*)(l_Lean_IR_EmitC_emitStructBoxFn___lam__0___boxed), 5, 1); +lean_closure_set(x_34, 0, x_33); +x_35 = lean_array_get_size(x_33); +lean_dec_ref(x_33); +x_36 = l_Lean_IR_EmitC_emitStructIncDecFn___closed__0; +x_37 = l_Lean_IR_EmitC_emitUnionSwitch(x_35, x_36, x_34, x_3, x_32); +if (lean_obj_tag(x_37) == 0) { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_16 = lean_ctor_get(x_15, 0); -lean_inc_ref(x_16); -x_17 = lean_ctor_get(x_15, 1); -lean_inc(x_17); -lean_dec_ref(x_15); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec_ref(x_16); -x_19 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitCase_spec__4___closed__0; -x_20 = lean_string_append(x_6, x_19); -x_21 = l_Nat_reprFast(x_18); -x_22 = lean_string_append(x_20, x_21); -lean_dec_ref(x_21); -x_23 = l_Lean_IR_EmitC_emitJPs___closed__0; -x_24 = lean_string_append(x_22, x_23); -x_25 = l_Lean_IR_EmitC_emitLn___closed__0; -x_26 = lean_string_append(x_24, x_25); -lean_inc_ref(x_5); -x_27 = l_Lean_IR_EmitC_emitFnBody(x_17, x_5, x_26); -x_7 = x_27; -goto block_13; +lean_object* x_38; +x_38 = lean_ctor_get(x_37, 1); +lean_inc(x_38); +lean_dec_ref(x_37); +x_5 = x_38; +goto block_15; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_28 = lean_ctor_get(x_15, 0); -lean_inc(x_28); -lean_dec_ref(x_15); -x_29 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitCase_spec__4___closed__1; -x_30 = lean_string_append(x_6, x_29); -x_31 = l_Lean_IR_EmitC_emitLn___closed__0; -x_32 = lean_string_append(x_30, x_31); -lean_inc_ref(x_5); -x_33 = l_Lean_IR_EmitC_emitFnBody(x_28, x_5, x_32); -x_7 = x_33; -goto block_13; +return x_37; } } else { -lean_object* x_34; -lean_dec_ref(x_5); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_4); -lean_ctor_set(x_34, 1, x_6); -return x_34; -} -block_13: -{ -if (lean_obj_tag(x_7) == 0) +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_unsigned_to_nat(0u); +x_40 = l_Lean_IR_EmitC_emitStructBoxFn___closed__2; +x_41 = l_Lean_IR_EmitC_emitStructBox(x_1, x_39, x_40, x_3, x_32); +if (lean_obj_tag(x_41) == 0) { -lean_object* x_8; lean_object* x_9; size_t x_10; size_t x_11; -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_7, 1); -lean_inc(x_9); -lean_dec_ref(x_7); -x_10 = 1; -x_11 = lean_usize_add(x_2, x_10); -x_2 = x_11; -x_4 = x_8; -x_6 = x_9; -goto _start; +lean_object* x_42; +x_42 = lean_ctor_get(x_41, 1); +lean_inc(x_42); +lean_dec_ref(x_41); +x_5 = x_42; +goto block_15; } else { -lean_dec_ref(x_5); -return x_7; +return x_41; +} } +block_15: +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_6 = l_Lean_IR_EmitC_emitStructReshapeFn___closed__0; +x_7 = lean_string_append(x_5, x_6); +x_8 = l_Lean_IR_EmitC_emitLn___closed__0; +x_9 = lean_string_append(x_7, x_8); +x_10 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; +x_11 = lean_string_append(x_9, x_10); +x_12 = lean_box(0); +x_13 = lean_string_append(x_11, x_8); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +return x_14; } } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCase(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructUnboxFn_spec__0___redArg___closed__0() { _start: { -lean_object* x_6; lean_object* x_14; lean_object* x_17; -x_17 = l_Lean_IR_EmitC_isIf(x_3); -if (lean_obj_tag(x_17) == 1) -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -lean_dec_ref(x_3); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -lean_dec_ref(x_17); -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 0); -lean_inc(x_20); -lean_dec(x_18); -x_21 = lean_ctor_get(x_19, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_19, 1); -lean_inc(x_22); -lean_dec(x_19); -x_23 = l_Lean_IR_EmitC_emitIf(x_1, x_2, x_20, x_21, x_22, x_4, x_5); -return x_23; +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" = lean_ctor_get(x, ", 20, 20); +return x_1; } -else -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; -lean_dec(x_17); -x_24 = l_Lean_IR_EmitC_emitCase___closed__0; -x_25 = lean_string_append(x_5, x_24); -x_26 = l_Lean_IR_EmitC_emitTag___redArg(x_1, x_2, x_25); -if (lean_obj_tag(x_26) == 0) +} +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructUnboxFn_spec__0___redArg___closed__1() { +_start: { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; -x_27 = lean_ctor_get(x_26, 1); -lean_inc(x_27); -lean_dec_ref(x_26); -x_28 = l_Lean_IR_EmitC_emitCase___closed__1; -x_29 = lean_string_append(x_27, x_28); -x_30 = l_Lean_IR_EmitC_emitLn___closed__0; -x_31 = lean_string_append(x_29, x_30); -x_32 = l_Lean_IR_ensureHasDefault(x_3); -x_33 = lean_unsigned_to_nat(0u); -x_34 = lean_array_get_size(x_32); -x_35 = lean_nat_dec_lt(x_33, x_34); -if (x_35 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("(lean_ctor_get(x, ", 18, 18); +return x_1; +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructUnboxFn_spec__0___redArg(lean_object* x_1, lean_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_dec_ref(x_32); -lean_dec_ref(x_4); -x_6 = x_31; -goto block_13; +lean_object* x_8; lean_object* x_9; uint8_t x_26; +x_26 = lean_nat_dec_lt(x_4, x_2); +if (x_26 == 0) +{ +lean_object* x_27; +lean_dec(x_4); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_5); +lean_ctor_set(x_27, 1, x_7); +return x_27; } else { -lean_object* x_36; uint8_t x_37; -x_36 = lean_box(0); -x_37 = lean_nat_dec_le(x_34, x_34); -if (x_37 == 0) -{ -if (x_35 == 0) +lean_object* x_28; +x_28 = lean_array_fget_borrowed(x_3, x_4); +switch (lean_obj_tag(x_28)) { +case 6: { -lean_dec_ref(x_32); -lean_dec_ref(x_4); -x_6 = x_31; +x_8 = x_1; +x_9 = x_7; goto block_13; } -else +case 13: { -size_t x_38; size_t x_39; lean_object* x_40; -x_38 = 0; -x_39 = lean_usize_of_nat(x_34); -x_40 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitCase_spec__4(x_32, x_38, x_39, x_36, x_4, x_31); -lean_dec_ref(x_32); -x_14 = x_40; -goto block_16; +x_8 = x_1; +x_9 = x_7; +goto block_13; } +default: +{ +switch (lean_obj_tag(x_28)) { +case 7: +{ +goto block_25; } -else +case 8: { -size_t x_41; size_t x_42; lean_object* x_43; -x_41 = 0; -x_42 = lean_usize_of_nat(x_34); -x_43 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitCase_spec__4(x_32, x_41, x_42, x_36, x_4, x_31); -lean_dec_ref(x_32); -x_14 = x_43; -goto block_16; +goto block_25; } +case 12: +{ +goto block_25; } +default: +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_29 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__0; +x_30 = lean_string_append(x_7, x_29); +lean_inc(x_4); +x_31 = l_Nat_reprFast(x_4); +x_32 = lean_string_append(x_30, x_31); +x_33 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__1; +x_34 = lean_string_append(x_32, x_33); +x_35 = l_Lean_IR_EmitC_emitUnboxFn(x_28, x_6, x_34); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_36 = lean_ctor_get(x_35, 1); +lean_inc(x_36); +lean_dec_ref(x_35); +x_37 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructUnboxFn_spec__0___redArg___closed__1; +x_38 = lean_string_append(x_36, x_37); +x_39 = lean_string_append(x_38, x_31); +lean_dec_ref(x_31); +x_40 = l_Lean_IR_EmitC_emitSProj___closed__0; +x_41 = lean_string_append(x_39, x_40); +x_42 = l_Lean_IR_EmitC_emitLn___closed__0; +x_43 = lean_string_append(x_41, x_42); +x_8 = x_1; +x_9 = x_43; +goto block_13; } else { -lean_dec_ref(x_4); -lean_dec_ref(x_3); -return x_26; +lean_dec_ref(x_31); +lean_dec(x_4); +return x_35; +} +} +} +} } } block_13: { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_7 = l_Lean_IR_EmitC_emitMainFn___closed__12; -x_8 = lean_string_append(x_6, x_7); -x_9 = l_Lean_IR_EmitC_emitLn___closed__0; -x_10 = lean_box(0); -x_11 = lean_string_append(x_8, x_9); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_10); -lean_ctor_set(x_12, 1, x_11); -return x_12; +lean_object* x_10; lean_object* x_11; +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_add(x_4, x_10); +lean_dec(x_4); +x_4 = x_11; +x_5 = x_8; +x_7 = x_9; +goto _start; } -block_16: -{ -if (lean_obj_tag(x_14) == 0) +block_25: { -lean_object* x_15; -x_15 = lean_ctor_get(x_14, 1); -lean_inc(x_15); -lean_dec_ref(x_14); -x_6 = x_15; +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_14 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__0; +x_15 = lean_string_append(x_7, x_14); +lean_inc(x_4); +x_16 = l_Nat_reprFast(x_4); +x_17 = lean_string_append(x_15, x_16); +x_18 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructUnboxFn_spec__0___redArg___closed__0; +x_19 = lean_string_append(x_17, x_18); +x_20 = lean_string_append(x_19, x_16); +lean_dec_ref(x_16); +x_21 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_22 = lean_string_append(x_20, x_21); +x_23 = l_Lean_IR_EmitC_emitLn___closed__0; +x_24 = lean_string_append(x_22, x_23); +x_8 = x_1; +x_9 = x_24; goto block_13; } -else +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructUnboxFn_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructUnboxFn_spec__0___redArg(x_1, x_2, x_3, x_6, x_7, x_9, x_10); +return x_11; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitStructUnboxFn___lam__0___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("(x);", 4, 4); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructUnboxFn___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_7 = l_Lean_IR_EmitC_emitStructReshapeFn___lam__1___closed__0; +x_8 = lean_string_append(x_6, x_7); +lean_inc(x_3); +x_9 = l_Nat_reprFast(x_3); +x_10 = lean_string_append(x_8, x_9); +lean_dec_ref(x_9); +x_11 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__1; +x_12 = lean_string_append(x_10, x_11); +x_13 = lean_array_fget_borrowed(x_1, x_3); +lean_dec(x_3); +x_14 = l_Lean_IR_EmitC_emitUnboxFn(x_13, x_5, x_12); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) { +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_16 = lean_ctor_get(x_14, 1); +x_17 = lean_ctor_get(x_14, 0); +lean_dec(x_17); +x_18 = l_Lean_IR_EmitC_emitStructUnboxFn___lam__0___closed__0; +x_19 = lean_string_append(x_16, x_18); +x_20 = lean_box(0); +x_21 = lean_string_append(x_19, x_2); +lean_ctor_set(x_14, 1, x_21); +lean_ctor_set(x_14, 0, x_20); return x_14; } +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_dec(x_14); +x_23 = l_Lean_IR_EmitC_emitStructUnboxFn___lam__0___closed__0; +x_24 = lean_string_append(x_22, x_23); +x_25 = lean_box(0); +x_26 = lean_string_append(x_24, x_2); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; } } } -static lean_object* _init_l_Lean_IR_EmitC_emitBlock___closed__0() { +static lean_object* _init_l_Lean_IR_EmitC_emitStructUnboxFn___closed__0() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("return ", 7, 7); +x_1 = lean_mk_string_unchecked("(lean_object* x) {", 18, 18); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitBlock___closed__1() { +static lean_object* _init_l_Lean_IR_EmitC_emitStructUnboxFn___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("lean_internal_panic_unreachable();", 34, 34); +x_1 = lean_mk_string_unchecked("y.tag = lean_obj_tag(x);", 24, 24); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitBlock(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_IR_EmitC_emitStructUnboxFn___closed__2() { _start: { -switch (lean_obj_tag(x_1)) { -case 0: -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); -x_5 = lean_ctor_get(x_1, 1); -lean_inc(x_5); -x_6 = lean_ctor_get(x_1, 2); -lean_inc_ref(x_6); -x_7 = lean_ctor_get(x_1, 3); -lean_inc(x_7); -x_8 = lean_ctor_get(x_2, 3); -x_9 = l_Lean_IR_isTailCallTo(x_8, x_1); -lean_dec_ref(x_1); -if (x_9 == 0) -{ -lean_object* x_10; -lean_inc_ref(x_2); -x_10 = l_Lean_IR_EmitC_emitVDecl(x_4, x_5, x_6, x_2, x_3); -if (lean_obj_tag(x_10) == 0) -{ -lean_object* x_11; -x_11 = lean_ctor_get(x_10, 1); -lean_inc(x_11); -lean_dec_ref(x_10); -x_1 = x_7; -x_3 = x_11; -goto _start; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("y.tag", 5, 5); +return x_1; } -else +} +static lean_object* _init_l_Lean_IR_EmitC_emitStructUnboxFn___closed__3() { +_start: { -lean_dec(x_7); -lean_dec_ref(x_2); -return x_10; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("memcpy(y.u, lean_ctor_scalar_cptr(x), sizeof(size_t)*", 53, 53); +return x_1; } } -else +static lean_object* _init_l_Lean_IR_EmitC_emitStructUnboxFn___closed__4() { +_start: { -lean_object* x_13; -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_4); -x_13 = l_Lean_IR_EmitC_emitTailCall(x_6, x_2, x_3); -lean_dec_ref(x_2); -return x_13; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.IR.EmitC.emitStructUnboxFn", 31, 31); +return x_1; } } -case 1: +static lean_object* _init_l_Lean_IR_EmitC_emitStructUnboxFn___closed__5() { +_start: { -lean_object* x_14; -x_14 = lean_ctor_get(x_1, 3); -lean_inc(x_14); -lean_dec_ref(x_1); -x_1 = x_14; -goto _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_IR_EmitC_emitUSet___closed__5; +x_2 = lean_unsigned_to_nat(36u); +x_3 = lean_unsigned_to_nat(1052u); +x_4 = l_Lean_IR_EmitC_emitStructUnboxFn___closed__4; +x_5 = l_Lean_IR_EmitC_emitUSet___closed__3; +x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); +return x_6; } -case 2: -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_16 = lean_ctor_get(x_1, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_1, 1); -lean_inc(x_17); -x_18 = lean_ctor_get(x_1, 2); -lean_inc(x_18); -x_19 = lean_ctor_get(x_1, 3); -lean_inc(x_19); -lean_dec_ref(x_1); -x_20 = l_Lean_IR_EmitC_emitSet___redArg(x_16, x_17, x_18, x_3); -if (lean_obj_tag(x_20) == 0) -{ -lean_object* x_21; -x_21 = lean_ctor_get(x_20, 1); -lean_inc(x_21); -lean_dec_ref(x_20); -x_1 = x_19; -x_3 = x_21; -goto _start; } -else +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructUnboxFn___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -lean_dec(x_19); +lean_object* x_7; +x_7 = l_Lean_IR_EmitC_emitStructUnboxFn___lam__0(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec_ref(x_5); lean_dec_ref(x_2); -return x_20; +lean_dec_ref(x_1); +return x_7; } } -case 3: +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructUnboxFn(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_23 = lean_ctor_get(x_1, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_1, 1); -lean_inc(x_24); -x_25 = lean_ctor_get(x_1, 2); -lean_inc(x_25); +lean_object* x_5; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_16 = l_Lean_IR_EmitC_emitFnDeclAux___closed__5; +x_17 = lean_string_append(x_4, x_16); +lean_inc(x_2); +x_18 = l_Lean_IR_EmitC_structType(x_2); +x_19 = lean_string_append(x_17, x_18); +x_20 = l_Lean_IR_EmitC_emitSpreadArg___closed__0; +x_21 = lean_string_append(x_19, x_20); +x_22 = l_Lean_IR_EmitC_structUnboxFnPrefix___closed__0; +x_23 = lean_string_append(x_21, x_22); +x_24 = l_Nat_reprFast(x_2); +x_25 = lean_string_append(x_23, x_24); +lean_dec_ref(x_24); +x_26 = l_Lean_IR_EmitC_emitStructUnboxFn___closed__0; +x_27 = lean_string_append(x_25, x_26); +x_28 = l_Lean_IR_EmitC_emitLn___closed__0; +x_29 = lean_string_append(x_27, x_28); +x_30 = lean_string_append(x_29, x_18); +lean_dec_ref(x_18); +x_31 = l_Lean_IR_EmitC_emitStructReshapeFn___closed__1; +x_32 = lean_string_append(x_30, x_31); +x_33 = lean_string_append(x_32, x_28); +if (lean_obj_tag(x_1) == 11) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_34 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_34); lean_dec_ref(x_1); -x_26 = l_Lean_IR_EmitC_emitSetTag___redArg(x_23, x_24, x_3); -if (lean_obj_tag(x_26) == 0) +lean_inc_ref(x_34); +x_35 = lean_alloc_closure((void*)(l_Lean_IR_EmitC_emitStructUnboxFn___lam__0___boxed), 6, 2); +lean_closure_set(x_35, 0, x_34); +lean_closure_set(x_35, 1, x_28); +x_36 = l_Lean_IR_EmitC_emitStructUnboxFn___closed__1; +x_37 = lean_string_append(x_33, x_36); +x_38 = lean_string_append(x_37, x_28); +x_39 = lean_array_get_size(x_34); +lean_dec_ref(x_34); +x_40 = l_Lean_IR_EmitC_emitStructUnboxFn___closed__2; +x_41 = l_Lean_IR_EmitC_emitUnionSwitch(x_39, x_40, x_35, x_3, x_38); +if (lean_obj_tag(x_41) == 0) { -lean_object* x_27; -x_27 = lean_ctor_get(x_26, 1); -lean_inc(x_27); -lean_dec_ref(x_26); -x_1 = x_25; -x_3 = x_27; -goto _start; +lean_object* x_42; +x_42 = lean_ctor_get(x_41, 1); +lean_inc(x_42); +lean_dec_ref(x_41); +x_5 = x_42; +goto block_15; } else { -lean_dec(x_25); -lean_dec_ref(x_2); -return x_26; +return x_41; } } -case 4: +else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_29 = lean_ctor_get(x_1, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_1, 1); -lean_inc(x_30); -x_31 = lean_ctor_get(x_1, 2); -lean_inc(x_31); -x_32 = lean_ctor_get(x_1, 3); -lean_inc(x_32); +if (lean_obj_tag(x_1) == 10) +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_66; uint8_t x_67; +x_43 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_43); +x_44 = lean_ctor_get(x_1, 2); +lean_inc(x_44); +x_45 = lean_ctor_get(x_1, 3); +lean_inc(x_45); lean_dec_ref(x_1); -x_33 = l_Lean_IR_EmitC_emitUSet___redArg(x_29, x_30, x_31, x_3); -if (lean_obj_tag(x_33) == 0) +x_66 = lean_unsigned_to_nat(0u); +x_67 = lean_nat_dec_eq(x_44, x_66); +if (x_67 == 0) { -lean_object* x_34; -x_34 = lean_ctor_get(x_33, 1); -lean_inc(x_34); -lean_dec_ref(x_33); -x_1 = x_32; -x_3 = x_34; -goto _start; +goto block_65; } else { -lean_dec(x_32); -lean_dec_ref(x_2); -return x_33; +uint8_t x_68; +x_68 = lean_nat_dec_eq(x_45, x_66); +if (x_68 == 0) +{ +goto block_65; +} +else +{ +lean_dec(x_45); +lean_dec(x_44); +x_46 = x_3; +x_47 = x_33; +goto block_53; } } -case 5: +block_53: { -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_36 = lean_ctor_get(x_1, 0); -lean_inc(x_36); -x_37 = lean_ctor_get(x_1, 1); -lean_inc(x_37); -x_38 = lean_ctor_get(x_1, 2); -lean_inc(x_38); -x_39 = lean_ctor_get(x_1, 3); -lean_inc(x_39); -x_40 = lean_ctor_get(x_1, 4); -lean_inc(x_40); -x_41 = lean_ctor_get(x_1, 5); -lean_inc(x_41); -lean_dec_ref(x_1); -x_42 = l_Lean_IR_EmitC_emitSSet___redArg(x_36, x_37, x_38, x_39, x_40, x_3); -lean_dec(x_40); -if (lean_obj_tag(x_42) == 0) +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_48 = lean_array_get_size(x_43); +x_49 = lean_unsigned_to_nat(0u); +x_50 = lean_box(0); +x_51 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructUnboxFn_spec__0___redArg(x_50, x_48, x_43, x_49, x_50, x_46, x_47); +lean_dec_ref(x_46); +lean_dec_ref(x_43); +if (lean_obj_tag(x_51) == 0) { -lean_object* x_43; -x_43 = lean_ctor_get(x_42, 1); -lean_inc(x_43); -lean_dec_ref(x_42); -x_1 = x_41; -x_3 = x_43; -goto _start; +lean_object* x_52; +x_52 = lean_ctor_get(x_51, 1); +lean_inc(x_52); +lean_dec_ref(x_51); +x_5 = x_52; +goto block_15; } else { -lean_dec(x_41); -lean_dec_ref(x_2); -return x_42; +return x_51; } } -case 6: +block_65: { -uint8_t x_45; -x_45 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 1); -if (x_45 == 0) +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_54 = l_Lean_IR_EmitC_emitStructUnboxFn___closed__3; +x_55 = lean_string_append(x_33, x_54); +x_56 = l_Nat_reprFast(x_44); +x_57 = lean_string_append(x_55, x_56); +lean_dec_ref(x_56); +x_58 = l_Lean_IR_EmitC_emitStructReshapeFn___closed__4; +x_59 = lean_string_append(x_57, x_58); +x_60 = l_Nat_reprFast(x_45); +x_61 = lean_string_append(x_59, x_60); +lean_dec_ref(x_60); +x_62 = l_Lean_IR_EmitC_emitIncOfType___closed__0; +x_63 = lean_string_append(x_61, x_62); +x_64 = lean_string_append(x_63, x_28); +x_46 = x_3; +x_47 = x_64; +goto block_53; +} +} +else { -lean_object* x_46; lean_object* x_47; uint8_t x_48; lean_object* x_49; lean_object* x_50; -x_46 = lean_ctor_get(x_1, 0); -lean_inc(x_46); -x_47 = lean_ctor_get(x_1, 1); -lean_inc(x_47); -x_48 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); -x_49 = lean_ctor_get(x_1, 2); -lean_inc(x_49); -lean_dec_ref(x_1); -x_50 = l_Lean_IR_EmitC_emitInc___redArg(x_46, x_47, x_48, x_3); -if (lean_obj_tag(x_50) == 0) +lean_object* x_69; lean_object* x_70; +lean_dec(x_1); +x_69 = l_Lean_IR_EmitC_emitStructUnboxFn___closed__5; +x_70 = l_panic___at___00Lean_IR_EmitC_emitUSet_spec__0(x_69, x_3, x_33); +if (lean_obj_tag(x_70) == 0) { -lean_object* x_51; -x_51 = lean_ctor_get(x_50, 1); -lean_inc(x_51); -lean_dec_ref(x_50); -x_1 = x_49; -x_3 = x_51; -goto _start; +lean_object* x_71; +x_71 = lean_ctor_get(x_70, 1); +lean_inc(x_71); +lean_dec_ref(x_70); +x_5 = x_71; +goto block_15; } else { -lean_dec(x_49); -lean_dec_ref(x_2); -return x_50; +return x_70; } } -else +} +block_15: { -lean_object* x_53; -x_53 = lean_ctor_get(x_1, 2); -lean_inc(x_53); -lean_dec_ref(x_1); -x_1 = x_53; -goto _start; +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_6 = l_Lean_IR_EmitC_emitStructReshapeFn___closed__0; +x_7 = lean_string_append(x_5, x_6); +x_8 = l_Lean_IR_EmitC_emitLn___closed__0; +x_9 = lean_string_append(x_7, x_8); +x_10 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; +x_11 = lean_string_append(x_9, x_10); +x_12 = lean_box(0); +x_13 = lean_string_append(x_11, x_8); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +return x_14; } } -case 7: +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructUnboxFn_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -uint8_t x_55; -x_55 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 1); -if (x_55 == 0) +lean_object* x_11; +x_11 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructUnboxFn_spec__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec_ref(x_9); +lean_dec_ref(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructUnboxFn_spec__0___redArg___boxed(lean_object* x_1, lean_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_56; lean_object* x_57; uint8_t x_58; lean_object* x_59; lean_object* x_60; -x_56 = lean_ctor_get(x_1, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_1, 1); -lean_inc(x_57); -x_58 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); -x_59 = lean_ctor_get(x_1, 2); -lean_inc(x_59); -lean_dec_ref(x_1); -x_60 = l_Lean_IR_EmitC_emitDec___redArg(x_56, x_57, x_58, x_3); -if (lean_obj_tag(x_60) == 0) +lean_object* x_8; +x_8 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructUnboxFn_spec__0___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec_ref(x_6); +lean_dec_ref(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructDeclsFor(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_61; -x_61 = lean_ctor_get(x_60, 1); -lean_inc(x_61); -lean_dec_ref(x_60); -x_1 = x_59; -x_3 = x_61; -goto _start; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_12; lean_object* x_13; +x_5 = lean_ctor_get(x_2, 0); +lean_inc(x_5); +lean_dec_ref(x_2); +lean_inc(x_1); +x_12 = l_Lean_IR_EmitC_structType(x_1); +lean_inc_ref(x_3); +lean_inc(x_5); +x_13 = l_Lean_IR_EmitC_emitStructDefn(x_5, x_12, x_3, x_4); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec_ref(x_13); +x_15 = l_Lean_IR_EmitC_needsRC(x_5); +if (x_15 == 0) +{ +x_6 = x_3; +x_7 = x_14; +goto block_11; } else { -lean_dec(x_59); -lean_dec_ref(x_2); -return x_60; +lean_object* x_16; +lean_inc_ref(x_3); +lean_inc(x_1); +lean_inc(x_5); +x_16 = l_Lean_IR_EmitC_emitStructIncDecFn(x_5, x_1, x_15, x_3, x_14); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; uint8_t x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_16, 1); +lean_inc(x_17); +lean_dec_ref(x_16); +x_18 = 0; +lean_inc_ref(x_3); +lean_inc(x_1); +lean_inc(x_5); +x_19 = l_Lean_IR_EmitC_emitStructIncDecFn(x_5, x_1, x_18, x_3, x_17); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; +x_20 = lean_ctor_get(x_19, 1); +lean_inc(x_20); +lean_dec_ref(x_19); +x_6 = x_3; +x_7 = x_20; +goto block_11; +} +else +{ +lean_dec(x_5); +lean_dec_ref(x_3); +lean_dec(x_1); +return x_19; } } else { -lean_object* x_63; -x_63 = lean_ctor_get(x_1, 2); -lean_inc(x_63); -lean_dec_ref(x_1); -x_1 = x_63; -goto _start; +lean_dec(x_5); +lean_dec_ref(x_3); +lean_dec(x_1); +return x_16; } } -case 8: +} +else { -lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_65 = lean_ctor_get(x_1, 0); -lean_inc(x_65); -x_66 = lean_ctor_get(x_1, 1); -lean_inc(x_66); -lean_dec_ref(x_1); -x_67 = l_Lean_IR_EmitC_emitDel___redArg(x_65, x_3); -if (lean_obj_tag(x_67) == 0) +lean_dec(x_5); +lean_dec_ref(x_3); +lean_dec(x_1); +return x_13; +} +block_11: { -lean_object* x_68; -x_68 = lean_ctor_get(x_67, 1); -lean_inc(x_68); -lean_dec_ref(x_67); -x_1 = x_66; -x_3 = x_68; -goto _start; +lean_object* x_8; +lean_inc_ref(x_6); +lean_inc(x_1); +lean_inc(x_5); +x_8 = l_Lean_IR_EmitC_emitStructBoxFn(x_5, x_1, x_6, x_7); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_ctor_get(x_8, 1); +lean_inc(x_9); +lean_dec_ref(x_8); +x_10 = l_Lean_IR_EmitC_emitStructUnboxFn(x_5, x_1, x_6, x_9); +return x_10; } else { -lean_dec(x_66); -lean_dec_ref(x_2); -return x_67; +lean_dec_ref(x_6); +lean_dec(x_5); +lean_dec(x_1); +return x_8; } } -case 9: -{ -lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_70 = lean_ctor_get(x_1, 1); -lean_inc(x_70); -x_71 = lean_ctor_get(x_1, 2); -lean_inc(x_71); -x_72 = lean_ctor_get(x_1, 3); -lean_inc_ref(x_72); -lean_dec_ref(x_1); -x_73 = l_Lean_IR_EmitC_emitCase(x_70, x_71, x_72, x_2, x_3); -lean_dec(x_71); -return x_73; } -case 10: +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDecls_spec__2___redArg(lean_object* x_1, lean_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_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; -lean_dec_ref(x_2); -x_74 = lean_ctor_get(x_1, 0); -lean_inc(x_74); -lean_dec_ref(x_1); -x_75 = l_Lean_IR_EmitC_emitBlock___closed__0; -x_76 = lean_string_append(x_3, x_75); -x_77 = l_Lean_IR_EmitC_emitArg___redArg(x_74, x_76); -if (lean_obj_tag(x_77) == 0) -{ -uint8_t x_78; -x_78 = !lean_is_exclusive(x_77); -if (x_78 == 0) -{ -lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_79 = lean_ctor_get(x_77, 1); -x_80 = lean_ctor_get(x_77, 0); -lean_dec(x_80); -x_81 = l_Lean_IR_EmitC_emitFnDeclAux___closed__0; -x_82 = lean_string_append(x_79, x_81); -x_83 = l_Lean_IR_EmitC_emitLn___closed__0; -x_84 = lean_box(0); -x_85 = lean_string_append(x_82, x_83); -lean_ctor_set(x_77, 1, x_85); -lean_ctor_set(x_77, 0, x_84); -return x_77; +uint8_t x_8; +x_8 = lean_nat_dec_lt(x_4, x_1); +if (x_8 == 0) +{ +lean_object* x_9; +lean_dec_ref(x_6); +lean_dec(x_4); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_5); +lean_ctor_set(x_9, 1, x_7); +return x_9; } else { -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; -x_86 = lean_ctor_get(x_77, 1); -lean_inc(x_86); -lean_dec(x_77); -x_87 = l_Lean_IR_EmitC_emitFnDeclAux___closed__0; -x_88 = lean_string_append(x_86, x_87); -x_89 = l_Lean_IR_EmitC_emitLn___closed__0; -x_90 = lean_box(0); -x_91 = lean_string_append(x_88, x_89); -x_92 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_92, 0, x_90); -lean_ctor_set(x_92, 1, x_91); -return x_92; +lean_object* x_10; lean_object* x_11; +x_10 = lean_array_fget_borrowed(x_2, x_4); +lean_inc_ref(x_6); +lean_inc(x_10); +lean_inc(x_4); +x_11 = l_Lean_IR_EmitC_emitStructDeclsFor(x_4, x_10, x_6, x_7); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +lean_dec_ref(x_11); +x_13 = lean_unsigned_to_nat(1u); +x_14 = lean_nat_add(x_4, x_13); +lean_dec(x_4); +{ +lean_object* _tmp_3 = x_14; +lean_object* _tmp_4 = x_3; +lean_object* _tmp_6 = x_12; +x_4 = _tmp_3; +x_5 = _tmp_4; +x_7 = _tmp_6; } +goto _start; } else { -return x_77; -} +lean_dec_ref(x_6); +lean_dec(x_4); +return x_11; } -case 11: -{ -lean_object* x_93; lean_object* x_94; lean_object* x_95; -x_93 = lean_ctor_get(x_1, 0); -lean_inc(x_93); -x_94 = lean_ctor_get(x_1, 1); -lean_inc_ref(x_94); -lean_dec_ref(x_1); -x_95 = l_Lean_IR_EmitC_emitJmp(x_93, x_94, x_2, x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_94); -return x_95; } -default: -{ -lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; -lean_dec_ref(x_2); -x_96 = l_Lean_IR_EmitC_emitBlock___closed__1; -x_97 = lean_string_append(x_3, x_96); -x_98 = l_Lean_IR_EmitC_emitLn___closed__0; -x_99 = lean_box(0); -x_100 = lean_string_append(x_97, x_98); -x_101 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_101, 0, x_99); -lean_ctor_set(x_101, 1, x_100); -return x_101; } } +static lean_object* _init_l_Lean_IR_EmitC_emitStructDecls___redArg___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_IR_instInhabitedStructTypeInfo_default; +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitJPs(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitStructDecls_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, size_t x_7, size_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -if (lean_obj_tag(x_1) == 1) +uint8_t x_12; +x_12 = lean_usize_dec_lt(x_8, x_7); +if (x_12 == 0) { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); -x_5 = lean_ctor_get(x_1, 2); -lean_inc(x_5); -x_6 = lean_ctor_get(x_1, 3); -lean_inc(x_6); +lean_object* x_13; +lean_dec_ref(x_10); +lean_dec(x_4); +lean_dec_ref(x_2); lean_dec_ref(x_1); -x_7 = l_Lean_IR_EmitC_emitJmp___closed__2; -x_8 = l_Nat_reprFast(x_4); -x_9 = lean_string_append(x_7, x_8); -lean_dec_ref(x_8); -x_10 = lean_string_append(x_3, x_9); -lean_dec_ref(x_9); -x_11 = l_Lean_IR_EmitC_emitJPs___closed__0; -x_12 = lean_string_append(x_10, x_11); -x_13 = l_Lean_IR_EmitC_emitLn___closed__0; -x_14 = lean_string_append(x_12, x_13); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_9); +lean_ctor_set(x_13, 1, x_11); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_14 = lean_ctor_get(x_1, 0); +x_15 = lean_array_uget(x_6, x_8); lean_inc_ref(x_2); -x_15 = l_Lean_IR_EmitC_emitFnBody(x_5, x_2, x_14); -if (lean_obj_tag(x_15) == 0) +x_16 = lean_array_get_borrowed(x_2, x_3, x_15); +x_17 = lean_ctor_get(x_16, 0); +lean_inc_ref(x_10); +lean_inc(x_4); +lean_inc(x_17); +lean_inc(x_14); +x_18 = l_Lean_IR_EmitC_emitStructReshapeFn(x_14, x_17, x_4, x_15, x_10, x_11); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_16; -x_16 = lean_ctor_get(x_15, 1); -lean_inc(x_16); -lean_dec_ref(x_15); -x_1 = x_6; -x_3 = x_16; +lean_object* x_19; size_t x_20; size_t x_21; +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +lean_dec_ref(x_18); +x_20 = 1; +x_21 = lean_usize_add(x_8, x_20); +{ +size_t _tmp_7 = x_21; +lean_object* _tmp_8 = x_5; +lean_object* _tmp_10 = x_19; +x_8 = _tmp_7; +x_9 = _tmp_8; +x_11 = _tmp_10; +} goto _start; } else { -lean_dec(x_6); +lean_dec_ref(x_10); +lean_dec(x_4); lean_dec_ref(x_2); -return x_15; +lean_dec_ref(x_1); +return x_18; +} +} } } +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDecls_spec__1___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +uint8_t x_9; +x_9 = lean_nat_dec_lt(x_5, x_1); +if (x_9 == 0) +{ +lean_object* x_10; +lean_dec_ref(x_7); +lean_dec(x_5); +lean_dec_ref(x_3); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_6); +lean_ctor_set(x_10, 1, x_8); +return x_10; +} else { -uint8_t x_18; -x_18 = l_Lean_IR_FnBody_isTerminal(x_1); -if (x_18 == 0) +lean_object* x_11; lean_object* x_12; size_t x_13; size_t x_14; lean_object* x_15; +x_11 = lean_array_fget_borrowed(x_2, x_5); +x_12 = lean_ctor_get(x_11, 1); +x_13 = lean_array_size(x_12); +x_14 = 0; +lean_inc_ref(x_7); +lean_inc(x_5); +lean_inc_ref(x_3); +lean_inc(x_11); +x_15 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitStructDecls_spec__0(x_11, x_3, x_2, x_5, x_4, x_12, x_13, x_14, x_4, x_7, x_8); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_19; -x_19 = l_Lean_IR_FnBody_body(x_1); -lean_dec(x_1); -x_1 = x_19; +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec_ref(x_15); +x_17 = lean_unsigned_to_nat(1u); +x_18 = lean_nat_add(x_5, x_17); +lean_dec(x_5); +{ +lean_object* _tmp_4 = x_18; +lean_object* _tmp_5 = x_4; +lean_object* _tmp_7 = x_16; +x_5 = _tmp_4; +x_6 = _tmp_5; +x_8 = _tmp_7; +} goto _start; } else { -lean_object* x_21; lean_object* x_22; -lean_dec_ref(x_2); -lean_dec(x_1); -x_21 = lean_box(0); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_3); -return x_22; +lean_dec_ref(x_7); +lean_dec(x_5); +lean_dec_ref(x_3); +return x_15; } } } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFnBody(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitStructDecls___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_4; lean_object* x_5; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; -x_25 = l_Lean_IR_EmitC_emitTailCall___closed__1; -x_26 = lean_string_append(x_3, x_25); -x_27 = l_Lean_IR_EmitC_emitLn___closed__0; -x_28 = lean_string_append(x_26, x_27); -x_29 = 0; -lean_inc(x_1); -x_30 = l_Lean_IR_EmitC_declareVars(x_1, x_29, x_2, x_28); -if (lean_obj_tag(x_30) == 0) +uint8_t x_4; +x_4 = !lean_is_exclusive(x_2); +if (x_4 == 0) { -lean_object* x_31; uint8_t x_32; -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -x_32 = lean_unbox(x_31); -lean_dec(x_31); -if (x_32 == 0) +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_5 = lean_ctor_get(x_2, 0); +x_6 = lean_ctor_get(x_2, 6); +lean_dec(x_6); +lean_inc_ref(x_5); +x_7 = l_Lean_IR_getDecls(x_5); +x_8 = l_Lean_IR_collectStructTypes(x_7); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +lean_dec_ref(x_8); +x_11 = lean_array_get_size(x_9); +x_12 = lean_unsigned_to_nat(0u); +x_13 = lean_box(0); +lean_ctor_set(x_2, 6, x_10); +lean_inc_ref(x_2); +x_14 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDecls_spec__2___redArg(x_11, x_9, x_13, x_12, x_13, x_2, x_3); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_33; -x_33 = lean_ctor_get(x_30, 1); -lean_inc(x_33); -lean_dec_ref(x_30); -x_4 = x_2; -x_5 = x_33; -goto block_24; +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +lean_dec_ref(x_14); +x_16 = l_Lean_IR_EmitC_emitStructDecls___redArg___closed__0; +lean_inc_ref(x_2); +x_17 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDecls_spec__1___redArg(x_11, x_9, x_16, x_13, x_12, x_13, x_2, x_15); +lean_dec(x_9); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +lean_dec_ref(x_17); +x_19 = lean_apply_2(x_1, x_2, x_18); +return x_19; } else { -lean_object* x_34; lean_object* x_35; -x_34 = lean_ctor_get(x_30, 1); -lean_inc(x_34); -lean_dec_ref(x_30); -x_35 = lean_string_append(x_34, x_27); -x_4 = x_2; -x_5 = x_35; -goto block_24; +uint8_t x_20; +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_20 = !lean_is_exclusive(x_17); +if (x_20 == 0) +{ +return x_17; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_17, 0); +x_22 = lean_ctor_get(x_17, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_17); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} } } else { -uint8_t x_36; +uint8_t x_24; lean_dec_ref(x_2); -lean_dec(x_1); -x_36 = !lean_is_exclusive(x_30); -if (x_36 == 0) +lean_dec(x_9); +lean_dec_ref(x_1); +x_24 = !lean_is_exclusive(x_14); +if (x_24 == 0) { -return x_30; +return x_14; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_30, 0); -x_38 = lean_ctor_get(x_30, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_30); -x_39 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_39, 0, x_37); -lean_ctor_set(x_39, 1, x_38); -return x_39; +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_14, 0); +x_26 = lean_ctor_get(x_14, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_14); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; } } -block_24: -{ -lean_object* x_6; -lean_inc_ref(x_4); -lean_inc(x_1); -x_6 = l_Lean_IR_EmitC_emitBlock(x_1, x_4, x_5); -if (lean_obj_tag(x_6) == 0) +} +else { -lean_object* x_7; lean_object* x_8; -x_7 = lean_ctor_get(x_6, 1); -lean_inc(x_7); -lean_dec_ref(x_6); -x_8 = l_Lean_IR_EmitC_emitJPs(x_1, x_4, x_7); -if (lean_obj_tag(x_8) == 0) +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_28 = lean_ctor_get(x_2, 0); +x_29 = lean_ctor_get(x_2, 1); +x_30 = lean_ctor_get(x_2, 2); +x_31 = lean_ctor_get(x_2, 3); +x_32 = lean_ctor_get(x_2, 4); +x_33 = lean_ctor_get(x_2, 5); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_2); +lean_inc_ref(x_28); +x_34 = l_Lean_IR_getDecls(x_28); +x_35 = l_Lean_IR_collectStructTypes(x_34); +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec_ref(x_35); +x_38 = lean_array_get_size(x_36); +x_39 = lean_unsigned_to_nat(0u); +x_40 = lean_box(0); +x_41 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_41, 0, x_28); +lean_ctor_set(x_41, 1, x_29); +lean_ctor_set(x_41, 2, x_30); +lean_ctor_set(x_41, 3, x_31); +lean_ctor_set(x_41, 4, x_32); +lean_ctor_set(x_41, 5, x_33); +lean_ctor_set(x_41, 6, x_37); +lean_inc_ref(x_41); +x_42 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDecls_spec__2___redArg(x_38, x_36, x_40, x_39, x_40, x_41, x_3); +if (lean_obj_tag(x_42) == 0) { -uint8_t x_9; -x_9 = !lean_is_exclusive(x_8); -if (x_9 == 0) +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_42, 1); +lean_inc(x_43); +lean_dec_ref(x_42); +x_44 = l_Lean_IR_EmitC_emitStructDecls___redArg___closed__0; +lean_inc_ref(x_41); +x_45 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDecls_spec__1___redArg(x_38, x_36, x_44, x_40, x_39, x_40, x_41, x_43); +lean_dec(x_36); +if (lean_obj_tag(x_45) == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_10 = lean_ctor_get(x_8, 1); -x_11 = lean_ctor_get(x_8, 0); -lean_dec(x_11); -x_12 = l_Lean_IR_EmitC_emitMainFn___closed__12; -x_13 = lean_string_append(x_10, x_12); -x_14 = l_Lean_IR_EmitC_emitLn___closed__0; -x_15 = lean_box(0); -x_16 = lean_string_append(x_13, x_14); -lean_ctor_set(x_8, 1, x_16); -lean_ctor_set(x_8, 0, x_15); -return x_8; +lean_object* x_46; lean_object* x_47; +x_46 = lean_ctor_get(x_45, 1); +lean_inc(x_46); +lean_dec_ref(x_45); +x_47 = lean_apply_2(x_1, x_41, x_46); +return x_47; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_17 = lean_ctor_get(x_8, 1); -lean_inc(x_17); -lean_dec(x_8); -x_18 = l_Lean_IR_EmitC_emitMainFn___closed__12; -x_19 = lean_string_append(x_17, x_18); -x_20 = l_Lean_IR_EmitC_emitLn___closed__0; -x_21 = lean_box(0); -x_22 = lean_string_append(x_19, x_20); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -return x_23; +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +lean_dec_ref(x_41); +lean_dec_ref(x_1); +x_48 = lean_ctor_get(x_45, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_45, 1); +lean_inc(x_49); +if (lean_is_exclusive(x_45)) { + lean_ctor_release(x_45, 0); + lean_ctor_release(x_45, 1); + x_50 = x_45; +} else { + lean_dec_ref(x_45); + x_50 = lean_box(0); } +if (lean_is_scalar(x_50)) { + x_51 = lean_alloc_ctor(1, 2, 0); +} else { + x_51 = x_50; } -else -{ -return x_8; +lean_ctor_set(x_51, 0, x_48); +lean_ctor_set(x_51, 1, x_49); +return x_51; } } else { -lean_dec_ref(x_4); -lean_dec(x_1); -return x_6; +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +lean_dec_ref(x_41); +lean_dec(x_36); +lean_dec_ref(x_1); +x_52 = lean_ctor_get(x_42, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_42, 1); +lean_inc(x_53); +if (lean_is_exclusive(x_42)) { + lean_ctor_release(x_42, 0); + lean_ctor_release(x_42, 1); + x_54 = x_42; +} else { + lean_dec_ref(x_42); + x_54 = lean_box(0); +} +if (lean_is_scalar(x_54)) { + x_55 = lean_alloc_ctor(1, 2, 0); +} else { + x_55 = x_54; } +lean_ctor_set(x_55, 0, x_52); +lean_ctor_set(x_55, 1, x_53); +return x_55; } } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitIf(lean_object* x_1, lean_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_IR_EmitC_emitStructDecls(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_8 = l_Lean_IR_EmitC_emitIf___closed__0; -x_9 = lean_string_append(x_7, x_8); -x_10 = l_Lean_IR_EmitC_emitTag___redArg(x_1, x_2, x_9); -if (lean_obj_tag(x_10) == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_11 = lean_ctor_get(x_10, 1); -lean_inc(x_11); -lean_dec_ref(x_10); -x_12 = l_Lean_IR_EmitC_emitIf___closed__1; -x_13 = lean_string_append(x_11, x_12); -x_14 = l_Nat_reprFast(x_3); -x_15 = lean_string_append(x_13, x_14); -lean_dec_ref(x_14); -x_16 = l_Lean_IR_EmitC_emitFnDeclAux___closed__1; -x_17 = lean_string_append(x_15, x_16); -x_18 = l_Lean_IR_EmitC_emitLn___closed__0; -x_19 = lean_string_append(x_17, x_18); -lean_inc_ref(x_6); -x_20 = l_Lean_IR_EmitC_emitFnBody(x_4, x_6, x_19); -if (lean_obj_tag(x_20) == 0) +lean_object* x_5; +x_5 = l_Lean_IR_EmitC_emitStructDecls___redArg(x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDecls_spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_21 = lean_ctor_get(x_20, 1); -lean_inc(x_21); -lean_dec_ref(x_20); -x_22 = l_Lean_IR_EmitC_emitIf___closed__2; -x_23 = lean_string_append(x_21, x_22); -x_24 = lean_string_append(x_23, x_18); -x_25 = l_Lean_IR_EmitC_emitFnBody(x_5, x_6, x_24); -return x_25; +lean_object* x_12; +x_12 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDecls_spec__1___redArg(x_1, x_2, x_3, x_4, x_7, x_8, x_10, x_11); +return x_12; } -else +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDecls_spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -lean_dec_ref(x_6); -lean_dec(x_5); -return x_20; +lean_object* x_11; +x_11 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDecls_spec__2___redArg(x_1, x_2, x_3, x_6, x_7, x_9, x_10); +return x_11; } } -else +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDecls_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -lean_dec_ref(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_10; +lean_object* x_12; +x_12 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDecls_spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec_ref(x_2); +lean_dec(x_1); +return x_12; +} } +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDecls_spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDecls_spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec_ref(x_2); +lean_dec(x_1); +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitIf___boxed(lean_object* x_1, lean_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_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDecls_spec__2___redArg___boxed(lean_object* x_1, lean_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_IR_EmitC_emitIf(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_2); +x_8 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDecls_spec__2___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec_ref(x_2); +lean_dec(x_1); return x_8; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitCase_spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDecls_spec__1___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_8 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_9 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitCase_spec__4(x_1, x_7, x_8, x_4, x_5, x_6); -lean_dec_ref(x_1); +lean_object* x_9; +x_9 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDecls_spec__1___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec_ref(x_2); +lean_dec(x_1); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCase___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitStructDecls_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_6; -x_6 = l_Lean_IR_EmitC_emitCase(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_2); -return x_6; +size_t x_12; size_t x_13; lean_object* x_14; +x_12 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_13 = lean_unbox_usize(x_8); +lean_dec(x_8); +x_14 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitStructDecls_spec__0(x_1, x_2, x_3, x_4, x_5, x_6, x_12, x_13, x_9, x_10, x_11); +lean_dec_ref(x_6); +lean_dec_ref(x_3); +return x_14; } } static lean_object* _init_l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__0___redArg___closed__0() { @@ -11887,14 +20468,6 @@ x_1 = lean_mk_string_unchecked(" = _args[", 9, 9); return x_1; } } -static lean_object* _init_l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__0___redArg___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("];", 2, 2); -return x_1; -} -} LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__0___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { @@ -11920,7 +20493,7 @@ lean_dec(x_3); x_11 = lean_nat_sub(x_2, x_10); x_12 = lean_nat_sub(x_11, x_9); lean_dec(x_11); -x_13 = l_Lean_IR_EmitC_emitSimpleExternalCall___closed__0; +x_13 = l_Lean_IR_EmitC_emitFullAppArgs___closed__0; x_14 = lean_array_get_borrowed(x_13, x_1, x_12); x_15 = lean_ctor_get(x_14, 0); x_16 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__0___redArg___closed__0; @@ -11937,7 +20510,7 @@ x_23 = lean_string_append(x_21, x_22); x_24 = l_Nat_reprFast(x_12); x_25 = lean_string_append(x_23, x_24); lean_dec_ref(x_24); -x_26 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__0___redArg___closed__2; +x_26 = l_Lean_IR_EmitC_emitUProj___closed__0; x_27 = lean_string_append(x_25, x_26); x_28 = l_Lean_IR_EmitC_emitLn___closed__0; x_29 = lean_string_append(x_27, x_28); @@ -11955,78 +20528,6 @@ x_7 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_ return x_7; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__1___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; uint8_t x_8; -x_7 = lean_unsigned_to_nat(0u); -x_8 = lean_nat_dec_eq(x_5, x_7); -if (x_8 == 1) -{ -lean_object* x_9; lean_object* x_10; -lean_dec(x_5); -x_9 = lean_box(0); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_6); -return x_10; -} -else -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_28; -x_11 = lean_unsigned_to_nat(1u); -x_12 = lean_nat_sub(x_5, x_11); -lean_dec(x_5); -x_13 = lean_nat_sub(x_4, x_12); -x_14 = lean_nat_sub(x_13, x_11); -lean_dec(x_13); -x_28 = lean_nat_dec_lt(x_3, x_14); -if (x_28 == 0) -{ -x_15 = x_6; -goto block_27; -} -else -{ -lean_object* x_29; lean_object* x_30; -x_29 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___closed__0; -x_30 = lean_string_append(x_6, x_29); -x_15 = x_30; -goto block_27; -} -block_27: -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_16 = lean_array_fget_borrowed(x_1, x_14); -lean_dec(x_14); -x_17 = lean_ctor_get(x_16, 0); -x_18 = lean_ctor_get(x_16, 1); -x_19 = l_Lean_IR_EmitC_toCType(x_18); -x_20 = lean_string_append(x_15, x_19); -lean_dec_ref(x_19); -x_21 = lean_string_append(x_20, x_2); -x_22 = l_Lean_IR_EmitC_argToCString___closed__0; -lean_inc(x_17); -x_23 = l_Nat_reprFast(x_17); -x_24 = lean_string_append(x_22, x_23); -lean_dec_ref(x_23); -x_25 = lean_string_append(x_21, x_24); -lean_dec_ref(x_24); -x_5 = x_12; -x_6 = x_25; -goto _start; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__1___redArg(x_1, x_2, x_3, x_4, x_5, x_8); -return x_9; -} -} static lean_object* _init_l_Lean_IR_EmitC_emitDeclAux___closed__0() { _start: { @@ -12039,19 +20540,11 @@ static lean_object* _init_l_Lean_IR_EmitC_emitDeclAux___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked(" {", 2, 2); -return x_1; -} -} -static lean_object* _init_l_Lean_IR_EmitC_emitDeclAux___closed__2() { -_start: -{ -lean_object* x_1; x_1 = lean_mk_string_unchecked("lean_object** _args", 19, 19); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_emitDeclAux___closed__3() { +static lean_object* _init_l_Lean_IR_EmitC_emitDeclAux___closed__2() { _start: { lean_object* x_1; @@ -12068,187 +20561,220 @@ x_4 = l_Lean_IR_mkVarJPMaps(x_1); x_5 = !lean_is_exclusive(x_4); if (x_5 == 0) { -lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_6 = lean_ctor_get(x_4, 1); -x_7 = lean_ctor_get(x_4, 0); -lean_dec(x_7); -x_8 = !lean_is_exclusive(x_2); -if (x_8 == 0) +uint8_t x_6; +x_6 = !lean_is_exclusive(x_2); +if (x_6 == 0) { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_7 = lean_ctor_get(x_4, 0); +x_8 = lean_ctor_get(x_4, 1); x_9 = lean_ctor_get(x_2, 0); x_10 = lean_ctor_get(x_2, 1); x_11 = lean_ctor_get(x_2, 3); x_12 = lean_ctor_get(x_2, 4); -x_13 = lean_ctor_get(x_2, 2); -lean_dec(x_13); -x_14 = l_Lean_IR_Decl_name(x_1); -lean_inc(x_14); +x_13 = lean_ctor_get(x_2, 6); +x_14 = lean_ctor_get(x_2, 5); +lean_dec(x_14); +x_15 = lean_ctor_get(x_2, 2); +lean_dec(x_15); +x_16 = l_Lean_IR_Decl_name(x_1); +lean_inc(x_16); lean_inc_ref(x_9); -x_15 = l_Lean_hasInitAttr(x_9, x_14); -if (x_15 == 0) +x_17 = l_Lean_hasInitAttr(x_9, x_16); +if (x_17 == 0) { if (lean_obj_tag(x_1) == 0) { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; lean_object* x_67; lean_object* x_68; lean_object* x_78; lean_object* x_79; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; lean_object* x_95; +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; lean_object* x_73; lean_object* x_74; lean_object* x_84; lean_object* x_85; lean_object* x_89; lean_free_object(x_4); -x_16 = lean_ctor_get(x_1, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_1, 1); -lean_inc_ref(x_17); -x_18 = lean_ctor_get(x_1, 2); +x_18 = lean_ctor_get(x_1, 0); lean_inc(x_18); -x_19 = lean_ctor_get(x_1, 3); -lean_inc(x_19); +x_19 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_19); +x_20 = lean_ctor_get(x_1, 2); +lean_inc(x_20); +x_21 = lean_ctor_get(x_1, 3); +lean_inc(x_21); lean_dec_ref(x_1); -lean_ctor_set(x_2, 2, x_6); +lean_ctor_set(x_2, 5, x_7); +lean_ctor_set(x_2, 2, x_8); lean_inc_ref(x_2); -lean_inc(x_16); -x_95 = l_Lean_IR_EmitC_toCName(x_16, x_2, x_3); -if (lean_obj_tag(x_95) == 0) +lean_inc(x_18); +x_89 = l_Lean_IR_EmitC_toCName(x_18, x_2, x_3); +if (lean_obj_tag(x_89) == 0) +{ +lean_object* x_90; lean_object* x_91; uint8_t x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; uint8_t x_96; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_113; lean_object* x_114; lean_object* x_138; lean_object* x_139; uint8_t x_140; +x_90 = lean_ctor_get(x_89, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_89, 1); +lean_inc(x_91); +lean_dec_ref(x_89); +x_92 = 1; +x_138 = lean_array_get_size(x_19); +x_139 = lean_unsigned_to_nat(0u); +x_140 = lean_nat_dec_eq(x_138, x_139); +if (x_140 == 0) +{ +lean_object* x_141; lean_object* x_142; +x_141 = l_Lean_IR_EmitC_emitFnDeclAux___closed__4; +x_142 = lean_string_append(x_91, x_141); +x_113 = x_2; +x_114 = x_142; +goto block_137; +} +else +{ +lean_object* x_143; lean_object* x_144; +x_143 = l_Lean_IR_EmitC_emitFnDeclAux___closed__5; +x_144 = lean_string_append(x_91, x_143); +x_113 = x_2; +x_114 = x_144; +goto block_137; +} +block_101: { -lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_111; lean_object* x_112; lean_object* x_134; lean_object* x_135; uint8_t x_136; -x_96 = lean_ctor_get(x_95, 0); -lean_inc(x_96); -x_97 = lean_ctor_get(x_95, 1); -lean_inc(x_97); -lean_dec_ref(x_95); -x_134 = lean_array_get_size(x_17); -x_135 = lean_unsigned_to_nat(0u); -x_136 = lean_nat_dec_eq(x_134, x_135); -if (x_136 == 0) +if (x_96 == 0) +{ +lean_object* x_97; +x_97 = l_Lean_IR_EmitC_emitSpreadArgs(x_94, x_92, x_93, x_95); +lean_dec_ref(x_94); +if (lean_obj_tag(x_97) == 0) { -lean_object* x_137; lean_object* x_138; -x_137 = l_Lean_IR_EmitC_emitFnDeclAux___closed__7; -x_138 = lean_string_append(x_97, x_137); -x_111 = x_2; -x_112 = x_138; -goto block_133; +lean_object* x_98; +x_98 = lean_ctor_get(x_97, 1); +lean_inc(x_98); +lean_dec_ref(x_97); +x_84 = x_93; +x_85 = x_98; +goto block_88; } else { -lean_object* x_139; lean_object* x_140; -x_139 = l_Lean_IR_EmitC_emitFnDeclAux___closed__8; -x_140 = lean_string_append(x_97, x_139); -x_111 = x_2; -x_112 = x_140; -goto block_133; +lean_dec_ref(x_93); +lean_dec(x_21); +lean_dec_ref(x_19); +lean_dec(x_18); +lean_dec(x_16); +return x_97; } -block_110: +} +else { -lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; uint8_t x_108; -x_103 = lean_string_append(x_99, x_96); -lean_dec(x_96); -x_104 = l_Lean_IR_EmitC_emitFnDeclAux___closed__5; -x_105 = lean_string_append(x_103, x_104); -x_106 = l_Lean_IR_EmitC_emitFnDeclAux___closed__3; -x_107 = lean_array_get_size(x_102); -x_108 = lean_nat_dec_lt(x_106, x_107); -if (x_108 == 0) +lean_object* x_99; lean_object* x_100; +lean_dec_ref(x_94); +x_99 = l_Lean_IR_EmitC_emitDeclAux___closed__1; +x_100 = lean_string_append(x_95, x_99); +x_84 = x_93; +x_85 = x_100; +goto block_88; +} +} +block_112: +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; uint8_t x_110; +x_105 = lean_string_append(x_103, x_90); +lean_dec(x_90); +x_106 = l_Lean_IR_EmitC_emitFnDeclAux___closed__2; +x_107 = lean_string_append(x_105, x_106); +x_108 = l_Lean_IR_EmitC_emitFnDeclAux___closed__1; +x_109 = lean_array_get_size(x_104); +x_110 = lean_nat_dec_lt(x_108, x_109); +if (x_110 == 0) { -x_83 = x_98; -x_84 = x_105; -x_85 = x_107; -x_86 = x_100; -x_87 = x_101; -x_88 = x_102; -x_89 = x_108; -goto block_94; +x_93 = x_102; +x_94 = x_104; +x_95 = x_107; +x_96 = x_110; +goto block_101; } else { -uint8_t x_109; -x_109 = l_Lean_IR_ExplicitBoxing_isBoxedName(x_14); -x_83 = x_98; -x_84 = x_105; -x_85 = x_107; -x_86 = x_100; -x_87 = x_101; -x_88 = x_102; -x_89 = x_109; -goto block_94; +uint8_t x_111; +x_111 = l_Lean_IR_ExplicitBoxing_isBoxedName(x_16); +x_93 = x_102; +x_94 = x_104; +x_95 = x_107; +x_96 = x_111; +goto block_101; } } -block_133: +block_137: { -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_113 = l_Lean_IR_EmitC_toCType(x_18); -lean_dec(x_18); -x_114 = lean_string_append(x_112, x_113); -lean_dec_ref(x_113); -x_115 = l_Lean_IR_EmitC_emitFnDeclAux___closed__4; -x_116 = lean_string_append(x_114, x_115); -x_117 = lean_unsigned_to_nat(0u); -x_118 = lean_array_get_size(x_17); -x_119 = lean_nat_dec_lt(x_117, x_118); -if (x_119 == 0) -{ -lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; -x_120 = l_Lean_IR_EmitC_toCInitName___closed__0; -x_121 = lean_string_append(x_120, x_96); -lean_dec(x_96); -x_122 = l_Lean_IR_EmitC_emitDeclAux___closed__3; -x_123 = lean_string_append(x_121, x_122); -x_124 = lean_string_append(x_116, x_123); -lean_dec_ref(x_123); -x_67 = x_111; -x_68 = x_124; -goto block_77; +lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; uint8_t x_123; +x_115 = l_Lean_IR_EmitC_toCType(x_20, x_113, x_114); +lean_dec(x_20); +x_116 = lean_ctor_get(x_115, 0); +lean_inc(x_116); +x_117 = lean_ctor_get(x_115, 1); +lean_inc(x_117); +lean_dec_ref(x_115); +x_118 = lean_string_append(x_117, x_116); +lean_dec(x_116); +x_119 = l_Lean_IR_EmitC_emitSpreadArg___closed__0; +x_120 = lean_string_append(x_118, x_119); +x_121 = lean_unsigned_to_nat(0u); +x_122 = lean_array_get_size(x_19); +x_123 = lean_nat_dec_lt(x_121, x_122); +if (x_123 == 0) +{ +lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; +x_124 = l_Lean_IR_EmitC_toCInitName___closed__0; +x_125 = lean_string_append(x_124, x_90); +lean_dec(x_90); +x_126 = l_Lean_IR_EmitC_emitDeclAux___closed__2; +x_127 = lean_string_append(x_125, x_126); +x_128 = lean_string_append(x_120, x_127); +lean_dec_ref(x_127); +x_73 = x_113; +x_74 = x_128; +goto block_83; } else { -lean_object* x_125; -x_125 = l_Lean_IR_EmitC_emitFnDeclAux___closed__6; -if (x_119 == 0) +lean_object* x_129; +x_129 = l_Lean_IR_EmitC_emitFnDeclAux___closed__3; +if (x_123 == 0) { -x_98 = x_117; -x_99 = x_116; -x_100 = x_115; -x_101 = x_111; -x_102 = x_125; -goto block_110; +x_102 = x_113; +x_103 = x_120; +x_104 = x_129; +goto block_112; } else { -uint8_t x_126; -x_126 = lean_nat_dec_le(x_118, x_118); -if (x_126 == 0) +uint8_t x_130; +x_130 = lean_nat_dec_le(x_122, x_122); +if (x_130 == 0) { -if (x_119 == 0) +if (x_123 == 0) { -x_98 = x_117; -x_99 = x_116; -x_100 = x_115; -x_101 = x_111; -x_102 = x_125; -goto block_110; +x_102 = x_113; +x_103 = x_120; +x_104 = x_129; +goto block_112; } else { -size_t x_127; size_t x_128; lean_object* x_129; -x_127 = 0; -x_128 = lean_usize_of_nat(x_118); -x_129 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__2(x_17, x_127, x_128, x_125); -x_98 = x_117; -x_99 = x_116; -x_100 = x_115; -x_101 = x_111; -x_102 = x_129; -goto block_110; +size_t x_131; size_t x_132; lean_object* x_133; +x_131 = 0; +x_132 = lean_usize_of_nat(x_122); +x_133 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__1(x_19, x_131, x_132, x_129); +x_102 = x_113; +x_103 = x_120; +x_104 = x_133; +goto block_112; } } else { -size_t x_130; size_t x_131; lean_object* x_132; -x_130 = 0; -x_131 = lean_usize_of_nat(x_118); -x_132 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__2(x_17, x_130, x_131, x_125); -x_98 = x_117; -x_99 = x_116; -x_100 = x_115; -x_101 = x_111; -x_102 = x_132; -goto block_110; +size_t x_134; size_t x_135; lean_object* x_136; +x_134 = 0; +x_135 = lean_usize_of_nat(x_122); +x_136 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__1(x_19, x_134, x_135, x_129); +x_102 = x_113; +x_103 = x_120; +x_104 = x_136; +goto block_112; } } } @@ -12256,454 +20782,486 @@ goto block_110; } else { -uint8_t x_141; +uint8_t x_145; lean_dec_ref(x_2); -lean_dec(x_19); +lean_dec(x_21); +lean_dec(x_20); +lean_dec_ref(x_19); lean_dec(x_18); -lean_dec_ref(x_17); lean_dec(x_16); -lean_dec(x_14); -x_141 = !lean_is_exclusive(x_95); -if (x_141 == 0) +x_145 = !lean_is_exclusive(x_89); +if (x_145 == 0) { -return x_95; +return x_89; } else { -lean_object* x_142; lean_object* x_143; lean_object* x_144; -x_142 = lean_ctor_get(x_95, 0); -x_143 = lean_ctor_get(x_95, 1); -lean_inc(x_143); -lean_inc(x_142); -lean_dec(x_95); -x_144 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_144, 0, x_142); -lean_ctor_set(x_144, 1, x_143); -return x_144; +lean_object* x_146; lean_object* x_147; lean_object* x_148; +x_146 = lean_ctor_get(x_89, 0); +x_147 = lean_ctor_get(x_89, 1); +lean_inc(x_147); +lean_inc(x_146); +lean_dec(x_89); +x_148 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_148, 0, x_146); +lean_ctor_set(x_148, 1, x_147); +return x_148; } } -block_59: +block_65: { -uint8_t x_22; -x_22 = !lean_is_exclusive(x_20); -if (x_22 == 0) +lean_object* x_24; +x_24 = l_Lean_IR_EmitC_emitSpreads(x_19, x_22, x_23); +if (lean_obj_tag(x_24) == 0) { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_23 = lean_ctor_get(x_20, 4); -lean_dec(x_23); -x_24 = lean_ctor_get(x_20, 3); -lean_dec(x_24); -x_25 = l_Lean_IR_EmitC_emitDeclAux___closed__0; -x_26 = lean_string_append(x_21, x_25); -x_27 = l_Lean_IR_EmitC_emitLn___closed__0; -x_28 = lean_string_append(x_26, x_27); -lean_ctor_set(x_20, 4, x_17); -lean_ctor_set(x_20, 3, x_16); -x_29 = l_Lean_IR_EmitC_emitFnBody(x_19, x_20, x_28); -if (lean_obj_tag(x_29) == 0) +lean_object* x_25; uint8_t x_26; +x_25 = lean_ctor_get(x_24, 1); +lean_inc(x_25); +lean_dec_ref(x_24); +x_26 = !lean_is_exclusive(x_22); +if (x_26 == 0) { -uint8_t x_30; -x_30 = !lean_is_exclusive(x_29); -if (x_30 == 0) +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_27 = lean_ctor_get(x_22, 4); +lean_dec(x_27); +x_28 = lean_ctor_get(x_22, 3); +lean_dec(x_28); +x_29 = l_Lean_IR_EmitC_emitDeclAux___closed__0; +x_30 = lean_string_append(x_25, x_29); +x_31 = l_Lean_IR_EmitC_emitLn___closed__0; +x_32 = lean_string_append(x_30, x_31); +lean_ctor_set(x_22, 4, x_19); +lean_ctor_set(x_22, 3, x_18); +x_33 = l_Lean_IR_EmitC_emitFnBody(x_21, x_22, x_32); +if (lean_obj_tag(x_33) == 0) { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_31 = lean_ctor_get(x_29, 1); -x_32 = lean_ctor_get(x_29, 0); -lean_dec(x_32); -x_33 = l_Lean_IR_EmitC_emitMainFn___closed__12; -x_34 = lean_string_append(x_31, x_33); -x_35 = lean_box(0); -x_36 = lean_string_append(x_34, x_27); -lean_ctor_set(x_29, 1, x_36); -lean_ctor_set(x_29, 0, x_35); -return x_29; +uint8_t x_34; +x_34 = !lean_is_exclusive(x_33); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_35 = lean_ctor_get(x_33, 1); +x_36 = lean_ctor_get(x_33, 0); +lean_dec(x_36); +x_37 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; +x_38 = lean_string_append(x_35, x_37); +x_39 = lean_box(0); +x_40 = lean_string_append(x_38, x_31); +lean_ctor_set(x_33, 1, x_40); +lean_ctor_set(x_33, 0, x_39); +return x_33; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_37 = lean_ctor_get(x_29, 1); -lean_inc(x_37); -lean_dec(x_29); -x_38 = l_Lean_IR_EmitC_emitMainFn___closed__12; -x_39 = lean_string_append(x_37, x_38); -x_40 = lean_box(0); -x_41 = lean_string_append(x_39, x_27); -x_42 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -return x_42; +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_41 = lean_ctor_get(x_33, 1); +lean_inc(x_41); +lean_dec(x_33); +x_42 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; +x_43 = lean_string_append(x_41, x_42); +x_44 = lean_box(0); +x_45 = lean_string_append(x_43, x_31); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; } } else { -return x_29; +return x_33; } } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_43 = lean_ctor_get(x_20, 0); -x_44 = lean_ctor_get(x_20, 1); -x_45 = lean_ctor_get(x_20, 2); -lean_inc(x_45); -lean_inc(x_44); -lean_inc(x_43); -lean_dec(x_20); -x_46 = l_Lean_IR_EmitC_emitDeclAux___closed__0; -x_47 = lean_string_append(x_21, x_46); -x_48 = l_Lean_IR_EmitC_emitLn___closed__0; -x_49 = lean_string_append(x_47, x_48); -x_50 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_50, 0, x_43); -lean_ctor_set(x_50, 1, x_44); -lean_ctor_set(x_50, 2, x_45); -lean_ctor_set(x_50, 3, x_16); -lean_ctor_set(x_50, 4, x_17); -x_51 = l_Lean_IR_EmitC_emitFnBody(x_19, x_50, x_49); -if (lean_obj_tag(x_51) == 0) -{ -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_52 = lean_ctor_get(x_51, 1); -lean_inc(x_52); -if (lean_is_exclusive(x_51)) { - lean_ctor_release(x_51, 0); - lean_ctor_release(x_51, 1); - x_53 = x_51; +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_47 = lean_ctor_get(x_22, 0); +x_48 = lean_ctor_get(x_22, 1); +x_49 = lean_ctor_get(x_22, 2); +x_50 = lean_ctor_get(x_22, 5); +x_51 = lean_ctor_get(x_22, 6); +lean_inc(x_51); +lean_inc(x_50); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_dec(x_22); +x_52 = l_Lean_IR_EmitC_emitDeclAux___closed__0; +x_53 = lean_string_append(x_25, x_52); +x_54 = l_Lean_IR_EmitC_emitLn___closed__0; +x_55 = lean_string_append(x_53, x_54); +x_56 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_56, 0, x_47); +lean_ctor_set(x_56, 1, x_48); +lean_ctor_set(x_56, 2, x_49); +lean_ctor_set(x_56, 3, x_18); +lean_ctor_set(x_56, 4, x_19); +lean_ctor_set(x_56, 5, x_50); +lean_ctor_set(x_56, 6, x_51); +x_57 = l_Lean_IR_EmitC_emitFnBody(x_21, x_56, x_55); +if (lean_obj_tag(x_57) == 0) +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_58 = lean_ctor_get(x_57, 1); +lean_inc(x_58); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + lean_ctor_release(x_57, 1); + x_59 = x_57; } else { - lean_dec_ref(x_51); - x_53 = lean_box(0); -} -x_54 = l_Lean_IR_EmitC_emitMainFn___closed__12; -x_55 = lean_string_append(x_52, x_54); -x_56 = lean_box(0); -x_57 = lean_string_append(x_55, x_48); -if (lean_is_scalar(x_53)) { - x_58 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_57); + x_59 = lean_box(0); +} +x_60 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; +x_61 = lean_string_append(x_58, x_60); +x_62 = lean_box(0); +x_63 = lean_string_append(x_61, x_54); +if (lean_is_scalar(x_59)) { + x_64 = lean_alloc_ctor(0, 2, 0); } else { - x_58 = x_53; + x_64 = x_59; } -lean_ctor_set(x_58, 0, x_56); -lean_ctor_set(x_58, 1, x_57); -return x_58; +lean_ctor_set(x_64, 0, x_62); +lean_ctor_set(x_64, 1, x_63); +return x_64; } else { -return x_51; -} +return x_57; } } -block_66: -{ -if (x_63 == 0) -{ -lean_dec(x_62); -x_20 = x_61; -x_21 = x_60; -goto block_59; } else { -lean_object* x_64; lean_object* x_65; -lean_inc(x_62); -x_64 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__0___redArg(x_17, x_62, x_62, x_60); -lean_dec(x_62); -x_65 = lean_ctor_get(x_64, 1); -lean_inc(x_65); -lean_dec_ref(x_64); -x_20 = x_61; -x_21 = x_65; -goto block_59; +lean_dec_ref(x_22); +lean_dec(x_21); +lean_dec_ref(x_19); +lean_dec(x_18); +return x_24; } } -block_77: +block_72: { -lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; -x_69 = l_Lean_IR_EmitC_emitDeclAux___closed__1; -x_70 = lean_string_append(x_68, x_69); -x_71 = l_Lean_IR_EmitC_emitLn___closed__0; -x_72 = lean_string_append(x_70, x_71); -x_73 = l_Lean_IR_EmitC_emitFnDeclAux___closed__3; -x_74 = lean_array_get_size(x_17); -x_75 = lean_nat_dec_lt(x_73, x_74); -if (x_75 == 0) +if (x_69 == 0) { -lean_dec(x_14); -x_60 = x_72; -x_61 = x_67; -x_62 = x_74; -x_63 = x_75; -goto block_66; +lean_dec(x_68); +x_22 = x_66; +x_23 = x_67; +goto block_65; } else { -uint8_t x_76; -x_76 = l_Lean_IR_ExplicitBoxing_isBoxedName(x_14); -lean_dec(x_14); -x_60 = x_72; -x_61 = x_67; -x_62 = x_74; -x_63 = x_76; -goto block_66; -} +lean_object* x_70; lean_object* x_71; +lean_inc(x_68); +x_70 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__0___redArg(x_19, x_68, x_68, x_67); +lean_dec(x_68); +x_71 = lean_ctor_get(x_70, 1); +lean_inc(x_71); +lean_dec_ref(x_70); +x_22 = x_66; +x_23 = x_71; +goto block_65; } -block_82: -{ -lean_object* x_80; lean_object* x_81; -x_80 = l_Lean_IR_EmitC_emitFnDeclAux___closed__1; -x_81 = lean_string_append(x_79, x_80); -x_67 = x_78; -x_68 = x_81; -goto block_77; } -block_94: +block_83: { -if (x_89 == 0) +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; uint8_t x_81; +x_75 = l_Lean_IR_EmitC_emitStructDefn___closed__0; +x_76 = lean_string_append(x_74, x_75); +x_77 = l_Lean_IR_EmitC_emitLn___closed__0; +x_78 = lean_string_append(x_76, x_77); +x_79 = l_Lean_IR_EmitC_emitFnDeclAux___closed__1; +x_80 = lean_array_get_size(x_19); +x_81 = lean_nat_dec_lt(x_79, x_80); +if (x_81 == 0) { -lean_object* x_90; lean_object* x_91; -lean_inc(x_85); -x_90 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__1___redArg(x_88, x_86, x_83, x_85, x_85, x_84); -lean_dec(x_85); -lean_dec_ref(x_86); -lean_dec_ref(x_88); -x_91 = lean_ctor_get(x_90, 1); -lean_inc(x_91); -lean_dec_ref(x_90); -x_78 = x_87; -x_79 = x_91; -goto block_82; +lean_dec(x_16); +x_66 = x_73; +x_67 = x_78; +x_68 = x_80; +x_69 = x_81; +goto block_72; } else { -lean_object* x_92; lean_object* x_93; -lean_dec_ref(x_88); -lean_dec_ref(x_86); -lean_dec(x_85); -x_92 = l_Lean_IR_EmitC_emitDeclAux___closed__2; -x_93 = lean_string_append(x_84, x_92); -x_78 = x_87; -x_79 = x_93; -goto block_82; +uint8_t x_82; +x_82 = l_Lean_IR_ExplicitBoxing_isBoxedName(x_16); +lean_dec(x_16); +x_66 = x_73; +x_67 = x_78; +x_68 = x_80; +x_69 = x_82; +goto block_72; } } +block_88: +{ +lean_object* x_86; lean_object* x_87; +x_86 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__1; +x_87 = lean_string_append(x_85, x_86); +x_73 = x_84; +x_74 = x_87; +goto block_83; +} } else { -lean_object* x_145; -lean_dec(x_14); +lean_object* x_149; +lean_dec(x_16); lean_free_object(x_2); +lean_dec_ref(x_13); lean_dec_ref(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec_ref(x_9); -lean_dec(x_6); +lean_dec(x_8); +lean_dec(x_7); lean_dec_ref(x_1); -x_145 = lean_box(0); +x_149 = lean_box(0); lean_ctor_set(x_4, 1, x_3); -lean_ctor_set(x_4, 0, x_145); +lean_ctor_set(x_4, 0, x_149); return x_4; } } else { -lean_object* x_146; -lean_dec(x_14); +lean_object* x_150; +lean_dec(x_16); lean_free_object(x_2); +lean_dec_ref(x_13); lean_dec_ref(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec_ref(x_9); -lean_dec(x_6); +lean_dec(x_8); +lean_dec(x_7); lean_dec_ref(x_1); -x_146 = lean_box(0); +x_150 = lean_box(0); lean_ctor_set(x_4, 1, x_3); -lean_ctor_set(x_4, 0, x_146); +lean_ctor_set(x_4, 0, x_150); return x_4; } } else { -lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; uint8_t x_152; -x_147 = lean_ctor_get(x_2, 0); -x_148 = lean_ctor_get(x_2, 1); -x_149 = lean_ctor_get(x_2, 3); -x_150 = lean_ctor_get(x_2, 4); -lean_inc(x_150); -lean_inc(x_149); -lean_inc(x_148); -lean_inc(x_147); +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_ctor_get(x_4, 0); +x_152 = lean_ctor_get(x_4, 1); +x_153 = lean_ctor_get(x_2, 0); +x_154 = lean_ctor_get(x_2, 1); +x_155 = lean_ctor_get(x_2, 3); +x_156 = lean_ctor_get(x_2, 4); +x_157 = lean_ctor_get(x_2, 6); +lean_inc(x_157); +lean_inc(x_156); +lean_inc(x_155); +lean_inc(x_154); +lean_inc(x_153); lean_dec(x_2); -x_151 = l_Lean_IR_Decl_name(x_1); -lean_inc(x_151); -lean_inc_ref(x_147); -x_152 = l_Lean_hasInitAttr(x_147, x_151); -if (x_152 == 0) +x_158 = l_Lean_IR_Decl_name(x_1); +lean_inc(x_158); +lean_inc_ref(x_153); +x_159 = l_Lean_hasInitAttr(x_153, x_158); +if (x_159 == 0) { if (lean_obj_tag(x_1) == 0) { -lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_177; lean_object* x_178; lean_object* x_179; uint8_t x_180; lean_object* x_184; lean_object* x_185; lean_object* x_195; lean_object* x_196; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; uint8_t x_206; lean_object* x_212; lean_object* x_213; +lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_188; lean_object* x_189; lean_object* x_190; uint8_t x_191; lean_object* x_195; lean_object* x_196; lean_object* x_206; lean_object* x_207; lean_object* x_211; lean_object* x_212; lean_free_object(x_4); -x_153 = lean_ctor_get(x_1, 0); -lean_inc(x_153); -x_154 = lean_ctor_get(x_1, 1); -lean_inc_ref(x_154); -x_155 = lean_ctor_get(x_1, 2); -lean_inc(x_155); -x_156 = lean_ctor_get(x_1, 3); -lean_inc(x_156); +x_160 = lean_ctor_get(x_1, 0); +lean_inc(x_160); +x_161 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_161); +x_162 = lean_ctor_get(x_1, 2); +lean_inc(x_162); +x_163 = lean_ctor_get(x_1, 3); +lean_inc(x_163); lean_dec_ref(x_1); -x_212 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_212, 0, x_147); -lean_ctor_set(x_212, 1, x_148); -lean_ctor_set(x_212, 2, x_6); -lean_ctor_set(x_212, 3, x_149); -lean_ctor_set(x_212, 4, x_150); -lean_inc_ref(x_212); -lean_inc(x_153); -x_213 = l_Lean_IR_EmitC_toCName(x_153, x_212, x_3); -if (lean_obj_tag(x_213) == 0) +x_211 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_211, 0, x_153); +lean_ctor_set(x_211, 1, x_154); +lean_ctor_set(x_211, 2, x_152); +lean_ctor_set(x_211, 3, x_155); +lean_ctor_set(x_211, 4, x_156); +lean_ctor_set(x_211, 5, x_151); +lean_ctor_set(x_211, 6, x_157); +lean_inc_ref(x_211); +lean_inc(x_160); +x_212 = l_Lean_IR_EmitC_toCName(x_160, x_211, x_3); +if (lean_obj_tag(x_212) == 0) { -lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_229; lean_object* x_230; lean_object* x_252; lean_object* x_253; uint8_t x_254; -x_214 = lean_ctor_get(x_213, 0); +lean_object* x_213; lean_object* x_214; uint8_t x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; uint8_t x_219; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_236; lean_object* x_237; lean_object* x_261; lean_object* x_262; uint8_t x_263; +x_213 = lean_ctor_get(x_212, 0); +lean_inc(x_213); +x_214 = lean_ctor_get(x_212, 1); lean_inc(x_214); -x_215 = lean_ctor_get(x_213, 1); -lean_inc(x_215); -lean_dec_ref(x_213); -x_252 = lean_array_get_size(x_154); -x_253 = lean_unsigned_to_nat(0u); -x_254 = lean_nat_dec_eq(x_252, x_253); -if (x_254 == 0) -{ -lean_object* x_255; lean_object* x_256; -x_255 = l_Lean_IR_EmitC_emitFnDeclAux___closed__7; -x_256 = lean_string_append(x_215, x_255); -x_229 = x_212; -x_230 = x_256; -goto block_251; -} -else -{ -lean_object* x_257; lean_object* x_258; -x_257 = l_Lean_IR_EmitC_emitFnDeclAux___closed__8; -x_258 = lean_string_append(x_215, x_257); -x_229 = x_212; -x_230 = x_258; -goto block_251; -} -block_228: -{ -lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; uint8_t x_226; -x_221 = lean_string_append(x_217, x_214); -lean_dec(x_214); -x_222 = l_Lean_IR_EmitC_emitFnDeclAux___closed__5; -x_223 = lean_string_append(x_221, x_222); -x_224 = l_Lean_IR_EmitC_emitFnDeclAux___closed__3; -x_225 = lean_array_get_size(x_220); -x_226 = lean_nat_dec_lt(x_224, x_225); -if (x_226 == 0) -{ -x_200 = x_216; -x_201 = x_223; -x_202 = x_225; -x_203 = x_218; -x_204 = x_219; -x_205 = x_220; -x_206 = x_226; -goto block_211; -} -else -{ -uint8_t x_227; -x_227 = l_Lean_IR_ExplicitBoxing_isBoxedName(x_151); -x_200 = x_216; -x_201 = x_223; -x_202 = x_225; -x_203 = x_218; -x_204 = x_219; -x_205 = x_220; -x_206 = x_227; -goto block_211; -} -} -block_251: -{ -lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; uint8_t x_237; -x_231 = l_Lean_IR_EmitC_toCType(x_155); -lean_dec(x_155); -x_232 = lean_string_append(x_230, x_231); -lean_dec_ref(x_231); -x_233 = l_Lean_IR_EmitC_emitFnDeclAux___closed__4; -x_234 = lean_string_append(x_232, x_233); -x_235 = lean_unsigned_to_nat(0u); -x_236 = lean_array_get_size(x_154); -x_237 = lean_nat_dec_lt(x_235, x_236); -if (x_237 == 0) -{ -lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; -x_238 = l_Lean_IR_EmitC_toCInitName___closed__0; -x_239 = lean_string_append(x_238, x_214); -lean_dec(x_214); -x_240 = l_Lean_IR_EmitC_emitDeclAux___closed__3; -x_241 = lean_string_append(x_239, x_240); -x_242 = lean_string_append(x_234, x_241); -lean_dec_ref(x_241); -x_184 = x_229; -x_185 = x_242; -goto block_194; +lean_dec_ref(x_212); +x_215 = 1; +x_261 = lean_array_get_size(x_161); +x_262 = lean_unsigned_to_nat(0u); +x_263 = lean_nat_dec_eq(x_261, x_262); +if (x_263 == 0) +{ +lean_object* x_264; lean_object* x_265; +x_264 = l_Lean_IR_EmitC_emitFnDeclAux___closed__4; +x_265 = lean_string_append(x_214, x_264); +x_236 = x_211; +x_237 = x_265; +goto block_260; +} +else +{ +lean_object* x_266; lean_object* x_267; +x_266 = l_Lean_IR_EmitC_emitFnDeclAux___closed__5; +x_267 = lean_string_append(x_214, x_266); +x_236 = x_211; +x_237 = x_267; +goto block_260; +} +block_224: +{ +if (x_219 == 0) +{ +lean_object* x_220; +x_220 = l_Lean_IR_EmitC_emitSpreadArgs(x_217, x_215, x_216, x_218); +lean_dec_ref(x_217); +if (lean_obj_tag(x_220) == 0) +{ +lean_object* x_221; +x_221 = lean_ctor_get(x_220, 1); +lean_inc(x_221); +lean_dec_ref(x_220); +x_206 = x_216; +x_207 = x_221; +goto block_210; +} +else +{ +lean_dec_ref(x_216); +lean_dec(x_163); +lean_dec_ref(x_161); +lean_dec(x_160); +lean_dec(x_158); +return x_220; +} +} +else +{ +lean_object* x_222; lean_object* x_223; +lean_dec_ref(x_217); +x_222 = l_Lean_IR_EmitC_emitDeclAux___closed__1; +x_223 = lean_string_append(x_218, x_222); +x_206 = x_216; +x_207 = x_223; +goto block_210; +} +} +block_235: +{ +lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; uint8_t x_233; +x_228 = lean_string_append(x_226, x_213); +lean_dec(x_213); +x_229 = l_Lean_IR_EmitC_emitFnDeclAux___closed__2; +x_230 = lean_string_append(x_228, x_229); +x_231 = l_Lean_IR_EmitC_emitFnDeclAux___closed__1; +x_232 = lean_array_get_size(x_227); +x_233 = lean_nat_dec_lt(x_231, x_232); +if (x_233 == 0) +{ +x_216 = x_225; +x_217 = x_227; +x_218 = x_230; +x_219 = x_233; +goto block_224; +} +else +{ +uint8_t x_234; +x_234 = l_Lean_IR_ExplicitBoxing_isBoxedName(x_158); +x_216 = x_225; +x_217 = x_227; +x_218 = x_230; +x_219 = x_234; +goto block_224; +} +} +block_260: +{ +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_238 = l_Lean_IR_EmitC_toCType(x_162, x_236, x_237); +lean_dec(x_162); +x_239 = lean_ctor_get(x_238, 0); +lean_inc(x_239); +x_240 = lean_ctor_get(x_238, 1); +lean_inc(x_240); +lean_dec_ref(x_238); +x_241 = lean_string_append(x_240, x_239); +lean_dec(x_239); +x_242 = l_Lean_IR_EmitC_emitSpreadArg___closed__0; +x_243 = lean_string_append(x_241, x_242); +x_244 = lean_unsigned_to_nat(0u); +x_245 = lean_array_get_size(x_161); +x_246 = lean_nat_dec_lt(x_244, x_245); +if (x_246 == 0) +{ +lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; +x_247 = l_Lean_IR_EmitC_toCInitName___closed__0; +x_248 = lean_string_append(x_247, x_213); +lean_dec(x_213); +x_249 = l_Lean_IR_EmitC_emitDeclAux___closed__2; +x_250 = lean_string_append(x_248, x_249); +x_251 = lean_string_append(x_243, x_250); +lean_dec_ref(x_250); +x_195 = x_236; +x_196 = x_251; +goto block_205; } else { -lean_object* x_243; -x_243 = l_Lean_IR_EmitC_emitFnDeclAux___closed__6; -if (x_237 == 0) +lean_object* x_252; +x_252 = l_Lean_IR_EmitC_emitFnDeclAux___closed__3; +if (x_246 == 0) { -x_216 = x_235; -x_217 = x_234; -x_218 = x_233; -x_219 = x_229; -x_220 = x_243; -goto block_228; +x_225 = x_236; +x_226 = x_243; +x_227 = x_252; +goto block_235; } else { -uint8_t x_244; -x_244 = lean_nat_dec_le(x_236, x_236); -if (x_244 == 0) +uint8_t x_253; +x_253 = lean_nat_dec_le(x_245, x_245); +if (x_253 == 0) { -if (x_237 == 0) +if (x_246 == 0) { -x_216 = x_235; -x_217 = x_234; -x_218 = x_233; -x_219 = x_229; -x_220 = x_243; -goto block_228; +x_225 = x_236; +x_226 = x_243; +x_227 = x_252; +goto block_235; } else { -size_t x_245; size_t x_246; lean_object* x_247; -x_245 = 0; -x_246 = lean_usize_of_nat(x_236); -x_247 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__2(x_154, x_245, x_246, x_243); -x_216 = x_235; -x_217 = x_234; -x_218 = x_233; -x_219 = x_229; -x_220 = x_247; -goto block_228; +size_t x_254; size_t x_255; lean_object* x_256; +x_254 = 0; +x_255 = lean_usize_of_nat(x_245); +x_256 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__1(x_161, x_254, x_255, x_252); +x_225 = x_236; +x_226 = x_243; +x_227 = x_256; +goto block_235; } } else { -size_t x_248; size_t x_249; lean_object* x_250; -x_248 = 0; -x_249 = lean_usize_of_nat(x_236); -x_250 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__2(x_154, x_248, x_249, x_243); -x_216 = x_235; -x_217 = x_234; -x_218 = x_233; -x_219 = x_229; -x_220 = x_250; -goto block_228; +size_t x_257; size_t x_258; lean_object* x_259; +x_257 = 0; +x_258 = lean_usize_of_nat(x_245); +x_259 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__1(x_161, x_257, x_258, x_252); +x_225 = x_236; +x_226 = x_243; +x_227 = x_259; +goto block_235; } } } @@ -12711,427 +21269,463 @@ goto block_228; } else { -lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; -lean_dec_ref(x_212); -lean_dec(x_156); -lean_dec(x_155); -lean_dec_ref(x_154); -lean_dec(x_153); -lean_dec(x_151); -x_259 = lean_ctor_get(x_213, 0); -lean_inc(x_259); -x_260 = lean_ctor_get(x_213, 1); -lean_inc(x_260); -if (lean_is_exclusive(x_213)) { - lean_ctor_release(x_213, 0); - lean_ctor_release(x_213, 1); - x_261 = x_213; +lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; +lean_dec_ref(x_211); +lean_dec(x_163); +lean_dec(x_162); +lean_dec_ref(x_161); +lean_dec(x_160); +lean_dec(x_158); +x_268 = lean_ctor_get(x_212, 0); +lean_inc(x_268); +x_269 = lean_ctor_get(x_212, 1); +lean_inc(x_269); +if (lean_is_exclusive(x_212)) { + lean_ctor_release(x_212, 0); + lean_ctor_release(x_212, 1); + x_270 = x_212; } else { - lean_dec_ref(x_213); - x_261 = lean_box(0); + lean_dec_ref(x_212); + x_270 = lean_box(0); } -if (lean_is_scalar(x_261)) { - x_262 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_270)) { + x_271 = lean_alloc_ctor(1, 2, 0); } else { - x_262 = x_261; + x_271 = x_270; } -lean_ctor_set(x_262, 0, x_259); -lean_ctor_set(x_262, 1, x_260); -return x_262; +lean_ctor_set(x_271, 0, x_268); +lean_ctor_set(x_271, 1, x_269); +return x_271; } -block_176: +block_187: { -lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; -x_159 = lean_ctor_get(x_157, 0); -lean_inc_ref(x_159); -x_160 = lean_ctor_get(x_157, 1); -lean_inc(x_160); -x_161 = lean_ctor_get(x_157, 2); -lean_inc_ref(x_161); -if (lean_is_exclusive(x_157)) { - lean_ctor_release(x_157, 0); - lean_ctor_release(x_157, 1); - lean_ctor_release(x_157, 2); - lean_ctor_release(x_157, 3); - lean_ctor_release(x_157, 4); - x_162 = x_157; +lean_object* x_166; +x_166 = l_Lean_IR_EmitC_emitSpreads(x_161, x_164, x_165); +if (lean_obj_tag(x_166) == 0) +{ +lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; +x_167 = lean_ctor_get(x_166, 1); +lean_inc(x_167); +lean_dec_ref(x_166); +x_168 = lean_ctor_get(x_164, 0); +lean_inc_ref(x_168); +x_169 = lean_ctor_get(x_164, 1); +lean_inc(x_169); +x_170 = lean_ctor_get(x_164, 2); +lean_inc_ref(x_170); +x_171 = lean_ctor_get(x_164, 5); +lean_inc_ref(x_171); +x_172 = lean_ctor_get(x_164, 6); +lean_inc_ref(x_172); +if (lean_is_exclusive(x_164)) { + lean_ctor_release(x_164, 0); + lean_ctor_release(x_164, 1); + lean_ctor_release(x_164, 2); + lean_ctor_release(x_164, 3); + lean_ctor_release(x_164, 4); + lean_ctor_release(x_164, 5); + lean_ctor_release(x_164, 6); + x_173 = x_164; } else { - lean_dec_ref(x_157); - x_162 = lean_box(0); -} -x_163 = l_Lean_IR_EmitC_emitDeclAux___closed__0; -x_164 = lean_string_append(x_158, x_163); -x_165 = l_Lean_IR_EmitC_emitLn___closed__0; -x_166 = lean_string_append(x_164, x_165); -if (lean_is_scalar(x_162)) { - x_167 = lean_alloc_ctor(0, 5, 0); + lean_dec_ref(x_164); + x_173 = lean_box(0); +} +x_174 = l_Lean_IR_EmitC_emitDeclAux___closed__0; +x_175 = lean_string_append(x_167, x_174); +x_176 = l_Lean_IR_EmitC_emitLn___closed__0; +x_177 = lean_string_append(x_175, x_176); +if (lean_is_scalar(x_173)) { + x_178 = lean_alloc_ctor(0, 7, 0); } else { - x_167 = x_162; -} -lean_ctor_set(x_167, 0, x_159); -lean_ctor_set(x_167, 1, x_160); -lean_ctor_set(x_167, 2, x_161); -lean_ctor_set(x_167, 3, x_153); -lean_ctor_set(x_167, 4, x_154); -x_168 = l_Lean_IR_EmitC_emitFnBody(x_156, x_167, x_166); -if (lean_obj_tag(x_168) == 0) -{ -lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; -x_169 = lean_ctor_get(x_168, 1); -lean_inc(x_169); -if (lean_is_exclusive(x_168)) { - lean_ctor_release(x_168, 0); - lean_ctor_release(x_168, 1); - x_170 = x_168; + x_178 = x_173; +} +lean_ctor_set(x_178, 0, x_168); +lean_ctor_set(x_178, 1, x_169); +lean_ctor_set(x_178, 2, x_170); +lean_ctor_set(x_178, 3, x_160); +lean_ctor_set(x_178, 4, x_161); +lean_ctor_set(x_178, 5, x_171); +lean_ctor_set(x_178, 6, x_172); +x_179 = l_Lean_IR_EmitC_emitFnBody(x_163, x_178, x_177); +if (lean_obj_tag(x_179) == 0) +{ +lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; +x_180 = lean_ctor_get(x_179, 1); +lean_inc(x_180); +if (lean_is_exclusive(x_179)) { + lean_ctor_release(x_179, 0); + lean_ctor_release(x_179, 1); + x_181 = x_179; } else { - lean_dec_ref(x_168); - x_170 = lean_box(0); -} -x_171 = l_Lean_IR_EmitC_emitMainFn___closed__12; -x_172 = lean_string_append(x_169, x_171); -x_173 = lean_box(0); -x_174 = lean_string_append(x_172, x_165); -if (lean_is_scalar(x_170)) { - x_175 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_179); + x_181 = lean_box(0); +} +x_182 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; +x_183 = lean_string_append(x_180, x_182); +x_184 = lean_box(0); +x_185 = lean_string_append(x_183, x_176); +if (lean_is_scalar(x_181)) { + x_186 = lean_alloc_ctor(0, 2, 0); } else { - x_175 = x_170; + x_186 = x_181; } -lean_ctor_set(x_175, 0, x_173); -lean_ctor_set(x_175, 1, x_174); -return x_175; +lean_ctor_set(x_186, 0, x_184); +lean_ctor_set(x_186, 1, x_185); +return x_186; } else { -return x_168; -} +return x_179; } -block_183: -{ -if (x_180 == 0) -{ -lean_dec(x_179); -x_157 = x_178; -x_158 = x_177; -goto block_176; } else { -lean_object* x_181; lean_object* x_182; -lean_inc(x_179); -x_181 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__0___redArg(x_154, x_179, x_179, x_177); -lean_dec(x_179); -x_182 = lean_ctor_get(x_181, 1); -lean_inc(x_182); -lean_dec_ref(x_181); -x_157 = x_178; -x_158 = x_182; -goto block_176; +lean_dec_ref(x_164); +lean_dec(x_163); +lean_dec_ref(x_161); +lean_dec(x_160); +return x_166; } } block_194: { -lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; uint8_t x_192; -x_186 = l_Lean_IR_EmitC_emitDeclAux___closed__1; -x_187 = lean_string_append(x_185, x_186); -x_188 = l_Lean_IR_EmitC_emitLn___closed__0; -x_189 = lean_string_append(x_187, x_188); -x_190 = l_Lean_IR_EmitC_emitFnDeclAux___closed__3; -x_191 = lean_array_get_size(x_154); -x_192 = lean_nat_dec_lt(x_190, x_191); -if (x_192 == 0) +if (x_191 == 0) { -lean_dec(x_151); -x_177 = x_189; -x_178 = x_184; -x_179 = x_191; -x_180 = x_192; -goto block_183; +lean_dec(x_190); +x_164 = x_188; +x_165 = x_189; +goto block_187; } else { -uint8_t x_193; -x_193 = l_Lean_IR_ExplicitBoxing_isBoxedName(x_151); -lean_dec(x_151); -x_177 = x_189; -x_178 = x_184; -x_179 = x_191; -x_180 = x_193; -goto block_183; +lean_object* x_192; lean_object* x_193; +lean_inc(x_190); +x_192 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__0___redArg(x_161, x_190, x_190, x_189); +lean_dec(x_190); +x_193 = lean_ctor_get(x_192, 1); +lean_inc(x_193); +lean_dec_ref(x_192); +x_164 = x_188; +x_165 = x_193; +goto block_187; } } -block_199: +block_205: { -lean_object* x_197; lean_object* x_198; -x_197 = l_Lean_IR_EmitC_emitFnDeclAux___closed__1; +lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; uint8_t x_203; +x_197 = l_Lean_IR_EmitC_emitStructDefn___closed__0; x_198 = lean_string_append(x_196, x_197); -x_184 = x_195; -x_185 = x_198; +x_199 = l_Lean_IR_EmitC_emitLn___closed__0; +x_200 = lean_string_append(x_198, x_199); +x_201 = l_Lean_IR_EmitC_emitFnDeclAux___closed__1; +x_202 = lean_array_get_size(x_161); +x_203 = lean_nat_dec_lt(x_201, x_202); +if (x_203 == 0) +{ +lean_dec(x_158); +x_188 = x_195; +x_189 = x_200; +x_190 = x_202; +x_191 = x_203; goto block_194; } -block_211: -{ -if (x_206 == 0) -{ -lean_object* x_207; lean_object* x_208; -lean_inc(x_202); -x_207 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__1___redArg(x_205, x_203, x_200, x_202, x_202, x_201); -lean_dec(x_202); -lean_dec_ref(x_203); -lean_dec_ref(x_205); -x_208 = lean_ctor_get(x_207, 1); -lean_inc(x_208); -lean_dec_ref(x_207); -x_195 = x_204; -x_196 = x_208; -goto block_199; -} else { -lean_object* x_209; lean_object* x_210; -lean_dec_ref(x_205); -lean_dec_ref(x_203); -lean_dec(x_202); -x_209 = l_Lean_IR_EmitC_emitDeclAux___closed__2; -x_210 = lean_string_append(x_201, x_209); -x_195 = x_204; -x_196 = x_210; -goto block_199; +uint8_t x_204; +x_204 = l_Lean_IR_ExplicitBoxing_isBoxedName(x_158); +lean_dec(x_158); +x_188 = x_195; +x_189 = x_200; +x_190 = x_202; +x_191 = x_204; +goto block_194; } } +block_210: +{ +lean_object* x_208; lean_object* x_209; +x_208 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__1; +x_209 = lean_string_append(x_207, x_208); +x_195 = x_206; +x_196 = x_209; +goto block_205; +} } else { -lean_object* x_263; +lean_object* x_272; +lean_dec(x_158); +lean_dec_ref(x_157); +lean_dec_ref(x_156); +lean_dec(x_155); +lean_dec(x_154); +lean_dec_ref(x_153); +lean_dec(x_152); lean_dec(x_151); -lean_dec_ref(x_150); -lean_dec(x_149); -lean_dec(x_148); -lean_dec_ref(x_147); -lean_dec(x_6); lean_dec_ref(x_1); -x_263 = lean_box(0); +x_272 = lean_box(0); lean_ctor_set(x_4, 1, x_3); -lean_ctor_set(x_4, 0, x_263); +lean_ctor_set(x_4, 0, x_272); return x_4; } } else { -lean_object* x_264; +lean_object* x_273; +lean_dec(x_158); +lean_dec_ref(x_157); +lean_dec_ref(x_156); +lean_dec(x_155); +lean_dec(x_154); +lean_dec_ref(x_153); +lean_dec(x_152); lean_dec(x_151); -lean_dec_ref(x_150); -lean_dec(x_149); -lean_dec(x_148); -lean_dec_ref(x_147); -lean_dec(x_6); lean_dec_ref(x_1); -x_264 = lean_box(0); +x_273 = lean_box(0); lean_ctor_set(x_4, 1, x_3); -lean_ctor_set(x_4, 0, x_264); +lean_ctor_set(x_4, 0, x_273); return x_4; } } } else { -lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; uint8_t x_272; -x_265 = lean_ctor_get(x_4, 1); -lean_inc(x_265); +lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; uint8_t x_283; +x_274 = lean_ctor_get(x_4, 0); +x_275 = lean_ctor_get(x_4, 1); +lean_inc(x_275); +lean_inc(x_274); lean_dec(x_4); -x_266 = lean_ctor_get(x_2, 0); -lean_inc_ref(x_266); -x_267 = lean_ctor_get(x_2, 1); -lean_inc(x_267); -x_268 = lean_ctor_get(x_2, 3); -lean_inc(x_268); -x_269 = lean_ctor_get(x_2, 4); -lean_inc_ref(x_269); +x_276 = lean_ctor_get(x_2, 0); +lean_inc_ref(x_276); +x_277 = lean_ctor_get(x_2, 1); +lean_inc(x_277); +x_278 = lean_ctor_get(x_2, 3); +lean_inc(x_278); +x_279 = lean_ctor_get(x_2, 4); +lean_inc_ref(x_279); +x_280 = lean_ctor_get(x_2, 6); +lean_inc_ref(x_280); if (lean_is_exclusive(x_2)) { lean_ctor_release(x_2, 0); lean_ctor_release(x_2, 1); lean_ctor_release(x_2, 2); lean_ctor_release(x_2, 3); lean_ctor_release(x_2, 4); - x_270 = x_2; + lean_ctor_release(x_2, 5); + lean_ctor_release(x_2, 6); + x_281 = x_2; } else { lean_dec_ref(x_2); - x_270 = lean_box(0); + x_281 = lean_box(0); } -x_271 = l_Lean_IR_Decl_name(x_1); -lean_inc(x_271); -lean_inc_ref(x_266); -x_272 = l_Lean_hasInitAttr(x_266, x_271); -if (x_272 == 0) +x_282 = l_Lean_IR_Decl_name(x_1); +lean_inc(x_282); +lean_inc_ref(x_276); +x_283 = l_Lean_hasInitAttr(x_276, x_282); +if (x_283 == 0) { if (lean_obj_tag(x_1) == 0) { -lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_297; lean_object* x_298; lean_object* x_299; uint8_t x_300; lean_object* x_304; lean_object* x_305; lean_object* x_315; lean_object* x_316; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; uint8_t x_326; lean_object* x_332; lean_object* x_333; -x_273 = lean_ctor_get(x_1, 0); -lean_inc(x_273); -x_274 = lean_ctor_get(x_1, 1); -lean_inc_ref(x_274); -x_275 = lean_ctor_get(x_1, 2); -lean_inc(x_275); -x_276 = lean_ctor_get(x_1, 3); -lean_inc(x_276); +lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_312; lean_object* x_313; lean_object* x_314; uint8_t x_315; lean_object* x_319; lean_object* x_320; lean_object* x_330; lean_object* x_331; lean_object* x_335; lean_object* x_336; +x_284 = lean_ctor_get(x_1, 0); +lean_inc(x_284); +x_285 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_285); +x_286 = lean_ctor_get(x_1, 2); +lean_inc(x_286); +x_287 = lean_ctor_get(x_1, 3); +lean_inc(x_287); lean_dec_ref(x_1); -if (lean_is_scalar(x_270)) { - x_332 = lean_alloc_ctor(0, 5, 0); +if (lean_is_scalar(x_281)) { + x_335 = lean_alloc_ctor(0, 7, 0); } else { - x_332 = x_270; -} -lean_ctor_set(x_332, 0, x_266); -lean_ctor_set(x_332, 1, x_267); -lean_ctor_set(x_332, 2, x_265); -lean_ctor_set(x_332, 3, x_268); -lean_ctor_set(x_332, 4, x_269); -lean_inc_ref(x_332); -lean_inc(x_273); -x_333 = l_Lean_IR_EmitC_toCName(x_273, x_332, x_3); -if (lean_obj_tag(x_333) == 0) -{ -lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_349; lean_object* x_350; lean_object* x_372; lean_object* x_373; uint8_t x_374; -x_334 = lean_ctor_get(x_333, 0); -lean_inc(x_334); -x_335 = lean_ctor_get(x_333, 1); -lean_inc(x_335); -lean_dec_ref(x_333); -x_372 = lean_array_get_size(x_274); -x_373 = lean_unsigned_to_nat(0u); -x_374 = lean_nat_dec_eq(x_372, x_373); -if (x_374 == 0) -{ -lean_object* x_375; lean_object* x_376; -x_375 = l_Lean_IR_EmitC_emitFnDeclAux___closed__7; -x_376 = lean_string_append(x_335, x_375); -x_349 = x_332; -x_350 = x_376; -goto block_371; -} -else -{ -lean_object* x_377; lean_object* x_378; -x_377 = l_Lean_IR_EmitC_emitFnDeclAux___closed__8; -x_378 = lean_string_append(x_335, x_377); -x_349 = x_332; -x_350 = x_378; -goto block_371; + x_335 = x_281; +} +lean_ctor_set(x_335, 0, x_276); +lean_ctor_set(x_335, 1, x_277); +lean_ctor_set(x_335, 2, x_275); +lean_ctor_set(x_335, 3, x_278); +lean_ctor_set(x_335, 4, x_279); +lean_ctor_set(x_335, 5, x_274); +lean_ctor_set(x_335, 6, x_280); +lean_inc_ref(x_335); +lean_inc(x_284); +x_336 = l_Lean_IR_EmitC_toCName(x_284, x_335, x_3); +if (lean_obj_tag(x_336) == 0) +{ +lean_object* x_337; lean_object* x_338; uint8_t x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; uint8_t x_343; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_360; lean_object* x_361; lean_object* x_385; lean_object* x_386; uint8_t x_387; +x_337 = lean_ctor_get(x_336, 0); +lean_inc(x_337); +x_338 = lean_ctor_get(x_336, 1); +lean_inc(x_338); +lean_dec_ref(x_336); +x_339 = 1; +x_385 = lean_array_get_size(x_285); +x_386 = lean_unsigned_to_nat(0u); +x_387 = lean_nat_dec_eq(x_385, x_386); +if (x_387 == 0) +{ +lean_object* x_388; lean_object* x_389; +x_388 = l_Lean_IR_EmitC_emitFnDeclAux___closed__4; +x_389 = lean_string_append(x_338, x_388); +x_360 = x_335; +x_361 = x_389; +goto block_384; +} +else +{ +lean_object* x_390; lean_object* x_391; +x_390 = l_Lean_IR_EmitC_emitFnDeclAux___closed__5; +x_391 = lean_string_append(x_338, x_390); +x_360 = x_335; +x_361 = x_391; +goto block_384; } block_348: { -lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; uint8_t x_346; -x_341 = lean_string_append(x_337, x_334); -lean_dec(x_334); -x_342 = l_Lean_IR_EmitC_emitFnDeclAux___closed__5; -x_343 = lean_string_append(x_341, x_342); -x_344 = l_Lean_IR_EmitC_emitFnDeclAux___closed__3; -x_345 = lean_array_get_size(x_340); -x_346 = lean_nat_dec_lt(x_344, x_345); -if (x_346 == 0) -{ -x_320 = x_336; -x_321 = x_343; -x_322 = x_345; -x_323 = x_338; -x_324 = x_339; -x_325 = x_340; -x_326 = x_346; -goto block_331; -} -else -{ -uint8_t x_347; -x_347 = l_Lean_IR_ExplicitBoxing_isBoxedName(x_271); -x_320 = x_336; -x_321 = x_343; -x_322 = x_345; -x_323 = x_338; -x_324 = x_339; -x_325 = x_340; -x_326 = x_347; -goto block_331; -} -} -block_371: -{ -lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; uint8_t x_357; -x_351 = l_Lean_IR_EmitC_toCType(x_275); -lean_dec(x_275); -x_352 = lean_string_append(x_350, x_351); -lean_dec_ref(x_351); -x_353 = l_Lean_IR_EmitC_emitFnDeclAux___closed__4; +if (x_343 == 0) +{ +lean_object* x_344; +x_344 = l_Lean_IR_EmitC_emitSpreadArgs(x_341, x_339, x_340, x_342); +lean_dec_ref(x_341); +if (lean_obj_tag(x_344) == 0) +{ +lean_object* x_345; +x_345 = lean_ctor_get(x_344, 1); +lean_inc(x_345); +lean_dec_ref(x_344); +x_330 = x_340; +x_331 = x_345; +goto block_334; +} +else +{ +lean_dec_ref(x_340); +lean_dec(x_287); +lean_dec_ref(x_285); +lean_dec(x_284); +lean_dec(x_282); +return x_344; +} +} +else +{ +lean_object* x_346; lean_object* x_347; +lean_dec_ref(x_341); +x_346 = l_Lean_IR_EmitC_emitDeclAux___closed__1; +x_347 = lean_string_append(x_342, x_346); +x_330 = x_340; +x_331 = x_347; +goto block_334; +} +} +block_359: +{ +lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; uint8_t x_357; +x_352 = lean_string_append(x_350, x_337); +lean_dec(x_337); +x_353 = l_Lean_IR_EmitC_emitFnDeclAux___closed__2; x_354 = lean_string_append(x_352, x_353); -x_355 = lean_unsigned_to_nat(0u); -x_356 = lean_array_get_size(x_274); +x_355 = l_Lean_IR_EmitC_emitFnDeclAux___closed__1; +x_356 = lean_array_get_size(x_351); x_357 = lean_nat_dec_lt(x_355, x_356); if (x_357 == 0) { -lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; -x_358 = l_Lean_IR_EmitC_toCInitName___closed__0; -x_359 = lean_string_append(x_358, x_334); -lean_dec(x_334); -x_360 = l_Lean_IR_EmitC_emitDeclAux___closed__3; -x_361 = lean_string_append(x_359, x_360); -x_362 = lean_string_append(x_354, x_361); -lean_dec_ref(x_361); -x_304 = x_349; -x_305 = x_362; -goto block_314; +x_340 = x_349; +x_341 = x_351; +x_342 = x_354; +x_343 = x_357; +goto block_348; } else { -lean_object* x_363; -x_363 = l_Lean_IR_EmitC_emitFnDeclAux___closed__6; -if (x_357 == 0) -{ -x_336 = x_355; -x_337 = x_354; -x_338 = x_353; -x_339 = x_349; -x_340 = x_363; +uint8_t x_358; +x_358 = l_Lean_IR_ExplicitBoxing_isBoxedName(x_282); +x_340 = x_349; +x_341 = x_351; +x_342 = x_354; +x_343 = x_358; goto block_348; } +} +block_384: +{ +lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; uint8_t x_370; +x_362 = l_Lean_IR_EmitC_toCType(x_286, x_360, x_361); +lean_dec(x_286); +x_363 = lean_ctor_get(x_362, 0); +lean_inc(x_363); +x_364 = lean_ctor_get(x_362, 1); +lean_inc(x_364); +lean_dec_ref(x_362); +x_365 = lean_string_append(x_364, x_363); +lean_dec(x_363); +x_366 = l_Lean_IR_EmitC_emitSpreadArg___closed__0; +x_367 = lean_string_append(x_365, x_366); +x_368 = lean_unsigned_to_nat(0u); +x_369 = lean_array_get_size(x_285); +x_370 = lean_nat_dec_lt(x_368, x_369); +if (x_370 == 0) +{ +lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; +x_371 = l_Lean_IR_EmitC_toCInitName___closed__0; +x_372 = lean_string_append(x_371, x_337); +lean_dec(x_337); +x_373 = l_Lean_IR_EmitC_emitDeclAux___closed__2; +x_374 = lean_string_append(x_372, x_373); +x_375 = lean_string_append(x_367, x_374); +lean_dec_ref(x_374); +x_319 = x_360; +x_320 = x_375; +goto block_329; +} else { -uint8_t x_364; -x_364 = lean_nat_dec_le(x_356, x_356); -if (x_364 == 0) +lean_object* x_376; +x_376 = l_Lean_IR_EmitC_emitFnDeclAux___closed__3; +if (x_370 == 0) { -if (x_357 == 0) +x_349 = x_360; +x_350 = x_367; +x_351 = x_376; +goto block_359; +} +else { -x_336 = x_355; -x_337 = x_354; -x_338 = x_353; -x_339 = x_349; -x_340 = x_363; -goto block_348; +uint8_t x_377; +x_377 = lean_nat_dec_le(x_369, x_369); +if (x_377 == 0) +{ +if (x_370 == 0) +{ +x_349 = x_360; +x_350 = x_367; +x_351 = x_376; +goto block_359; } else { -size_t x_365; size_t x_366; lean_object* x_367; -x_365 = 0; -x_366 = lean_usize_of_nat(x_356); -x_367 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__2(x_274, x_365, x_366, x_363); -x_336 = x_355; -x_337 = x_354; -x_338 = x_353; -x_339 = x_349; -x_340 = x_367; -goto block_348; +size_t x_378; size_t x_379; lean_object* x_380; +x_378 = 0; +x_379 = lean_usize_of_nat(x_369); +x_380 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__1(x_285, x_378, x_379, x_376); +x_349 = x_360; +x_350 = x_367; +x_351 = x_380; +goto block_359; } } else { -size_t x_368; size_t x_369; lean_object* x_370; -x_368 = 0; -x_369 = lean_usize_of_nat(x_356); -x_370 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__2(x_274, x_368, x_369, x_363); -x_336 = x_355; -x_337 = x_354; -x_338 = x_353; -x_339 = x_349; -x_340 = x_370; -goto block_348; +size_t x_381; size_t x_382; lean_object* x_383; +x_381 = 0; +x_382 = lean_usize_of_nat(x_369); +x_383 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__1(x_285, x_381, x_382, x_376); +x_349 = x_360; +x_350 = x_367; +x_351 = x_383; +goto block_359; } } } @@ -13139,228 +21733,226 @@ goto block_348; } else { -lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; -lean_dec_ref(x_332); -lean_dec(x_276); -lean_dec(x_275); -lean_dec_ref(x_274); -lean_dec(x_273); -lean_dec(x_271); -x_379 = lean_ctor_get(x_333, 0); -lean_inc(x_379); -x_380 = lean_ctor_get(x_333, 1); -lean_inc(x_380); -if (lean_is_exclusive(x_333)) { - lean_ctor_release(x_333, 0); - lean_ctor_release(x_333, 1); - x_381 = x_333; -} else { - lean_dec_ref(x_333); - x_381 = lean_box(0); -} -if (lean_is_scalar(x_381)) { - x_382 = lean_alloc_ctor(1, 2, 0); +lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; +lean_dec_ref(x_335); +lean_dec(x_287); +lean_dec(x_286); +lean_dec_ref(x_285); +lean_dec(x_284); +lean_dec(x_282); +x_392 = lean_ctor_get(x_336, 0); +lean_inc(x_392); +x_393 = lean_ctor_get(x_336, 1); +lean_inc(x_393); +if (lean_is_exclusive(x_336)) { + lean_ctor_release(x_336, 0); + lean_ctor_release(x_336, 1); + x_394 = x_336; } else { - x_382 = x_381; -} -lean_ctor_set(x_382, 0, x_379); -lean_ctor_set(x_382, 1, x_380); -return x_382; + lean_dec_ref(x_336); + x_394 = lean_box(0); } -block_296: -{ -lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; -x_279 = lean_ctor_get(x_277, 0); -lean_inc_ref(x_279); -x_280 = lean_ctor_get(x_277, 1); -lean_inc(x_280); -x_281 = lean_ctor_get(x_277, 2); -lean_inc_ref(x_281); -if (lean_is_exclusive(x_277)) { - lean_ctor_release(x_277, 0); - lean_ctor_release(x_277, 1); - lean_ctor_release(x_277, 2); - lean_ctor_release(x_277, 3); - lean_ctor_release(x_277, 4); - x_282 = x_277; +if (lean_is_scalar(x_394)) { + x_395 = lean_alloc_ctor(1, 2, 0); } else { - lean_dec_ref(x_277); - x_282 = lean_box(0); -} -x_283 = l_Lean_IR_EmitC_emitDeclAux___closed__0; -x_284 = lean_string_append(x_278, x_283); -x_285 = l_Lean_IR_EmitC_emitLn___closed__0; -x_286 = lean_string_append(x_284, x_285); -if (lean_is_scalar(x_282)) { - x_287 = lean_alloc_ctor(0, 5, 0); -} else { - x_287 = x_282; -} -lean_ctor_set(x_287, 0, x_279); -lean_ctor_set(x_287, 1, x_280); -lean_ctor_set(x_287, 2, x_281); -lean_ctor_set(x_287, 3, x_273); -lean_ctor_set(x_287, 4, x_274); -x_288 = l_Lean_IR_EmitC_emitFnBody(x_276, x_287, x_286); -if (lean_obj_tag(x_288) == 0) -{ -lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; -x_289 = lean_ctor_get(x_288, 1); -lean_inc(x_289); + x_395 = x_394; +} +lean_ctor_set(x_395, 0, x_392); +lean_ctor_set(x_395, 1, x_393); +return x_395; +} +block_311: +{ +lean_object* x_290; +x_290 = l_Lean_IR_EmitC_emitSpreads(x_285, x_288, x_289); +if (lean_obj_tag(x_290) == 0) +{ +lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; +x_291 = lean_ctor_get(x_290, 1); +lean_inc(x_291); +lean_dec_ref(x_290); +x_292 = lean_ctor_get(x_288, 0); +lean_inc_ref(x_292); +x_293 = lean_ctor_get(x_288, 1); +lean_inc(x_293); +x_294 = lean_ctor_get(x_288, 2); +lean_inc_ref(x_294); +x_295 = lean_ctor_get(x_288, 5); +lean_inc_ref(x_295); +x_296 = lean_ctor_get(x_288, 6); +lean_inc_ref(x_296); if (lean_is_exclusive(x_288)) { lean_ctor_release(x_288, 0); lean_ctor_release(x_288, 1); - x_290 = x_288; + lean_ctor_release(x_288, 2); + lean_ctor_release(x_288, 3); + lean_ctor_release(x_288, 4); + lean_ctor_release(x_288, 5); + lean_ctor_release(x_288, 6); + x_297 = x_288; } else { lean_dec_ref(x_288); - x_290 = lean_box(0); -} -x_291 = l_Lean_IR_EmitC_emitMainFn___closed__12; -x_292 = lean_string_append(x_289, x_291); -x_293 = lean_box(0); -x_294 = lean_string_append(x_292, x_285); -if (lean_is_scalar(x_290)) { - x_295 = lean_alloc_ctor(0, 2, 0); + x_297 = lean_box(0); +} +x_298 = l_Lean_IR_EmitC_emitDeclAux___closed__0; +x_299 = lean_string_append(x_291, x_298); +x_300 = l_Lean_IR_EmitC_emitLn___closed__0; +x_301 = lean_string_append(x_299, x_300); +if (lean_is_scalar(x_297)) { + x_302 = lean_alloc_ctor(0, 7, 0); +} else { + x_302 = x_297; +} +lean_ctor_set(x_302, 0, x_292); +lean_ctor_set(x_302, 1, x_293); +lean_ctor_set(x_302, 2, x_294); +lean_ctor_set(x_302, 3, x_284); +lean_ctor_set(x_302, 4, x_285); +lean_ctor_set(x_302, 5, x_295); +lean_ctor_set(x_302, 6, x_296); +x_303 = l_Lean_IR_EmitC_emitFnBody(x_287, x_302, x_301); +if (lean_obj_tag(x_303) == 0) +{ +lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; +x_304 = lean_ctor_get(x_303, 1); +lean_inc(x_304); +if (lean_is_exclusive(x_303)) { + lean_ctor_release(x_303, 0); + lean_ctor_release(x_303, 1); + x_305 = x_303; } else { - x_295 = x_290; + lean_dec_ref(x_303); + x_305 = lean_box(0); +} +x_306 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; +x_307 = lean_string_append(x_304, x_306); +x_308 = lean_box(0); +x_309 = lean_string_append(x_307, x_300); +if (lean_is_scalar(x_305)) { + x_310 = lean_alloc_ctor(0, 2, 0); +} else { + x_310 = x_305; } -lean_ctor_set(x_295, 0, x_293); -lean_ctor_set(x_295, 1, x_294); -return x_295; +lean_ctor_set(x_310, 0, x_308); +lean_ctor_set(x_310, 1, x_309); +return x_310; } else { -return x_288; -} +return x_303; } -block_303: -{ -if (x_300 == 0) -{ -lean_dec(x_299); -x_277 = x_298; -x_278 = x_297; -goto block_296; } else { -lean_object* x_301; lean_object* x_302; -lean_inc(x_299); -x_301 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__0___redArg(x_274, x_299, x_299, x_297); -lean_dec(x_299); -x_302 = lean_ctor_get(x_301, 1); -lean_inc(x_302); -lean_dec_ref(x_301); -x_277 = x_298; -x_278 = x_302; -goto block_296; +lean_dec_ref(x_288); +lean_dec(x_287); +lean_dec_ref(x_285); +lean_dec(x_284); +return x_290; } } -block_314: +block_318: { -lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; uint8_t x_312; -x_306 = l_Lean_IR_EmitC_emitDeclAux___closed__1; -x_307 = lean_string_append(x_305, x_306); -x_308 = l_Lean_IR_EmitC_emitLn___closed__0; -x_309 = lean_string_append(x_307, x_308); -x_310 = l_Lean_IR_EmitC_emitFnDeclAux___closed__3; -x_311 = lean_array_get_size(x_274); -x_312 = lean_nat_dec_lt(x_310, x_311); -if (x_312 == 0) +if (x_315 == 0) { -lean_dec(x_271); -x_297 = x_309; -x_298 = x_304; -x_299 = x_311; -x_300 = x_312; -goto block_303; +lean_dec(x_314); +x_288 = x_312; +x_289 = x_313; +goto block_311; } else { -uint8_t x_313; -x_313 = l_Lean_IR_ExplicitBoxing_isBoxedName(x_271); -lean_dec(x_271); -x_297 = x_309; -x_298 = x_304; -x_299 = x_311; -x_300 = x_313; -goto block_303; +lean_object* x_316; lean_object* x_317; +lean_inc(x_314); +x_316 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__0___redArg(x_285, x_314, x_314, x_313); +lean_dec(x_314); +x_317 = lean_ctor_get(x_316, 1); +lean_inc(x_317); +lean_dec_ref(x_316); +x_288 = x_312; +x_289 = x_317; +goto block_311; } } -block_319: -{ -lean_object* x_317; lean_object* x_318; -x_317 = l_Lean_IR_EmitC_emitFnDeclAux___closed__1; -x_318 = lean_string_append(x_316, x_317); -x_304 = x_315; -x_305 = x_318; -goto block_314; -} -block_331: +block_329: { -if (x_326 == 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; uint8_t x_327; +x_321 = l_Lean_IR_EmitC_emitStructDefn___closed__0; +x_322 = lean_string_append(x_320, x_321); +x_323 = l_Lean_IR_EmitC_emitLn___closed__0; +x_324 = lean_string_append(x_322, x_323); +x_325 = l_Lean_IR_EmitC_emitFnDeclAux___closed__1; +x_326 = lean_array_get_size(x_285); +x_327 = lean_nat_dec_lt(x_325, x_326); +if (x_327 == 0) { -lean_object* x_327; lean_object* x_328; -lean_inc(x_322); -x_327 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__1___redArg(x_325, x_323, x_320, x_322, x_322, x_321); -lean_dec(x_322); -lean_dec_ref(x_323); -lean_dec_ref(x_325); -x_328 = lean_ctor_get(x_327, 1); -lean_inc(x_328); -lean_dec_ref(x_327); -x_315 = x_324; -x_316 = x_328; -goto block_319; +lean_dec(x_282); +x_312 = x_319; +x_313 = x_324; +x_314 = x_326; +x_315 = x_327; +goto block_318; } else { -lean_object* x_329; lean_object* x_330; -lean_dec_ref(x_325); -lean_dec_ref(x_323); -lean_dec(x_322); -x_329 = l_Lean_IR_EmitC_emitDeclAux___closed__2; -x_330 = lean_string_append(x_321, x_329); -x_315 = x_324; -x_316 = x_330; -goto block_319; +uint8_t x_328; +x_328 = l_Lean_IR_ExplicitBoxing_isBoxedName(x_282); +lean_dec(x_282); +x_312 = x_319; +x_313 = x_324; +x_314 = x_326; +x_315 = x_328; +goto block_318; } } +block_334: +{ +lean_object* x_332; lean_object* x_333; +x_332 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__1; +x_333 = lean_string_append(x_331, x_332); +x_319 = x_330; +x_320 = x_333; +goto block_329; +} } else { -lean_object* x_383; lean_object* x_384; -lean_dec(x_271); -lean_dec(x_270); -lean_dec_ref(x_269); -lean_dec(x_268); -lean_dec(x_267); -lean_dec_ref(x_266); -lean_dec(x_265); +lean_object* x_396; lean_object* x_397; +lean_dec(x_282); +lean_dec(x_281); +lean_dec_ref(x_280); +lean_dec_ref(x_279); +lean_dec(x_278); +lean_dec(x_277); +lean_dec_ref(x_276); +lean_dec(x_275); +lean_dec(x_274); lean_dec_ref(x_1); -x_383 = lean_box(0); -x_384 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_384, 0, x_383); -lean_ctor_set(x_384, 1, x_3); -return x_384; +x_396 = lean_box(0); +x_397 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_397, 0, x_396); +lean_ctor_set(x_397, 1, x_3); +return x_397; } } else { -lean_object* x_385; lean_object* x_386; -lean_dec(x_271); -lean_dec(x_270); -lean_dec_ref(x_269); -lean_dec(x_268); -lean_dec(x_267); -lean_dec_ref(x_266); -lean_dec(x_265); +lean_object* x_398; lean_object* x_399; +lean_dec(x_282); +lean_dec(x_281); +lean_dec_ref(x_280); +lean_dec_ref(x_279); +lean_dec(x_278); +lean_dec(x_277); +lean_dec_ref(x_276); +lean_dec(x_275); +lean_dec(x_274); lean_dec_ref(x_1); -x_385 = lean_box(0); -x_386 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_386, 0, x_385); -lean_ctor_set(x_386, 1, x_3); -return x_386; +x_398 = lean_box(0); +x_399 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_399, 0, x_398); +lean_ctor_set(x_399, 1, x_3); +return x_399; } } } @@ -13376,19 +21968,6 @@ lean_dec_ref(x_1); return x_7; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec_ref(x_7); -lean_dec(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -return x_9; -} -} LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { @@ -13399,18 +21978,6 @@ lean_dec_ref(x_1); return x_5; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__1___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; -x_7 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__1___redArg(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -return x_7; -} -} static lean_object* _init_l_Lean_IR_EmitC_emitDecl___closed__0() { _start: { @@ -13422,48 +21989,49 @@ return x_1; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitDecl(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_4; lean_object* x_5; +lean_object* x_4; lean_object* x_5; lean_object* x_6; x_4 = l_Lean_IR_Decl_normalizeIds(x_1); -lean_inc_ref(x_4); -x_5 = l_Lean_IR_EmitC_emitDeclAux(x_4, x_2, x_3); -if (lean_obj_tag(x_5) == 0) +x_5 = l_Lean_IR_StructRC_visitDecl(x_4); +lean_inc_ref(x_5); +x_6 = l_Lean_IR_EmitC_emitDeclAux(x_5, x_2, x_3); +if (lean_obj_tag(x_6) == 0) { -lean_dec_ref(x_4); -return x_5; +lean_dec_ref(x_5); +return x_6; } else { -uint8_t x_6; -x_6 = !lean_is_exclusive(x_5); -if (x_6 == 0) +uint8_t x_7; +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_7 = lean_ctor_get(x_5, 0); -x_8 = l_Lean_IR_EmitC_emitDecl___closed__0; -x_9 = lean_string_append(x_7, x_8); -x_10 = l_Lean_IR_declToString(x_4); -x_11 = lean_string_append(x_9, x_10); -lean_dec_ref(x_10); -lean_ctor_set(x_5, 0, x_11); -return x_5; +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_8 = lean_ctor_get(x_6, 0); +x_9 = l_Lean_IR_EmitC_emitDecl___closed__0; +x_10 = lean_string_append(x_8, x_9); +x_11 = l_Lean_IR_declToString(x_5); +x_12 = lean_string_append(x_10, x_11); +lean_dec_ref(x_11); +lean_ctor_set(x_6, 0, x_12); +return x_6; } else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_12 = lean_ctor_get(x_5, 0); -x_13 = lean_ctor_get(x_5, 1); +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_13 = lean_ctor_get(x_6, 0); +x_14 = lean_ctor_get(x_6, 1); +lean_inc(x_14); lean_inc(x_13); -lean_inc(x_12); -lean_dec(x_5); -x_14 = l_Lean_IR_EmitC_emitDecl___closed__0; -x_15 = lean_string_append(x_12, x_14); -x_16 = l_Lean_IR_declToString(x_4); -x_17 = lean_string_append(x_15, x_16); -lean_dec_ref(x_16); -x_18 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_13); -return x_18; +lean_dec(x_6); +x_15 = l_Lean_IR_EmitC_emitDecl___closed__0; +x_16 = lean_string_append(x_13, x_15); +x_17 = l_Lean_IR_declToString(x_5); +x_18 = lean_string_append(x_16, x_17); +lean_dec_ref(x_17); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_14); +return x_19; } } } @@ -13564,7 +22132,7 @@ lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean x_13 = lean_ctor_get(x_11, 1); x_14 = lean_ctor_get(x_11, 0); lean_dec(x_14); -x_15 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; +x_15 = l_Lean_IR_EmitC_emitIncOfType___closed__0; x_16 = lean_string_append(x_13, x_15); x_17 = l_Lean_IR_EmitC_emitLn___closed__0; x_18 = lean_box(0); @@ -13579,7 +22147,7 @@ lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean x_20 = lean_ctor_get(x_11, 1); lean_inc(x_20); lean_dec(x_11); -x_21 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; +x_21 = l_Lean_IR_EmitC_emitIncOfType___closed__0; x_22 = lean_string_append(x_20, x_21); x_23 = l_Lean_IR_EmitC_emitLn___closed__0; x_24 = lean_box(0); @@ -13610,7 +22178,7 @@ static lean_object* _init_l_Lean_IR_EmitC_emitDeclInit___closed__0() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("();", 3, 3); +x_1 = lean_mk_string_unchecked("(lean_box(0));", 14, 14); return x_1; } } @@ -13646,260 +22214,269 @@ x_1 = lean_mk_string_unchecked("if (builtin) {", 14, 14); return x_1; } } +static lean_object* _init_l_Lean_IR_EmitC_emitDeclInit___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("();", 3, 3); +return x_1; +} +} LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitDeclInit(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_4; uint8_t x_5; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_57; +lean_object* x_4; uint8_t x_5; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_58; x_13 = lean_ctor_get(x_2, 0); x_14 = l_Lean_IR_Decl_name(x_1); lean_inc(x_14); lean_inc_ref(x_13); -x_57 = l_Lean_isIOUnitInitFn(x_13, x_14); -if (x_57 == 0) +x_58 = l_Lean_isIOUnitInitFn(x_13, x_14); +if (x_58 == 0) { -lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; lean_object* x_62; -x_58 = l_Lean_IR_Decl_params(x_1); -x_59 = lean_array_get_size(x_58); -lean_dec_ref(x_58); -x_60 = lean_unsigned_to_nat(0u); -x_61 = lean_nat_dec_eq(x_59, x_60); -if (x_61 == 0) +lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; lean_object* x_63; +x_59 = l_Lean_IR_Decl_params(x_1); +x_60 = lean_array_get_size(x_59); +lean_dec_ref(x_59); +x_61 = lean_unsigned_to_nat(0u); +x_62 = lean_nat_dec_eq(x_60, x_61); +if (x_62 == 0) { -lean_object* x_69; lean_object* x_70; +lean_object* x_70; lean_object* x_71; lean_dec(x_14); lean_dec_ref(x_2); -x_69 = lean_box(0); -x_70 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_3); -return x_70; +x_70 = lean_box(0); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_3); +return x_71; } else { -lean_object* x_71; +lean_object* x_72; lean_inc(x_14); lean_inc_ref(x_13); -x_71 = lean_get_init_fn_name_for(x_13, x_14); -if (lean_obj_tag(x_71) == 1) +x_72 = lean_get_init_fn_name_for(x_13, x_14); +if (lean_obj_tag(x_72) == 1) { -lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_103; lean_object* x_107; +lean_object* x_73; lean_object* x_74; lean_object* x_75; uint8_t x_105; lean_object* x_109; lean_inc_ref(x_13); -x_72 = lean_ctor_get(x_71, 0); -lean_inc(x_72); -lean_dec_ref(x_71); +x_73 = lean_ctor_get(x_72, 0); +lean_inc(x_73); +lean_dec_ref(x_72); lean_inc(x_14); lean_inc_ref(x_13); -x_107 = l_Lean_getBuiltinInitFnNameFor_x3f(x_13, x_14); -if (lean_obj_tag(x_107) == 0) +x_109 = l_Lean_getBuiltinInitFnNameFor_x3f(x_13, x_14); +if (lean_obj_tag(x_109) == 0) { -x_103 = x_57; -goto block_106; +x_105 = x_58; +goto block_108; } else { -lean_dec_ref(x_107); -x_103 = x_61; -goto block_106; +lean_dec_ref(x_109); +x_105 = x_62; +goto block_108; } -block_102: -{ -lean_object* x_75; lean_object* x_76; lean_object* x_77; -x_75 = l_Lean_IR_EmitC_emitMainFn___closed__18; -x_76 = lean_string_append(x_74, x_75); -lean_inc_ref(x_73); -x_77 = l_Lean_IR_EmitC_emitCName(x_72, x_73, x_76); -if (lean_obj_tag(x_77) == 0) +block_104: { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; -x_78 = lean_ctor_get(x_77, 1); -lean_inc(x_78); -lean_dec_ref(x_77); -x_79 = l_Lean_IR_EmitC_emitDeclInit___closed__0; -x_80 = lean_string_append(x_78, x_79); -x_81 = l_Lean_IR_EmitC_emitLn___closed__0; +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_76 = l_Lean_IR_EmitC_emitMainFn___closed__17; +x_77 = lean_string_append(x_75, x_76); +x_78 = l_Lean_IR_ExplicitBoxing_mkBoxedName(x_73); +lean_inc_ref(x_74); +x_79 = l_Lean_IR_EmitC_emitCName(x_78, x_74, x_77); +if (lean_obj_tag(x_79) == 0) +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_80 = lean_ctor_get(x_79, 1); +lean_inc(x_80); +lean_dec_ref(x_79); +x_81 = l_Lean_IR_EmitC_emitDeclInit___closed__0; x_82 = lean_string_append(x_80, x_81); -x_83 = l_Lean_IR_EmitC_emitDeclInit___closed__1; +x_83 = l_Lean_IR_EmitC_emitLn___closed__0; x_84 = lean_string_append(x_82, x_83); -x_85 = lean_string_append(x_84, x_81); -lean_inc_ref(x_73); +x_85 = l_Lean_IR_EmitC_emitDeclInit___closed__1; +x_86 = lean_string_append(x_84, x_85); +x_87 = lean_string_append(x_86, x_83); +lean_inc_ref(x_74); lean_inc(x_14); -x_86 = l_Lean_IR_EmitC_emitCName(x_14, x_73, x_85); -if (lean_obj_tag(x_86) == 0) -{ -lean_object* x_87; lean_object* x_88; uint8_t x_89; -x_87 = lean_ctor_get(x_86, 1); -lean_inc(x_87); -lean_dec_ref(x_86); -x_88 = l_Lean_IR_Decl_resultType(x_1); -x_89 = l_Lean_IR_IRType_isScalar(x_88); -if (x_89 == 0) +x_88 = l_Lean_IR_EmitC_emitCName(x_14, x_74, x_87); +if (lean_obj_tag(x_88) == 0) { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -lean_dec(x_88); -x_90 = l_Lean_IR_EmitC_emitDeclInit___closed__2; -x_91 = lean_string_append(x_87, x_90); -x_92 = lean_string_append(x_91, x_81); +lean_object* x_89; lean_object* x_90; uint8_t x_91; +x_89 = lean_ctor_get(x_88, 1); +lean_inc(x_89); +lean_dec_ref(x_88); +x_90 = l_Lean_IR_Decl_resultType(x_1); +x_91 = l_Lean_IR_IRType_isScalarOrStruct(x_90); +if (x_91 == 0) +{ +lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +lean_dec(x_90); +x_92 = l_Lean_IR_EmitC_emitDeclInit___closed__2; +x_93 = lean_string_append(x_89, x_92); +x_94 = lean_string_append(x_93, x_83); lean_inc(x_14); -x_93 = l_Lean_IR_EmitC_emitMarkPersistent(x_1, x_14, x_73, x_92); -if (lean_obj_tag(x_93) == 0) +x_95 = l_Lean_IR_EmitC_emitMarkPersistent(x_1, x_14, x_74, x_94); +if (lean_obj_tag(x_95) == 0) { -lean_object* x_94; -x_94 = lean_ctor_get(x_93, 1); -lean_inc(x_94); -lean_dec_ref(x_93); -x_62 = x_94; -goto block_68; +lean_object* x_96; +x_96 = lean_ctor_get(x_95, 1); +lean_inc(x_96); +lean_dec_ref(x_95); +x_63 = x_96; +goto block_69; } else { lean_dec(x_14); lean_dec_ref(x_13); -return x_93; +return x_95; } } else { -lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; -lean_dec_ref(x_73); -x_95 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0___redArg___closed__0; -x_96 = l_Lean_IR_getUnboxOpName(x_88); -lean_dec(x_88); -x_97 = lean_string_append(x_95, x_96); -lean_dec_ref(x_96); -x_98 = l_Lean_IR_EmitC_emitDeclInit___closed__3; -x_99 = lean_string_append(x_97, x_98); -x_100 = lean_string_append(x_87, x_99); +lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; +x_97 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__1; +x_98 = lean_string_append(x_89, x_97); +x_99 = l_Lean_IR_EmitC_emitUnboxFn(x_90, x_74, x_98); +lean_dec_ref(x_74); +lean_dec(x_90); +x_100 = lean_ctor_get(x_99, 1); +lean_inc(x_100); lean_dec_ref(x_99); -x_101 = lean_string_append(x_100, x_81); -x_62 = x_101; -goto block_68; +x_101 = l_Lean_IR_EmitC_emitDeclInit___closed__3; +x_102 = lean_string_append(x_100, x_101); +x_103 = lean_string_append(x_102, x_83); +x_63 = x_103; +goto block_69; } } else { -lean_dec_ref(x_73); +lean_dec_ref(x_74); lean_dec(x_14); lean_dec_ref(x_13); -return x_86; +return x_88; } } else { -lean_dec_ref(x_73); +lean_dec_ref(x_74); lean_dec(x_14); lean_dec_ref(x_13); -return x_77; +return x_79; } } -block_106: +block_108: { -if (x_103 == 0) +if (x_105 == 0) { -x_73 = x_2; -x_74 = x_3; -goto block_102; +x_74 = x_2; +x_75 = x_3; +goto block_104; } else { -lean_object* x_104; lean_object* x_105; -x_104 = l_Lean_IR_EmitC_emitDeclInit___closed__4; -x_105 = lean_string_append(x_3, x_104); -x_73 = x_2; -x_74 = x_105; -goto block_102; +lean_object* x_106; lean_object* x_107; +x_106 = l_Lean_IR_EmitC_emitDeclInit___closed__4; +x_107 = lean_string_append(x_3, x_106); +x_74 = x_2; +x_75 = x_107; +goto block_104; } } } else { -lean_object* x_108; -lean_dec(x_71); +lean_object* x_110; +lean_dec(x_72); lean_inc_ref(x_2); lean_inc(x_14); -x_108 = l_Lean_IR_EmitC_emitCName(x_14, x_2, x_3); -if (lean_obj_tag(x_108) == 0) -{ -lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; -x_109 = lean_ctor_get(x_108, 1); -lean_inc(x_109); -lean_dec_ref(x_108); -x_110 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0___redArg___closed__0; -x_111 = lean_string_append(x_109, x_110); +x_110 = l_Lean_IR_EmitC_emitCName(x_14, x_2, x_3); +if (lean_obj_tag(x_110) == 0) +{ +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; +x_111 = lean_ctor_get(x_110, 1); +lean_inc(x_111); +lean_dec_ref(x_110); +x_112 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__1; +x_113 = lean_string_append(x_111, x_112); lean_inc_ref(x_2); lean_inc(x_14); -x_112 = l_Lean_IR_EmitC_emitCInitName(x_14, x_2, x_111); -if (lean_obj_tag(x_112) == 0) -{ -lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; -x_113 = lean_ctor_get(x_112, 1); -lean_inc(x_113); -lean_dec_ref(x_112); -x_114 = l_Lean_IR_EmitC_emitDeclInit___closed__0; -x_115 = lean_string_append(x_113, x_114); -x_116 = l_Lean_IR_EmitC_emitLn___closed__0; +x_114 = l_Lean_IR_EmitC_emitCInitName(x_14, x_2, x_113); +if (lean_obj_tag(x_114) == 0) +{ +lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_115 = lean_ctor_get(x_114, 1); +lean_inc(x_115); +lean_dec_ref(x_114); +x_116 = l_Lean_IR_EmitC_emitDeclInit___closed__5; x_117 = lean_string_append(x_115, x_116); -x_118 = l_Lean_IR_EmitC_emitMarkPersistent(x_1, x_14, x_2, x_117); -return x_118; +x_118 = l_Lean_IR_EmitC_emitLn___closed__0; +x_119 = lean_string_append(x_117, x_118); +x_120 = l_Lean_IR_EmitC_emitMarkPersistent(x_1, x_14, x_2, x_119); +return x_120; } else { lean_dec(x_14); lean_dec_ref(x_2); -return x_112; +return x_114; } } else { lean_dec(x_14); lean_dec_ref(x_2); -return x_108; +return x_110; } } } -block_68: +block_69: { -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_63 = l_Lean_IR_EmitC_emitMainFn___closed__23; -x_64 = lean_string_append(x_62, x_63); -x_65 = l_Lean_IR_EmitC_emitLn___closed__0; -x_66 = lean_string_append(x_64, x_65); -x_67 = l_Lean_getBuiltinInitFnNameFor_x3f(x_13, x_14); -if (lean_obj_tag(x_67) == 0) +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_64 = l_Lean_IR_EmitC_emitMainFn___closed__22; +x_65 = lean_string_append(x_63, x_64); +x_66 = l_Lean_IR_EmitC_emitLn___closed__0; +x_67 = lean_string_append(x_65, x_66); +x_68 = l_Lean_getBuiltinInitFnNameFor_x3f(x_13, x_14); +if (lean_obj_tag(x_68) == 0) { -x_4 = x_66; -x_5 = x_57; +x_4 = x_67; +x_5 = x_58; goto block_12; } else { -lean_dec_ref(x_67); -x_4 = x_66; -x_5 = x_61; +lean_dec_ref(x_68); +x_4 = x_67; +x_5 = x_62; goto block_12; } } } else { -uint8_t x_119; +uint8_t x_121; lean_inc_ref(x_13); lean_inc(x_14); lean_inc_ref(x_13); -x_119 = l_Lean_isIOUnitBuiltinInitFn(x_13, x_14); -if (x_119 == 0) +x_121 = l_Lean_isIOUnitBuiltinInitFn(x_13, x_14); +if (x_121 == 0) { x_15 = x_2; x_16 = x_3; -goto block_56; +goto block_57; } else { -lean_object* x_120; lean_object* x_121; -x_120 = l_Lean_IR_EmitC_emitDeclInit___closed__4; -x_121 = lean_string_append(x_3, x_120); +lean_object* x_122; lean_object* x_123; +x_122 = l_Lean_IR_EmitC_emitDeclInit___closed__4; +x_123 = lean_string_append(x_3, x_122); x_15 = x_2; -x_16 = x_121; -goto block_56; +x_16 = x_123; +goto block_57; } } block_12: @@ -13916,7 +22493,7 @@ return x_7; else { lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_8 = l_Lean_IR_EmitC_emitMainFn___closed__12; +x_8 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; x_9 = lean_box(0); x_10 = lean_string_append(x_4, x_8); x_11 = lean_alloc_ctor(0, 2, 0); @@ -13925,89 +22502,90 @@ lean_ctor_set(x_11, 1, x_10); return x_11; } } -block_56: +block_57: { -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = l_Lean_IR_EmitC_emitMainFn___closed__18; +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = l_Lean_IR_EmitC_emitMainFn___closed__17; x_18 = lean_string_append(x_16, x_17); lean_inc(x_14); -x_19 = l_Lean_IR_EmitC_emitCName(x_14, x_15, x_18); -if (lean_obj_tag(x_19) == 0) +x_19 = l_Lean_IR_ExplicitBoxing_mkBoxedName(x_14); +x_20 = l_Lean_IR_EmitC_emitCName(x_19, x_15, x_18); +if (lean_obj_tag(x_20) == 0) { -uint8_t x_20; -x_20 = !lean_is_exclusive(x_19); -if (x_20 == 0) +uint8_t x_21; +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; -x_21 = lean_ctor_get(x_19, 1); -x_22 = lean_ctor_get(x_19, 0); -lean_dec(x_22); -x_23 = l_Lean_IR_EmitC_emitDeclInit___closed__0; -x_24 = lean_string_append(x_21, x_23); -x_25 = l_Lean_IR_EmitC_emitLn___closed__0; -x_26 = lean_string_append(x_24, x_25); -x_27 = l_Lean_IR_EmitC_emitDeclInit___closed__1; -x_28 = lean_string_append(x_26, x_27); -x_29 = lean_string_append(x_28, x_25); -x_30 = l_Lean_IR_EmitC_emitMainFn___closed__23; -x_31 = lean_string_append(x_29, x_30); -x_32 = lean_string_append(x_31, x_25); -x_33 = l_Lean_isIOUnitBuiltinInitFn(x_13, x_14); -if (x_33 == 0) +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; +x_22 = lean_ctor_get(x_20, 1); +x_23 = lean_ctor_get(x_20, 0); +lean_dec(x_23); +x_24 = l_Lean_IR_EmitC_emitDeclInit___closed__0; +x_25 = lean_string_append(x_22, x_24); +x_26 = l_Lean_IR_EmitC_emitLn___closed__0; +x_27 = lean_string_append(x_25, x_26); +x_28 = l_Lean_IR_EmitC_emitDeclInit___closed__1; +x_29 = lean_string_append(x_27, x_28); +x_30 = lean_string_append(x_29, x_26); +x_31 = l_Lean_IR_EmitC_emitMainFn___closed__22; +x_32 = lean_string_append(x_30, x_31); +x_33 = lean_string_append(x_32, x_26); +x_34 = l_Lean_isIOUnitBuiltinInitFn(x_13, x_14); +if (x_34 == 0) { -lean_object* x_34; -x_34 = lean_box(0); -lean_ctor_set(x_19, 1, x_32); -lean_ctor_set(x_19, 0, x_34); -return x_19; +lean_object* x_35; +x_35 = lean_box(0); +lean_ctor_set(x_20, 1, x_33); +lean_ctor_set(x_20, 0, x_35); +return x_20; } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = l_Lean_IR_EmitC_emitMainFn___closed__12; -x_36 = lean_box(0); -x_37 = lean_string_append(x_32, x_35); -lean_ctor_set(x_19, 1, x_37); -lean_ctor_set(x_19, 0, x_36); -return x_19; +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; +x_37 = lean_box(0); +x_38 = lean_string_append(x_33, x_36); +lean_ctor_set(x_20, 1, x_38); +lean_ctor_set(x_20, 0, x_37); +return x_20; } } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; -x_38 = lean_ctor_get(x_19, 1); -lean_inc(x_38); -lean_dec(x_19); -x_39 = l_Lean_IR_EmitC_emitDeclInit___closed__0; -x_40 = lean_string_append(x_38, x_39); -x_41 = l_Lean_IR_EmitC_emitLn___closed__0; -x_42 = lean_string_append(x_40, x_41); -x_43 = l_Lean_IR_EmitC_emitDeclInit___closed__1; -x_44 = lean_string_append(x_42, x_43); -x_45 = lean_string_append(x_44, x_41); -x_46 = l_Lean_IR_EmitC_emitMainFn___closed__23; -x_47 = lean_string_append(x_45, x_46); -x_48 = lean_string_append(x_47, x_41); -x_49 = l_Lean_isIOUnitBuiltinInitFn(x_13, x_14); -if (x_49 == 0) +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; +x_39 = lean_ctor_get(x_20, 1); +lean_inc(x_39); +lean_dec(x_20); +x_40 = l_Lean_IR_EmitC_emitDeclInit___closed__0; +x_41 = lean_string_append(x_39, x_40); +x_42 = l_Lean_IR_EmitC_emitLn___closed__0; +x_43 = lean_string_append(x_41, x_42); +x_44 = l_Lean_IR_EmitC_emitDeclInit___closed__1; +x_45 = lean_string_append(x_43, x_44); +x_46 = lean_string_append(x_45, x_42); +x_47 = l_Lean_IR_EmitC_emitMainFn___closed__22; +x_48 = lean_string_append(x_46, x_47); +x_49 = lean_string_append(x_48, x_42); +x_50 = l_Lean_isIOUnitBuiltinInitFn(x_13, x_14); +if (x_50 == 0) { -lean_object* x_50; lean_object* x_51; -x_50 = lean_box(0); -x_51 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_48); -return x_51; +lean_object* x_51; lean_object* x_52; +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_49); +return x_52; } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_52 = l_Lean_IR_EmitC_emitMainFn___closed__12; -x_53 = lean_box(0); -x_54 = lean_string_append(x_48, x_52); -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_53); -lean_ctor_set(x_55, 1, x_54); -return x_55; +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_53 = l_Lean_IR_EmitC_emitSpreadValue___closed__1; +x_54 = lean_box(0); +x_55 = lean_string_append(x_49, x_53); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_55); +return x_56; } } } @@ -14015,7 +22593,7 @@ else { lean_dec(x_14); lean_dec_ref(x_13); -return x_19; +return x_20; } } } @@ -14046,13 +22624,13 @@ if (x_7 == 0) { lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; x_8 = lean_array_uget(x_2, x_3); -x_9 = l_Lean_IR_EmitC_emitMainFn___closed__18; +x_9 = l_Lean_IR_EmitC_emitMainFn___closed__17; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitInitFn_spec__2___closed__0; x_12 = lean_string_append(x_10, x_11); x_13 = l_Lean_IR_EmitC_emitDeclInit___closed__1; -x_14 = l_Lean_IR_EmitC_emitMainFn___closed__23; +x_14 = l_Lean_IR_EmitC_emitMainFn___closed__22; lean_inc(x_1); x_15 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_15, 0, x_14); @@ -14241,13 +22819,13 @@ if (x_8 == 0) { lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; x_9 = lean_array_uget(x_2, x_3); -x_10 = l_Lean_IR_EmitC_emitMainFn___closed__18; +x_10 = l_Lean_IR_EmitC_emitMainFn___closed__17; x_11 = lean_string_append(x_10, x_9); lean_dec(x_9); x_12 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitInitFn_spec__2___closed__0; x_13 = lean_string_append(x_11, x_12); x_14 = l_Lean_IR_EmitC_emitDeclInit___closed__1; -x_15 = l_Lean_IR_EmitC_emitMainFn___closed__23; +x_15 = l_Lean_IR_EmitC_emitMainFn___closed__22; lean_inc(x_1); x_16 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_16, 0, x_15); @@ -14350,7 +22928,7 @@ static lean_object* _init_l_Lean_IR_EmitC_emitInitFn___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_IR_EmitC_emitMainFn___closed__37; +x_1 = l_Lean_IR_EmitC_emitMainFn___closed__36; x_2 = l_Lean_IR_EmitC_emitInitFn___closed__6; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); @@ -14630,12 +23208,20 @@ lean_dec_ref(x_1); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitC_main(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_IR_EmitC_main___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("void* memcpy(void* restrict, const void* restrict, size_t);", 59, 59); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_main___lam__0(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_inc_ref(x_1); -x_3 = l_Lean_IR_EmitC_emitFileHeader(x_1, x_2); +x_3 = l_Lean_IR_EmitC_emitFnDecls(x_1, x_2); if (lean_obj_tag(x_3) == 0) { lean_object* x_4; lean_object* x_5; @@ -14643,7 +23229,7 @@ x_4 = lean_ctor_get(x_3, 1); lean_inc(x_4); lean_dec_ref(x_3); lean_inc_ref(x_1); -x_5 = l_Lean_IR_EmitC_emitFnDecls(x_1, x_4); +x_5 = l_Lean_IR_EmitC_emitFns(x_1, x_4); if (lean_obj_tag(x_5) == 0) { lean_object* x_6; lean_object* x_7; @@ -14651,54 +23237,67 @@ x_6 = lean_ctor_get(x_5, 1); lean_inc(x_6); lean_dec_ref(x_5); lean_inc_ref(x_1); -x_7 = l_Lean_IR_EmitC_emitFns(x_1, x_6); +x_7 = l_Lean_IR_EmitC_emitInitFn(x_1, x_6); if (lean_obj_tag(x_7) == 0) { lean_object* x_8; lean_object* x_9; x_8 = lean_ctor_get(x_7, 1); lean_inc(x_8); lean_dec_ref(x_7); -lean_inc_ref(x_1); -x_9 = l_Lean_IR_EmitC_emitInitFn(x_1, x_8); +x_9 = l_Lean_IR_EmitC_emitMainFnIfNeeded(x_1, x_8); if (lean_obj_tag(x_9) == 0) { lean_object* x_10; lean_object* x_11; x_10 = lean_ctor_get(x_9, 1); lean_inc(x_10); lean_dec_ref(x_9); -x_11 = l_Lean_IR_EmitC_emitMainFnIfNeeded(x_1, x_10); -if (lean_obj_tag(x_11) == 0) -{ -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_11, 1); -lean_inc(x_12); -lean_dec_ref(x_11); -x_13 = l_Lean_IR_EmitC_emitFileFooter___redArg(x_12); -return x_13; +x_11 = l_Lean_IR_EmitC_emitFileFooter___redArg(x_10); +return x_11; } else { -return x_11; +return x_9; } } else { lean_dec_ref(x_1); -return x_9; +return x_7; } } else { lean_dec_ref(x_1); -return x_7; +return x_5; } } else { lean_dec_ref(x_1); -return x_5; +return x_3; } } +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_main(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +lean_inc_ref(x_1); +x_3 = l_Lean_IR_EmitC_emitFileHeader(x_1, x_2); +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_4 = lean_ctor_get(x_3, 1); +lean_inc(x_4); +lean_dec_ref(x_3); +x_5 = lean_alloc_closure((void*)(l_Lean_IR_EmitC_main___lam__0), 2, 0); +x_6 = l_Lean_IR_EmitC_main___closed__0; +x_7 = lean_string_append(x_4, x_6); +x_8 = l_Lean_IR_EmitC_emitLn___closed__0; +x_9 = lean_string_append(x_7, x_8); +x_10 = l_Lean_IR_EmitC_emitStructDecls___redArg(x_5, x_1, x_9); +return x_10; +} else { lean_dec_ref(x_1); @@ -14734,14 +23333,16 @@ LEAN_EXPORT lean_object* l_Lean_IR_emitC(lean_object* x_1, lean_object* x_2) { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; x_3 = l_Lean_IR_emitC___closed__1; x_4 = lean_box(0); -x_5 = l_Lean_IR_EmitC_emitFnDeclAux___closed__6; -x_6 = lean_alloc_ctor(0, 5, 0); +x_5 = l_Lean_IR_EmitC_emitFnDeclAux___closed__3; +x_6 = lean_alloc_ctor(0, 7, 0); lean_ctor_set(x_6, 0, x_1); lean_ctor_set(x_6, 1, x_2); lean_ctor_set(x_6, 2, x_3); lean_ctor_set(x_6, 3, x_4); lean_ctor_set(x_6, 4, x_5); -x_7 = l_panic___at___00Lean_IR_EmitC_toCType_spec__0___closed__0; +lean_ctor_set(x_6, 5, x_3); +lean_ctor_set(x_6, 6, x_3); +x_7 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__1; x_8 = l_Lean_IR_EmitC_main(x_6, x_7); if (lean_obj_tag(x_8) == 0) { @@ -14770,6 +23371,7 @@ lean_object* initialize_Lean_Compiler_IR_EmitUtil(uint8_t builtin); lean_object* initialize_Lean_Compiler_IR_NormIds(uint8_t builtin); lean_object* initialize_Lean_Compiler_IR_SimpCase(uint8_t builtin); lean_object* initialize_Lean_Compiler_IR_Boxing(uint8_t builtin); +lean_object* initialize_Lean_Compiler_IR_StructRC(uint8_t builtin); lean_object* initialize_Lean_Compiler_ModPkgExt(uint8_t builtin); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Compiler_IR_EmitC(uint8_t builtin) { @@ -14791,6 +23393,9 @@ lean_dec_ref(res); res = initialize_Lean_Compiler_IR_Boxing(builtin); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Compiler_IR_StructRC(builtin); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Lean_Compiler_ModPkgExt(builtin); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); @@ -14832,8 +23437,20 @@ l_Lean_IR_EmitC_argToCString___closed__0 = _init_l_Lean_IR_EmitC_argToCString___ lean_mark_persistent(l_Lean_IR_EmitC_argToCString___closed__0); l_Lean_IR_EmitC_argToCString___closed__1 = _init_l_Lean_IR_EmitC_argToCString___closed__1(); lean_mark_persistent(l_Lean_IR_EmitC_argToCString___closed__1); -l_panic___at___00Lean_IR_EmitC_toCType_spec__0___closed__0 = _init_l_panic___at___00Lean_IR_EmitC_toCType_spec__0___closed__0(); -lean_mark_persistent(l_panic___at___00Lean_IR_EmitC_toCType_spec__0___closed__0); +l_Lean_IR_EmitC_lookupStruct___closed__0 = _init_l_Lean_IR_EmitC_lookupStruct___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitC_lookupStruct___closed__0); +l_Lean_IR_EmitC_lookupStruct___closed__1 = _init_l_Lean_IR_EmitC_lookupStruct___closed__1(); +lean_mark_persistent(l_Lean_IR_EmitC_lookupStruct___closed__1); +l_Lean_IR_EmitC_structType___closed__0 = _init_l_Lean_IR_EmitC_structType___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitC_structType___closed__0); +l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg___closed__2 = _init_l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg___closed__2(); +lean_mark_persistent(l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg___closed__2); +l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg___closed__1 = _init_l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg___closed__1(); +lean_mark_persistent(l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg___closed__1); +l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg___closed__0 = _init_l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg___closed__0(); +lean_mark_persistent(l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg___closed__0); +l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg___closed__3 = _init_l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg___closed__3(); +lean_mark_persistent(l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_EmitC_toCType_spec__0_spec__0___redArg___closed__3); l_Lean_IR_EmitC_toCType___closed__0 = _init_l_Lean_IR_EmitC_toCType___closed__0(); lean_mark_persistent(l_Lean_IR_EmitC_toCType___closed__0); l_Lean_IR_EmitC_toCType___closed__1 = _init_l_Lean_IR_EmitC_toCType___closed__1(); @@ -14850,16 +23467,6 @@ l_Lean_IR_EmitC_toCType___closed__6 = _init_l_Lean_IR_EmitC_toCType___closed__6( lean_mark_persistent(l_Lean_IR_EmitC_toCType___closed__6); l_Lean_IR_EmitC_toCType___closed__7 = _init_l_Lean_IR_EmitC_toCType___closed__7(); lean_mark_persistent(l_Lean_IR_EmitC_toCType___closed__7); -l_Lean_IR_EmitC_toCType___closed__8 = _init_l_Lean_IR_EmitC_toCType___closed__8(); -lean_mark_persistent(l_Lean_IR_EmitC_toCType___closed__8); -l_Lean_IR_EmitC_toCType___closed__9 = _init_l_Lean_IR_EmitC_toCType___closed__9(); -lean_mark_persistent(l_Lean_IR_EmitC_toCType___closed__9); -l_Lean_IR_EmitC_toCType___closed__10 = _init_l_Lean_IR_EmitC_toCType___closed__10(); -lean_mark_persistent(l_Lean_IR_EmitC_toCType___closed__10); -l_Lean_IR_EmitC_toCType___closed__11 = _init_l_Lean_IR_EmitC_toCType___closed__11(); -lean_mark_persistent(l_Lean_IR_EmitC_toCType___closed__11); -l_Lean_IR_EmitC_toCType___closed__12 = _init_l_Lean_IR_EmitC_toCType___closed__12(); -lean_mark_persistent(l_Lean_IR_EmitC_toCType___closed__12); l_Lean_IR_EmitC_throwInvalidExportName___redArg___closed__0 = _init_l_Lean_IR_EmitC_throwInvalidExportName___redArg___closed__0(); lean_mark_persistent(l_Lean_IR_EmitC_throwInvalidExportName___redArg___closed__0); l_Lean_IR_EmitC_toCName___closed__0 = _init_l_Lean_IR_EmitC_toCName___closed__0(); @@ -14868,8 +23475,28 @@ l_Lean_IR_EmitC_toCName___closed__1 = _init_l_Lean_IR_EmitC_toCName___closed__1( lean_mark_persistent(l_Lean_IR_EmitC_toCName___closed__1); l_Lean_IR_EmitC_toCInitName___closed__0 = _init_l_Lean_IR_EmitC_toCInitName___closed__0(); lean_mark_persistent(l_Lean_IR_EmitC_toCInitName___closed__0); -l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___closed__0 = _init_l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___closed__0(); -lean_mark_persistent(l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg___closed__0); +l_Lean_IR_EmitC_emitSpreadArg___closed__0 = _init_l_Lean_IR_EmitC_emitSpreadArg___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitC_emitSpreadArg___closed__0); +l_Lean_IR_EmitC_emitSpreadArg___closed__1 = _init_l_Lean_IR_EmitC_emitSpreadArg___closed__1(); +lean_mark_persistent(l_Lean_IR_EmitC_emitSpreadArg___closed__1); +l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadArg_spec__0___redArg___closed__0 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadArg_spec__0___redArg___closed__0(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadArg_spec__0___redArg___closed__0); +l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__0 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__0(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__0); +l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__1 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__1(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitSpreadValue_spec__0___redArg___lam__0___closed__1); +l_Lean_IR_EmitC_emitSpreadValue___closed__0 = _init_l_Lean_IR_EmitC_emitSpreadValue___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitC_emitSpreadValue___closed__0); +l_Lean_IR_EmitC_emitSpreadValue___closed__1 = _init_l_Lean_IR_EmitC_emitSpreadValue___closed__1(); +lean_mark_persistent(l_Lean_IR_EmitC_emitSpreadValue___closed__1); +l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__0 = _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__0(); +lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__0); +l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__1 = _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__1(); +lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__1); +l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2 = _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2(); +lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitC_emitSpreads_spec__0___closed__2); +l_Lean_IR_EmitC_emitFullAppArgs___closed__0 = _init_l_Lean_IR_EmitC_emitFullAppArgs___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitC_emitFullAppArgs___closed__0); l_Lean_IR_EmitC_emitFnDeclAux___closed__0 = _init_l_Lean_IR_EmitC_emitFnDeclAux___closed__0(); lean_mark_persistent(l_Lean_IR_EmitC_emitFnDeclAux___closed__0); l_Lean_IR_EmitC_emitFnDeclAux___closed__1 = _init_l_Lean_IR_EmitC_emitFnDeclAux___closed__1(); @@ -14884,12 +23511,6 @@ l_Lean_IR_EmitC_emitFnDeclAux___closed__5 = _init_l_Lean_IR_EmitC_emitFnDeclAux_ lean_mark_persistent(l_Lean_IR_EmitC_emitFnDeclAux___closed__5); l_Lean_IR_EmitC_emitFnDeclAux___closed__6 = _init_l_Lean_IR_EmitC_emitFnDeclAux___closed__6(); lean_mark_persistent(l_Lean_IR_EmitC_emitFnDeclAux___closed__6); -l_Lean_IR_EmitC_emitFnDeclAux___closed__7 = _init_l_Lean_IR_EmitC_emitFnDeclAux___closed__7(); -lean_mark_persistent(l_Lean_IR_EmitC_emitFnDeclAux___closed__7); -l_Lean_IR_EmitC_emitFnDeclAux___closed__8 = _init_l_Lean_IR_EmitC_emitFnDeclAux___closed__8(); -lean_mark_persistent(l_Lean_IR_EmitC_emitFnDeclAux___closed__8); -l_Lean_IR_EmitC_emitFnDeclAux___closed__9 = _init_l_Lean_IR_EmitC_emitFnDeclAux___closed__9(); -lean_mark_persistent(l_Lean_IR_EmitC_emitFnDeclAux___closed__9); l_Lean_IR_EmitC_emitFnDecls___closed__0 = _init_l_Lean_IR_EmitC_emitFnDecls___closed__0(); lean_mark_persistent(l_Lean_IR_EmitC_emitFnDecls___closed__0); l_List_forM___at___00Lean_IR_EmitC_emitFnDecls_spec__3___closed__0 = _init_l_List_forM___at___00Lean_IR_EmitC_emitFnDecls_spec__3___closed__0(); @@ -15010,8 +23631,6 @@ l_Lean_IR_EmitC_emitMainFn___closed__54 = _init_l_Lean_IR_EmitC_emitMainFn___clo lean_mark_persistent(l_Lean_IR_EmitC_emitMainFn___closed__54); l_Lean_IR_EmitC_emitMainFn___closed__55 = _init_l_Lean_IR_EmitC_emitMainFn___closed__55(); lean_mark_persistent(l_Lean_IR_EmitC_emitMainFn___closed__55); -l_Lean_IR_EmitC_emitMainFn___closed__56 = _init_l_Lean_IR_EmitC_emitMainFn___closed__56(); -lean_mark_persistent(l_Lean_IR_EmitC_emitMainFn___closed__56); l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__0 = _init_l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__0(); lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__0); l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__1 = _init_l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__1(); @@ -15020,6 +23639,8 @@ l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_Em lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__2); l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__3 = _init_l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__3(); lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__3); +l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__4 = _init_l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__4(); +lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFileHeader_spec__0___redArg___closed__4); l_Lean_IR_EmitC_emitFileHeader___closed__0 = _init_l_Lean_IR_EmitC_emitFileHeader___closed__0(); lean_mark_persistent(l_Lean_IR_EmitC_emitFileHeader___closed__0); l_Lean_IR_EmitC_emitFileHeader___closed__1 = _init_l_Lean_IR_EmitC_emitFileHeader___closed__1(); @@ -15078,24 +23699,46 @@ l_Lean_IR_EmitC_throwUnknownVar___redArg___closed__0 = _init_l_Lean_IR_EmitC_thr lean_mark_persistent(l_Lean_IR_EmitC_throwUnknownVar___redArg___closed__0); l_Lean_IR_EmitC_getJPParams___closed__0 = _init_l_Lean_IR_EmitC_getJPParams___closed__0(); lean_mark_persistent(l_Lean_IR_EmitC_getJPParams___closed__0); -l_Lean_IR_EmitC_declareVar___redArg___closed__0 = _init_l_Lean_IR_EmitC_declareVar___redArg___closed__0(); -lean_mark_persistent(l_Lean_IR_EmitC_declareVar___redArg___closed__0); +l_Lean_IR_EmitC_declareVar___closed__0 = _init_l_Lean_IR_EmitC_declareVar___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitC_declareVar___closed__0); l_Lean_IR_EmitC_emitTag___redArg___closed__0 = _init_l_Lean_IR_EmitC_emitTag___redArg___closed__0(); lean_mark_persistent(l_Lean_IR_EmitC_emitTag___redArg___closed__0); -l_Lean_IR_EmitC_emitInc___redArg___closed__0 = _init_l_Lean_IR_EmitC_emitInc___redArg___closed__0(); -lean_mark_persistent(l_Lean_IR_EmitC_emitInc___redArg___closed__0); -l_Lean_IR_EmitC_emitInc___redArg___closed__1 = _init_l_Lean_IR_EmitC_emitInc___redArg___closed__1(); -lean_mark_persistent(l_Lean_IR_EmitC_emitInc___redArg___closed__1); -l_Lean_IR_EmitC_emitInc___redArg___closed__2 = _init_l_Lean_IR_EmitC_emitInc___redArg___closed__2(); -lean_mark_persistent(l_Lean_IR_EmitC_emitInc___redArg___closed__2); -l_Lean_IR_EmitC_emitInc___redArg___closed__3 = _init_l_Lean_IR_EmitC_emitInc___redArg___closed__3(); -lean_mark_persistent(l_Lean_IR_EmitC_emitInc___redArg___closed__3); -l_Lean_IR_EmitC_emitInc___redArg___closed__4 = _init_l_Lean_IR_EmitC_emitInc___redArg___closed__4(); -lean_mark_persistent(l_Lean_IR_EmitC_emitInc___redArg___closed__4); -l_Lean_IR_EmitC_emitDec___redArg___closed__0 = _init_l_Lean_IR_EmitC_emitDec___redArg___closed__0(); -lean_mark_persistent(l_Lean_IR_EmitC_emitDec___redArg___closed__0); -l_Lean_IR_EmitC_emitDec___redArg___closed__1 = _init_l_Lean_IR_EmitC_emitDec___redArg___closed__1(); -lean_mark_persistent(l_Lean_IR_EmitC_emitDec___redArg___closed__1); +l_Lean_IR_EmitC_emitTag___redArg___closed__1 = _init_l_Lean_IR_EmitC_emitTag___redArg___closed__1(); +lean_mark_persistent(l_Lean_IR_EmitC_emitTag___redArg___closed__1); +l_Lean_IR_EmitC_structIncFnPrefix___closed__0 = _init_l_Lean_IR_EmitC_structIncFnPrefix___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitC_structIncFnPrefix___closed__0); +l_Lean_IR_EmitC_structIncFnPrefix = _init_l_Lean_IR_EmitC_structIncFnPrefix(); +lean_mark_persistent(l_Lean_IR_EmitC_structIncFnPrefix); +l_Lean_IR_EmitC_structDecFnPrefix___closed__0 = _init_l_Lean_IR_EmitC_structDecFnPrefix___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitC_structDecFnPrefix___closed__0); +l_Lean_IR_EmitC_structDecFnPrefix = _init_l_Lean_IR_EmitC_structDecFnPrefix(); +lean_mark_persistent(l_Lean_IR_EmitC_structDecFnPrefix); +l_Lean_IR_EmitC_structReshapeFnPrefix___closed__0 = _init_l_Lean_IR_EmitC_structReshapeFnPrefix___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitC_structReshapeFnPrefix___closed__0); +l_Lean_IR_EmitC_structReshapeFnPrefix = _init_l_Lean_IR_EmitC_structReshapeFnPrefix(); +lean_mark_persistent(l_Lean_IR_EmitC_structReshapeFnPrefix); +l_Lean_IR_EmitC_structBoxFnPrefix___closed__0 = _init_l_Lean_IR_EmitC_structBoxFnPrefix___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitC_structBoxFnPrefix___closed__0); +l_Lean_IR_EmitC_structBoxFnPrefix = _init_l_Lean_IR_EmitC_structBoxFnPrefix(); +lean_mark_persistent(l_Lean_IR_EmitC_structBoxFnPrefix); +l_Lean_IR_EmitC_structUnboxFnPrefix___closed__0 = _init_l_Lean_IR_EmitC_structUnboxFnPrefix___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitC_structUnboxFnPrefix___closed__0); +l_Lean_IR_EmitC_structUnboxFnPrefix = _init_l_Lean_IR_EmitC_structUnboxFnPrefix(); +lean_mark_persistent(l_Lean_IR_EmitC_structUnboxFnPrefix); +l_Lean_IR_EmitC_emitIncOfType___closed__0 = _init_l_Lean_IR_EmitC_emitIncOfType___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitC_emitIncOfType___closed__0); +l_Lean_IR_EmitC_emitIncOfType___closed__1 = _init_l_Lean_IR_EmitC_emitIncOfType___closed__1(); +lean_mark_persistent(l_Lean_IR_EmitC_emitIncOfType___closed__1); +l_Lean_IR_EmitC_emitIncOfType___closed__2 = _init_l_Lean_IR_EmitC_emitIncOfType___closed__2(); +lean_mark_persistent(l_Lean_IR_EmitC_emitIncOfType___closed__2); +l_Lean_IR_EmitC_emitIncOfType___closed__3 = _init_l_Lean_IR_EmitC_emitIncOfType___closed__3(); +lean_mark_persistent(l_Lean_IR_EmitC_emitIncOfType___closed__3); +l_Lean_IR_EmitC_emitIncOfType___closed__4 = _init_l_Lean_IR_EmitC_emitIncOfType___closed__4(); +lean_mark_persistent(l_Lean_IR_EmitC_emitIncOfType___closed__4); +l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDecOfType_spec__0___redArg___closed__0 = _init_l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDecOfType_spec__0___redArg___closed__0(); +lean_mark_persistent(l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDecOfType_spec__0___redArg___closed__0); +l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDecOfType_spec__0___redArg___closed__1 = _init_l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDecOfType_spec__0___redArg___closed__1(); +lean_mark_persistent(l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDecOfType_spec__0___redArg___closed__1); l_Lean_IR_EmitC_emitDel___redArg___closed__0 = _init_l_Lean_IR_EmitC_emitDel___redArg___closed__0(); lean_mark_persistent(l_Lean_IR_EmitC_emitDel___redArg___closed__0); l_Lean_IR_EmitC_emitSetTag___redArg___closed__0 = _init_l_Lean_IR_EmitC_emitSetTag___redArg___closed__0(); @@ -15106,24 +23749,46 @@ l_Lean_IR_EmitC_emitOffset___redArg___closed__0 = _init_l_Lean_IR_EmitC_emitOffs lean_mark_persistent(l_Lean_IR_EmitC_emitOffset___redArg___closed__0); l_Lean_IR_EmitC_emitOffset___redArg___closed__1 = _init_l_Lean_IR_EmitC_emitOffset___redArg___closed__1(); lean_mark_persistent(l_Lean_IR_EmitC_emitOffset___redArg___closed__1); -l_Lean_IR_EmitC_emitUSet___redArg___closed__0 = _init_l_Lean_IR_EmitC_emitUSet___redArg___closed__0(); -lean_mark_persistent(l_Lean_IR_EmitC_emitUSet___redArg___closed__0); -l_Lean_IR_EmitC_emitSSet___redArg___closed__0 = _init_l_Lean_IR_EmitC_emitSSet___redArg___closed__0(); -lean_mark_persistent(l_Lean_IR_EmitC_emitSSet___redArg___closed__0); -l_Lean_IR_EmitC_emitSSet___redArg___closed__1 = _init_l_Lean_IR_EmitC_emitSSet___redArg___closed__1(); -lean_mark_persistent(l_Lean_IR_EmitC_emitSSet___redArg___closed__1); -l_Lean_IR_EmitC_emitSSet___redArg___closed__2 = _init_l_Lean_IR_EmitC_emitSSet___redArg___closed__2(); -lean_mark_persistent(l_Lean_IR_EmitC_emitSSet___redArg___closed__2); -l_Lean_IR_EmitC_emitSSet___redArg___closed__3 = _init_l_Lean_IR_EmitC_emitSSet___redArg___closed__3(); -lean_mark_persistent(l_Lean_IR_EmitC_emitSSet___redArg___closed__3); -l_Lean_IR_EmitC_emitSSet___redArg___closed__4 = _init_l_Lean_IR_EmitC_emitSSet___redArg___closed__4(); -lean_mark_persistent(l_Lean_IR_EmitC_emitSSet___redArg___closed__4); -l_Lean_IR_EmitC_emitSSet___redArg___closed__5 = _init_l_Lean_IR_EmitC_emitSSet___redArg___closed__5(); -lean_mark_persistent(l_Lean_IR_EmitC_emitSSet___redArg___closed__5); -l_Lean_IR_EmitC_emitSSet___redArg___closed__6 = _init_l_Lean_IR_EmitC_emitSSet___redArg___closed__6(); -lean_mark_persistent(l_Lean_IR_EmitC_emitSSet___redArg___closed__6); -l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0___redArg___closed__0 = _init_l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0___redArg___closed__0(); -lean_mark_persistent(l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitJmp_spec__0___redArg___closed__0); +l_panic___at___00Lean_IR_EmitC_emitUSet_spec__0___closed__0 = _init_l_panic___at___00Lean_IR_EmitC_emitUSet_spec__0___closed__0(); +lean_mark_persistent(l_panic___at___00Lean_IR_EmitC_emitUSet_spec__0___closed__0); +l_Lean_IR_EmitC_emitUSet___closed__0 = _init_l_Lean_IR_EmitC_emitUSet___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitC_emitUSet___closed__0); +l_Lean_IR_EmitC_emitUSet___closed__1 = _init_l_Lean_IR_EmitC_emitUSet___closed__1(); +lean_mark_persistent(l_Lean_IR_EmitC_emitUSet___closed__1); +l_Lean_IR_EmitC_emitUSet___closed__2 = _init_l_Lean_IR_EmitC_emitUSet___closed__2(); +lean_mark_persistent(l_Lean_IR_EmitC_emitUSet___closed__2); +l_Lean_IR_EmitC_emitUSet___closed__3 = _init_l_Lean_IR_EmitC_emitUSet___closed__3(); +lean_mark_persistent(l_Lean_IR_EmitC_emitUSet___closed__3); +l_Lean_IR_EmitC_emitUSet___closed__4 = _init_l_Lean_IR_EmitC_emitUSet___closed__4(); +lean_mark_persistent(l_Lean_IR_EmitC_emitUSet___closed__4); +l_Lean_IR_EmitC_emitUSet___closed__5 = _init_l_Lean_IR_EmitC_emitUSet___closed__5(); +lean_mark_persistent(l_Lean_IR_EmitC_emitUSet___closed__5); +l_Lean_IR_EmitC_emitUSet___closed__6 = _init_l_Lean_IR_EmitC_emitUSet___closed__6(); +lean_mark_persistent(l_Lean_IR_EmitC_emitUSet___closed__6); +l_Lean_IR_EmitC_emitUSet___closed__7 = _init_l_Lean_IR_EmitC_emitUSet___closed__7(); +lean_mark_persistent(l_Lean_IR_EmitC_emitUSet___closed__7); +l_Lean_IR_EmitC_emitSSet___closed__0 = _init_l_Lean_IR_EmitC_emitSSet___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitC_emitSSet___closed__0); +l_Lean_IR_EmitC_emitSSet___closed__1 = _init_l_Lean_IR_EmitC_emitSSet___closed__1(); +lean_mark_persistent(l_Lean_IR_EmitC_emitSSet___closed__1); +l_Lean_IR_EmitC_emitSSet___closed__2 = _init_l_Lean_IR_EmitC_emitSSet___closed__2(); +lean_mark_persistent(l_Lean_IR_EmitC_emitSSet___closed__2); +l_Lean_IR_EmitC_emitSSet___closed__3 = _init_l_Lean_IR_EmitC_emitSSet___closed__3(); +lean_mark_persistent(l_Lean_IR_EmitC_emitSSet___closed__3); +l_Lean_IR_EmitC_emitSSet___closed__4 = _init_l_Lean_IR_EmitC_emitSSet___closed__4(); +lean_mark_persistent(l_Lean_IR_EmitC_emitSSet___closed__4); +l_Lean_IR_EmitC_emitSSet___closed__5 = _init_l_Lean_IR_EmitC_emitSSet___closed__5(); +lean_mark_persistent(l_Lean_IR_EmitC_emitSSet___closed__5); +l_Lean_IR_EmitC_emitSSet___closed__6 = _init_l_Lean_IR_EmitC_emitSSet___closed__6(); +lean_mark_persistent(l_Lean_IR_EmitC_emitSSet___closed__6); +l_Lean_IR_EmitC_emitSSet___closed__7 = _init_l_Lean_IR_EmitC_emitSSet___closed__7(); +lean_mark_persistent(l_Lean_IR_EmitC_emitSSet___closed__7); +l_Lean_IR_EmitC_emitSSet___closed__8 = _init_l_Lean_IR_EmitC_emitSSet___closed__8(); +lean_mark_persistent(l_Lean_IR_EmitC_emitSSet___closed__8); +l_Lean_IR_EmitC_emitSSet___closed__9 = _init_l_Lean_IR_EmitC_emitSSet___closed__9(); +lean_mark_persistent(l_Lean_IR_EmitC_emitSSet___closed__9); +l_Lean_IR_EmitC_emitSSet___closed__10 = _init_l_Lean_IR_EmitC_emitSSet___closed__10(); +lean_mark_persistent(l_Lean_IR_EmitC_emitSSet___closed__10); l_Lean_IR_EmitC_emitJmp___closed__0 = _init_l_Lean_IR_EmitC_emitJmp___closed__0(); lean_mark_persistent(l_Lean_IR_EmitC_emitJmp___closed__0); l_Lean_IR_EmitC_emitJmp___closed__1 = _init_l_Lean_IR_EmitC_emitJmp___closed__1(); @@ -15136,6 +23801,12 @@ l_Lean_IR_EmitC_emitAllocCtor___redArg___closed__0 = _init_l_Lean_IR_EmitC_emitA lean_mark_persistent(l_Lean_IR_EmitC_emitAllocCtor___redArg___closed__0); l_Lean_IR_EmitC_emitCtor___closed__0 = _init_l_Lean_IR_EmitC_emitCtor___closed__0(); lean_mark_persistent(l_Lean_IR_EmitC_emitCtor___closed__0); +l_Lean_IR_EmitC_emitCtor___closed__1 = _init_l_Lean_IR_EmitC_emitCtor___closed__1(); +lean_mark_persistent(l_Lean_IR_EmitC_emitCtor___closed__1); +l_Lean_IR_EmitC_emitCtor___closed__2 = _init_l_Lean_IR_EmitC_emitCtor___closed__2(); +lean_mark_persistent(l_Lean_IR_EmitC_emitCtor___closed__2); +l_Lean_IR_EmitC_emitCtor___closed__3 = _init_l_Lean_IR_EmitC_emitCtor___closed__3(); +lean_mark_persistent(l_Lean_IR_EmitC_emitCtor___closed__3); l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitReset_spec__0___redArg___closed__0 = _init_l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitReset_spec__0___redArg___closed__0(); lean_mark_persistent(l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitReset_spec__0___redArg___closed__0); l_Lean_IR_EmitC_emitReset___closed__0 = _init_l_Lean_IR_EmitC_emitReset___closed__0(); @@ -15150,28 +23821,32 @@ l_Lean_IR_EmitC_emitReuse___closed__0 = _init_l_Lean_IR_EmitC_emitReuse___closed lean_mark_persistent(l_Lean_IR_EmitC_emitReuse___closed__0); l_Lean_IR_EmitC_emitReuse___closed__1 = _init_l_Lean_IR_EmitC_emitReuse___closed__1(); lean_mark_persistent(l_Lean_IR_EmitC_emitReuse___closed__1); -l_Lean_IR_EmitC_emitProj___redArg___closed__0 = _init_l_Lean_IR_EmitC_emitProj___redArg___closed__0(); -lean_mark_persistent(l_Lean_IR_EmitC_emitProj___redArg___closed__0); -l_Lean_IR_EmitC_emitUProj___redArg___closed__0 = _init_l_Lean_IR_EmitC_emitUProj___redArg___closed__0(); -lean_mark_persistent(l_Lean_IR_EmitC_emitUProj___redArg___closed__0); -l_Lean_IR_EmitC_emitSProj___redArg___closed__0 = _init_l_Lean_IR_EmitC_emitSProj___redArg___closed__0(); -lean_mark_persistent(l_Lean_IR_EmitC_emitSProj___redArg___closed__0); -l_Lean_IR_EmitC_emitSProj___redArg___closed__1 = _init_l_Lean_IR_EmitC_emitSProj___redArg___closed__1(); -lean_mark_persistent(l_Lean_IR_EmitC_emitSProj___redArg___closed__1); -l_Lean_IR_EmitC_emitSProj___redArg___closed__2 = _init_l_Lean_IR_EmitC_emitSProj___redArg___closed__2(); -lean_mark_persistent(l_Lean_IR_EmitC_emitSProj___redArg___closed__2); -l_Lean_IR_EmitC_emitSProj___redArg___closed__3 = _init_l_Lean_IR_EmitC_emitSProj___redArg___closed__3(); -lean_mark_persistent(l_Lean_IR_EmitC_emitSProj___redArg___closed__3); -l_Lean_IR_EmitC_emitSProj___redArg___closed__4 = _init_l_Lean_IR_EmitC_emitSProj___redArg___closed__4(); -lean_mark_persistent(l_Lean_IR_EmitC_emitSProj___redArg___closed__4); -l_Lean_IR_EmitC_emitSProj___redArg___closed__5 = _init_l_Lean_IR_EmitC_emitSProj___redArg___closed__5(); -lean_mark_persistent(l_Lean_IR_EmitC_emitSProj___redArg___closed__5); -l_Lean_IR_EmitC_emitSimpleExternalCall___closed__0 = _init_l_Lean_IR_EmitC_emitSimpleExternalCall___closed__0(); -lean_mark_persistent(l_Lean_IR_EmitC_emitSimpleExternalCall___closed__0); +l_Lean_IR_EmitC_emitProj___closed__0 = _init_l_Lean_IR_EmitC_emitProj___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitC_emitProj___closed__0); +l_Lean_IR_EmitC_emitUProj___closed__0 = _init_l_Lean_IR_EmitC_emitUProj___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitC_emitUProj___closed__0); +l_Lean_IR_EmitC_emitUProj___closed__1 = _init_l_Lean_IR_EmitC_emitUProj___closed__1(); +lean_mark_persistent(l_Lean_IR_EmitC_emitUProj___closed__1); +l_Lean_IR_EmitC_emitUProj___closed__2 = _init_l_Lean_IR_EmitC_emitUProj___closed__2(); +lean_mark_persistent(l_Lean_IR_EmitC_emitUProj___closed__2); +l_Lean_IR_EmitC_emitUProj___closed__3 = _init_l_Lean_IR_EmitC_emitUProj___closed__3(); +lean_mark_persistent(l_Lean_IR_EmitC_emitUProj___closed__3); +l_Lean_IR_EmitC_emitSProj___closed__0 = _init_l_Lean_IR_EmitC_emitSProj___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitC_emitSProj___closed__0); +l_Lean_IR_EmitC_emitSProj___closed__1 = _init_l_Lean_IR_EmitC_emitSProj___closed__1(); +lean_mark_persistent(l_Lean_IR_EmitC_emitSProj___closed__1); +l_Lean_IR_EmitC_emitSProj___closed__2 = _init_l_Lean_IR_EmitC_emitSProj___closed__2(); +lean_mark_persistent(l_Lean_IR_EmitC_emitSProj___closed__2); +l_Lean_IR_EmitC_emitSProj___closed__3 = _init_l_Lean_IR_EmitC_emitSProj___closed__3(); +lean_mark_persistent(l_Lean_IR_EmitC_emitSProj___closed__3); +l_Lean_IR_EmitC_emitSProj___closed__4 = _init_l_Lean_IR_EmitC_emitSProj___closed__4(); +lean_mark_persistent(l_Lean_IR_EmitC_emitSProj___closed__4); +l_Lean_IR_EmitC_emitSProj___closed__5 = _init_l_Lean_IR_EmitC_emitSProj___closed__5(); +lean_mark_persistent(l_Lean_IR_EmitC_emitSProj___closed__5); +l_Lean_IR_EmitC_emitSProj___closed__6 = _init_l_Lean_IR_EmitC_emitSProj___closed__6(); +lean_mark_persistent(l_Lean_IR_EmitC_emitSProj___closed__6); l_Lean_IR_EmitC_emitExternCall___closed__0 = _init_l_Lean_IR_EmitC_emitExternCall___closed__0(); lean_mark_persistent(l_Lean_IR_EmitC_emitExternCall___closed__0); -l_Lean_IR_EmitC_emitFullApp___closed__0 = _init_l_Lean_IR_EmitC_emitFullApp___closed__0(); -lean_mark_persistent(l_Lean_IR_EmitC_emitFullApp___closed__0); l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitPartialApp_spec__0___redArg___closed__0 = _init_l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitPartialApp_spec__0___redArg___closed__0(); lean_mark_persistent(l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitPartialApp_spec__0___redArg___closed__0); l_Lean_IR_EmitC_emitPartialApp___closed__0 = _init_l_Lean_IR_EmitC_emitPartialApp___closed__0(); @@ -15188,18 +23863,18 @@ l_Lean_IR_EmitC_emitApp___closed__3 = _init_l_Lean_IR_EmitC_emitApp___closed__3( lean_mark_persistent(l_Lean_IR_EmitC_emitApp___closed__3); l_Lean_IR_EmitC_emitApp___closed__4 = _init_l_Lean_IR_EmitC_emitApp___closed__4(); lean_mark_persistent(l_Lean_IR_EmitC_emitApp___closed__4); -l_Lean_IR_EmitC_emitBoxFn___redArg___closed__0 = _init_l_Lean_IR_EmitC_emitBoxFn___redArg___closed__0(); -lean_mark_persistent(l_Lean_IR_EmitC_emitBoxFn___redArg___closed__0); -l_Lean_IR_EmitC_emitBoxFn___redArg___closed__1 = _init_l_Lean_IR_EmitC_emitBoxFn___redArg___closed__1(); -lean_mark_persistent(l_Lean_IR_EmitC_emitBoxFn___redArg___closed__1); -l_Lean_IR_EmitC_emitBoxFn___redArg___closed__2 = _init_l_Lean_IR_EmitC_emitBoxFn___redArg___closed__2(); -lean_mark_persistent(l_Lean_IR_EmitC_emitBoxFn___redArg___closed__2); -l_Lean_IR_EmitC_emitBoxFn___redArg___closed__3 = _init_l_Lean_IR_EmitC_emitBoxFn___redArg___closed__3(); -lean_mark_persistent(l_Lean_IR_EmitC_emitBoxFn___redArg___closed__3); -l_Lean_IR_EmitC_emitBoxFn___redArg___closed__4 = _init_l_Lean_IR_EmitC_emitBoxFn___redArg___closed__4(); -lean_mark_persistent(l_Lean_IR_EmitC_emitBoxFn___redArg___closed__4); -l_Lean_IR_EmitC_emitBoxFn___redArg___closed__5 = _init_l_Lean_IR_EmitC_emitBoxFn___redArg___closed__5(); -lean_mark_persistent(l_Lean_IR_EmitC_emitBoxFn___redArg___closed__5); +l_Lean_IR_EmitC_emitBoxFn___closed__0 = _init_l_Lean_IR_EmitC_emitBoxFn___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitC_emitBoxFn___closed__0); +l_Lean_IR_EmitC_emitBoxFn___closed__1 = _init_l_Lean_IR_EmitC_emitBoxFn___closed__1(); +lean_mark_persistent(l_Lean_IR_EmitC_emitBoxFn___closed__1); +l_Lean_IR_EmitC_emitBoxFn___closed__2 = _init_l_Lean_IR_EmitC_emitBoxFn___closed__2(); +lean_mark_persistent(l_Lean_IR_EmitC_emitBoxFn___closed__2); +l_Lean_IR_EmitC_emitBoxFn___closed__3 = _init_l_Lean_IR_EmitC_emitBoxFn___closed__3(); +lean_mark_persistent(l_Lean_IR_EmitC_emitBoxFn___closed__3); +l_Lean_IR_EmitC_emitBoxFn___closed__4 = _init_l_Lean_IR_EmitC_emitBoxFn___closed__4(); +lean_mark_persistent(l_Lean_IR_EmitC_emitBoxFn___closed__4); +l_Lean_IR_EmitC_emitBoxFn___closed__5 = _init_l_Lean_IR_EmitC_emitBoxFn___closed__5(); +lean_mark_persistent(l_Lean_IR_EmitC_emitBoxFn___closed__5); l_Lean_IR_EmitC_emitIsShared___redArg___closed__0 = _init_l_Lean_IR_EmitC_emitIsShared___redArg___closed__0(); lean_mark_persistent(l_Lean_IR_EmitC_emitIsShared___redArg___closed__0); l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__0 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_quoteString_spec__0___redArg___closed__0(); @@ -15248,8 +23923,6 @@ l_Lean_IR_EmitC_emitTailCall___closed__2 = _init_l_Lean_IR_EmitC_emitTailCall___ lean_mark_persistent(l_Lean_IR_EmitC_emitTailCall___closed__2); l_Lean_IR_EmitC_emitTailCall___closed__3 = _init_l_Lean_IR_EmitC_emitTailCall___closed__3(); lean_mark_persistent(l_Lean_IR_EmitC_emitTailCall___closed__3); -l_Lean_IR_EmitC_emitTailCall___closed__4 = _init_l_Lean_IR_EmitC_emitTailCall___closed__4(); -lean_mark_persistent(l_Lean_IR_EmitC_emitTailCall___closed__4); l_Lean_IR_EmitC_emitIf___closed__0 = _init_l_Lean_IR_EmitC_emitIf___closed__0(); lean_mark_persistent(l_Lean_IR_EmitC_emitIf___closed__0); l_Lean_IR_EmitC_emitIf___closed__1 = _init_l_Lean_IR_EmitC_emitIf___closed__1(); @@ -15270,20 +23943,130 @@ l_Lean_IR_EmitC_emitBlock___closed__0 = _init_l_Lean_IR_EmitC_emitBlock___closed lean_mark_persistent(l_Lean_IR_EmitC_emitBlock___closed__0); l_Lean_IR_EmitC_emitBlock___closed__1 = _init_l_Lean_IR_EmitC_emitBlock___closed__1(); lean_mark_persistent(l_Lean_IR_EmitC_emitBlock___closed__1); +l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__0___redArg___closed__0 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__0___redArg___closed__0(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__0___redArg___closed__0); +l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__1___redArg___closed__0 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__1___redArg___closed__0(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructDefn_spec__1___redArg___closed__0); +l_Lean_IR_EmitC_emitStructDefn___closed__0 = _init_l_Lean_IR_EmitC_emitStructDefn___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructDefn___closed__0); +l_Lean_IR_EmitC_emitStructDefn___closed__1 = _init_l_Lean_IR_EmitC_emitStructDefn___closed__1(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructDefn___closed__1); +l_Lean_IR_EmitC_emitStructDefn___closed__2 = _init_l_Lean_IR_EmitC_emitStructDefn___closed__2(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructDefn___closed__2); +l_Lean_IR_EmitC_emitStructDefn___closed__3 = _init_l_Lean_IR_EmitC_emitStructDefn___closed__3(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructDefn___closed__3); +l_Lean_IR_EmitC_emitStructDefn___closed__4 = _init_l_Lean_IR_EmitC_emitStructDefn___closed__4(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructDefn___closed__4); +l_Lean_IR_EmitC_emitStructDefn___closed__5 = _init_l_Lean_IR_EmitC_emitStructDefn___closed__5(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructDefn___closed__5); +l_Lean_IR_EmitC_emitStructDefn___closed__6 = _init_l_Lean_IR_EmitC_emitStructDefn___closed__6(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructDefn___closed__6); +l_Lean_IR_EmitC_emitStructDefn___closed__7 = _init_l_Lean_IR_EmitC_emitStructDefn___closed__7(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructDefn___closed__7); +l_Lean_IR_EmitC_emitStructDefn___closed__8 = _init_l_Lean_IR_EmitC_emitStructDefn___closed__8(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructDefn___closed__8); +l_Lean_IR_EmitC_emitStructDefn___closed__9 = _init_l_Lean_IR_EmitC_emitStructDefn___closed__9(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructDefn___closed__9); +l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitUnionSwitch_spec__0___redArg___closed__0 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitUnionSwitch_spec__0___redArg___closed__0(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitUnionSwitch_spec__0___redArg___closed__0); +l_Lean_IR_EmitC_emitUnionSwitch___closed__0 = _init_l_Lean_IR_EmitC_emitUnionSwitch___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitC_emitUnionSwitch___closed__0); +l_Lean_IR_EmitC_emitUnionSwitchWithImpossible___closed__0 = _init_l_Lean_IR_EmitC_emitUnionSwitchWithImpossible___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitC_emitUnionSwitchWithImpossible___closed__0); +l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructIncDecFn_spec__0___redArg___closed__0 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructIncDecFn_spec__0___redArg___closed__0(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructIncDecFn_spec__0___redArg___closed__0); +l_Lean_IR_EmitC_emitStructIncDecFn___lam__0___closed__1 = _init_l_Lean_IR_EmitC_emitStructIncDecFn___lam__0___closed__1(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructIncDecFn___lam__0___closed__1); +l_Lean_IR_EmitC_emitStructIncDecFn___lam__0___closed__0 = _init_l_Lean_IR_EmitC_emitStructIncDecFn___lam__0___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructIncDecFn___lam__0___closed__0); +l_Lean_IR_EmitC_emitStructIncDecFn___closed__0 = _init_l_Lean_IR_EmitC_emitStructIncDecFn___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructIncDecFn___closed__0); +l_Lean_IR_EmitC_emitStructIncDecFn___closed__1 = _init_l_Lean_IR_EmitC_emitStructIncDecFn___closed__1(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructIncDecFn___closed__1); +l_Lean_IR_EmitC_emitStructIncDecFn___closed__2 = _init_l_Lean_IR_EmitC_emitStructIncDecFn___closed__2(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructIncDecFn___closed__2); +l_Lean_IR_EmitC_emitStructIncDecFn___closed__3 = _init_l_Lean_IR_EmitC_emitStructIncDecFn___closed__3(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructIncDecFn___closed__3); +l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__0 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__0(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__0); +l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__1 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__1(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__1); +l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___lam__0___closed__0 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___lam__0___closed__0(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___lam__0___closed__0); +l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__2 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__2(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__2); +l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__3 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__3(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__3); +l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__4 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__4(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructReshapeFn_spec__0___redArg___closed__4); +l_Lean_IR_EmitC_emitStructReshapeFn___lam__1___closed__0 = _init_l_Lean_IR_EmitC_emitStructReshapeFn___lam__1___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructReshapeFn___lam__1___closed__0); +l_Lean_IR_EmitC_emitStructReshapeFn___lam__1___closed__1 = _init_l_Lean_IR_EmitC_emitStructReshapeFn___lam__1___closed__1(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructReshapeFn___lam__1___closed__1); +l_Lean_IR_EmitC_emitStructReshapeFn___closed__0 = _init_l_Lean_IR_EmitC_emitStructReshapeFn___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructReshapeFn___closed__0); +l_Lean_IR_EmitC_emitStructReshapeFn___closed__1 = _init_l_Lean_IR_EmitC_emitStructReshapeFn___closed__1(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructReshapeFn___closed__1); +l_Lean_IR_EmitC_emitStructReshapeFn___closed__2 = _init_l_Lean_IR_EmitC_emitStructReshapeFn___closed__2(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructReshapeFn___closed__2); +l_Lean_IR_EmitC_emitStructReshapeFn___closed__3 = _init_l_Lean_IR_EmitC_emitStructReshapeFn___closed__3(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructReshapeFn___closed__3); +l_Lean_IR_EmitC_emitStructReshapeFn___closed__4 = _init_l_Lean_IR_EmitC_emitStructReshapeFn___closed__4(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructReshapeFn___closed__4); +l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructBox_spec__0___redArg___closed__0 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructBox_spec__0___redArg___closed__0(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructBox_spec__0___redArg___closed__0); +l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructBox_spec__0___redArg___closed__1 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructBox_spec__0___redArg___closed__1(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructBox_spec__0___redArg___closed__1); +l_Lean_IR_EmitC_emitStructBox___closed__0 = _init_l_Lean_IR_EmitC_emitStructBox___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructBox___closed__0); +l_Lean_IR_EmitC_emitStructBox___closed__1 = _init_l_Lean_IR_EmitC_emitStructBox___closed__1(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructBox___closed__1); +l_Lean_IR_EmitC_emitStructBox___closed__2 = _init_l_Lean_IR_EmitC_emitStructBox___closed__2(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructBox___closed__2); +l_Lean_IR_EmitC_emitStructBox___closed__3 = _init_l_Lean_IR_EmitC_emitStructBox___closed__3(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructBox___closed__3); +l_Lean_IR_EmitC_emitStructBox___closed__4 = _init_l_Lean_IR_EmitC_emitStructBox___closed__4(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructBox___closed__4); +l_Lean_IR_EmitC_emitStructBox___closed__5 = _init_l_Lean_IR_EmitC_emitStructBox___closed__5(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructBox___closed__5); +l_Lean_IR_EmitC_emitStructBox___closed__6 = _init_l_Lean_IR_EmitC_emitStructBox___closed__6(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructBox___closed__6); +l_Lean_IR_EmitC_emitStructBoxFn___closed__0 = _init_l_Lean_IR_EmitC_emitStructBoxFn___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructBoxFn___closed__0); +l_Lean_IR_EmitC_emitStructBoxFn___closed__1 = _init_l_Lean_IR_EmitC_emitStructBoxFn___closed__1(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructBoxFn___closed__1); +l_Lean_IR_EmitC_emitStructBoxFn___closed__2 = _init_l_Lean_IR_EmitC_emitStructBoxFn___closed__2(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructBoxFn___closed__2); +l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructUnboxFn_spec__0___redArg___closed__0 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructUnboxFn_spec__0___redArg___closed__0(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructUnboxFn_spec__0___redArg___closed__0); +l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructUnboxFn_spec__0___redArg___closed__1 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructUnboxFn_spec__0___redArg___closed__1(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_IR_EmitC_emitStructUnboxFn_spec__0___redArg___closed__1); +l_Lean_IR_EmitC_emitStructUnboxFn___lam__0___closed__0 = _init_l_Lean_IR_EmitC_emitStructUnboxFn___lam__0___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructUnboxFn___lam__0___closed__0); +l_Lean_IR_EmitC_emitStructUnboxFn___closed__0 = _init_l_Lean_IR_EmitC_emitStructUnboxFn___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructUnboxFn___closed__0); +l_Lean_IR_EmitC_emitStructUnboxFn___closed__1 = _init_l_Lean_IR_EmitC_emitStructUnboxFn___closed__1(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructUnboxFn___closed__1); +l_Lean_IR_EmitC_emitStructUnboxFn___closed__2 = _init_l_Lean_IR_EmitC_emitStructUnboxFn___closed__2(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructUnboxFn___closed__2); +l_Lean_IR_EmitC_emitStructUnboxFn___closed__3 = _init_l_Lean_IR_EmitC_emitStructUnboxFn___closed__3(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructUnboxFn___closed__3); +l_Lean_IR_EmitC_emitStructUnboxFn___closed__4 = _init_l_Lean_IR_EmitC_emitStructUnboxFn___closed__4(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructUnboxFn___closed__4); +l_Lean_IR_EmitC_emitStructUnboxFn___closed__5 = _init_l_Lean_IR_EmitC_emitStructUnboxFn___closed__5(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructUnboxFn___closed__5); +l_Lean_IR_EmitC_emitStructDecls___redArg___closed__0 = _init_l_Lean_IR_EmitC_emitStructDecls___redArg___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitC_emitStructDecls___redArg___closed__0); l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__0___redArg___closed__0 = _init_l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__0___redArg___closed__0(); lean_mark_persistent(l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__0___redArg___closed__0); l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__0___redArg___closed__1 = _init_l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__0___redArg___closed__1(); lean_mark_persistent(l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__0___redArg___closed__1); -l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__0___redArg___closed__2 = _init_l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__0___redArg___closed__2(); -lean_mark_persistent(l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__0___redArg___closed__2); l_Lean_IR_EmitC_emitDeclAux___closed__0 = _init_l_Lean_IR_EmitC_emitDeclAux___closed__0(); lean_mark_persistent(l_Lean_IR_EmitC_emitDeclAux___closed__0); l_Lean_IR_EmitC_emitDeclAux___closed__1 = _init_l_Lean_IR_EmitC_emitDeclAux___closed__1(); lean_mark_persistent(l_Lean_IR_EmitC_emitDeclAux___closed__1); l_Lean_IR_EmitC_emitDeclAux___closed__2 = _init_l_Lean_IR_EmitC_emitDeclAux___closed__2(); lean_mark_persistent(l_Lean_IR_EmitC_emitDeclAux___closed__2); -l_Lean_IR_EmitC_emitDeclAux___closed__3 = _init_l_Lean_IR_EmitC_emitDeclAux___closed__3(); -lean_mark_persistent(l_Lean_IR_EmitC_emitDeclAux___closed__3); l_Lean_IR_EmitC_emitDecl___closed__0 = _init_l_Lean_IR_EmitC_emitDecl___closed__0(); lean_mark_persistent(l_Lean_IR_EmitC_emitDecl___closed__0); l_Lean_IR_EmitC_emitMarkPersistent___closed__0 = _init_l_Lean_IR_EmitC_emitMarkPersistent___closed__0(); @@ -15298,6 +24081,8 @@ l_Lean_IR_EmitC_emitDeclInit___closed__3 = _init_l_Lean_IR_EmitC_emitDeclInit___ lean_mark_persistent(l_Lean_IR_EmitC_emitDeclInit___closed__3); l_Lean_IR_EmitC_emitDeclInit___closed__4 = _init_l_Lean_IR_EmitC_emitDeclInit___closed__4(); lean_mark_persistent(l_Lean_IR_EmitC_emitDeclInit___closed__4); +l_Lean_IR_EmitC_emitDeclInit___closed__5 = _init_l_Lean_IR_EmitC_emitDeclInit___closed__5(); +lean_mark_persistent(l_Lean_IR_EmitC_emitDeclInit___closed__5); l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitInitFn_spec__2___closed__0 = _init_l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitInitFn_spec__2___closed__0(); lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitInitFn_spec__2___closed__0); l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_EmitC_emitInitFn_spec__0___redArg___closed__0 = _init_l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_EmitC_emitInitFn_spec__0___redArg___closed__0(); @@ -15326,6 +24111,8 @@ l_Lean_IR_EmitC_emitInitFn___closed__9 = _init_l_Lean_IR_EmitC_emitInitFn___clos lean_mark_persistent(l_Lean_IR_EmitC_emitInitFn___closed__9); l_Lean_IR_EmitC_emitInitFn___closed__10 = _init_l_Lean_IR_EmitC_emitInitFn___closed__10(); lean_mark_persistent(l_Lean_IR_EmitC_emitInitFn___closed__10); +l_Lean_IR_EmitC_main___closed__0 = _init_l_Lean_IR_EmitC_main___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitC_main___closed__0); l_Lean_IR_emitC___closed__0 = _init_l_Lean_IR_emitC___closed__0(); lean_mark_persistent(l_Lean_IR_emitC___closed__0); l_Lean_IR_emitC___closed__1 = _init_l_Lean_IR_emitC___closed__1(); diff --git a/stage0/stdlib/Lean/Compiler/IR/EmitLLVM.c b/stage0/stdlib/Lean/Compiler/IR/EmitLLVM.c index fd5586a8d8e6..a49fdc2c52cb 100644 --- a/stage0/stdlib/Lean/Compiler/IR/EmitLLVM.c +++ b/stage0/stdlib/Lean/Compiler/IR/EmitLLVM.c @@ -27,6 +27,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanIOMkWorld___boxed(lean_objec LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitArgVal___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitLLVM_emitJmp___closed__1; size_t lean_llvm_function_type(size_t, size_t, lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitProj___redArg(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_Format_pretty(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitFnDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_llvm_get_first_function(size_t, size_t); @@ -77,7 +78,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitDeclInit___lam__0(size_t, size_t LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanIOResultShowError___redArg(size_t, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanMainFn(size_t, size_t, lean_object*, size_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitLLVM_emitApp___closed__2; -LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitProj___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitProj___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_IR_EmitLLVM_emitReset___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_constIntUnsigned(size_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitLLVM_emitInitFn___closed__2; @@ -340,7 +341,7 @@ lean_object* lean_st_ref_take(lean_object*); static lean_object* l_Lean_IR_EmitLLVM_callLeanInitialize___redArg___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_getDecl___redArg(lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitLLVM_emitDeclInit___closed__3; -LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitProj(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitProj(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_llvm_add_attribute_at_index(size_t, size_t, uint64_t, size_t); uint8_t l_Lean_IR_usesModuleFrom(lean_object*, lean_object*); extern lean_object* l_Lean_IR_instInhabitedParam_default; @@ -785,6 +786,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callIODeclInitFn___redArg___boxed(le lean_object* l_List_reverse___redArg(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_ShouldForwardControlFlow_yes_elim___redArg(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_RefcountKind_dec_elim___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitProj___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00Lean_IR_EmitLLVM_emitFnDecls_spec__3___redArg(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitLLVM_addVarToState___redArg___closed__0; static lean_object* l_Lean_IR_EmitLLVM_emitMainFn___closed__15; @@ -8658,7 +8660,7 @@ return x_1; LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitFnDeclAux(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7) { _start: { -size_t x_9; lean_object* x_10; lean_object* x_15; size_t x_16; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; size_t x_24; lean_object* x_25; uint8_t x_34; +size_t x_9; lean_object* x_10; size_t x_15; lean_object* x_16; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; size_t x_24; lean_object* x_25; uint8_t x_34; x_20 = l_Lean_IR_EmitLLVM_getEnv___redArg(x_7); x_21 = lean_ctor_get(x_20, 0); lean_inc(x_21); @@ -9288,8 +9290,8 @@ return x_13; { uint64_t x_17; lean_object* x_18; x_17 = l_Lean_IR_EmitLLVM_emitFnDeclAux___closed__0; -x_18 = lean_llvm_set_dll_storage_class(x_1, x_16, x_17); -x_9 = x_16; +x_18 = lean_llvm_set_dll_storage_class(x_1, x_15, x_17); +x_9 = x_15; x_10 = lean_box(0); goto block_14; } @@ -9303,8 +9305,8 @@ if (x_26 == 0) lean_dec(x_22); if (x_5 == 0) { -x_15 = lean_box(0); -x_16 = x_24; +x_15 = x_24; +x_16 = lean_box(0); goto block_19; } else @@ -9317,8 +9319,8 @@ goto block_14; } else { -x_15 = lean_box(0); -x_16 = x_24; +x_15 = x_24; +x_16 = lean_box(0); goto block_19; } } @@ -23970,7 +23972,7 @@ lean_dec_ref(x_6); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitProj(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_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitProj___redArg(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) { _start: { lean_object* x_9; lean_object* x_10; @@ -24275,7 +24277,31 @@ return x_68; } } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitProj___boxed(lean_object* x_1, lean_object* x_2, lean_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_IR_EmitLLVM_emitProj(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_IR_EmitLLVM_emitProj___redArg(x_1, x_2, x_3, x_5, x_6, x_7, x_8); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitProj___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Lean_IR_EmitLLVM_emitProj(x_10, x_11, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitProj___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { size_t x_9; size_t x_10; lean_object* x_11; @@ -24283,7 +24309,7 @@ x_9 = lean_unbox_usize(x_1); lean_dec(x_1); x_10 = lean_unbox_usize(x_2); lean_dec(x_2); -x_11 = l_Lean_IR_EmitLLVM_emitProj(x_9, x_10, x_3, x_4, x_5, x_6, x_7); +x_11 = l_Lean_IR_EmitLLVM_emitProj___redArg(x_9, x_10, x_3, x_4, x_5, x_6, x_7); lean_dec_ref(x_7); lean_dec(x_6); lean_dec(x_4); @@ -31503,12 +31529,12 @@ case 3: { lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_dec(x_4); -x_20 = lean_ctor_get(x_5, 0); +x_20 = lean_ctor_get(x_5, 1); lean_inc(x_20); -x_21 = lean_ctor_get(x_5, 1); +x_21 = lean_ctor_get(x_5, 2); lean_inc(x_21); lean_dec_ref(x_5); -x_22 = l_Lean_IR_EmitLLVM_emitProj(x_1, x_2, x_3, x_20, x_21, x_6, x_7); +x_22 = l_Lean_IR_EmitLLVM_emitProj___redArg(x_1, x_2, x_3, x_20, x_21, x_6, x_7); lean_dec_ref(x_7); lean_dec(x_6); lean_dec(x_20); @@ -31518,9 +31544,9 @@ case 4: { lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_dec(x_4); -x_23 = lean_ctor_get(x_5, 0); +x_23 = lean_ctor_get(x_5, 1); lean_inc(x_23); -x_24 = lean_ctor_get(x_5, 1); +x_24 = lean_ctor_get(x_5, 2); lean_inc(x_24); lean_dec_ref(x_5); x_25 = l_Lean_IR_EmitLLVM_emitUProj(x_1, x_2, x_3, x_23, x_24, x_6, x_7); @@ -31532,11 +31558,11 @@ return x_25; case 5: { lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_26 = lean_ctor_get(x_5, 0); +x_26 = lean_ctor_get(x_5, 1); lean_inc(x_26); -x_27 = lean_ctor_get(x_5, 1); +x_27 = lean_ctor_get(x_5, 2); lean_inc(x_27); -x_28 = lean_ctor_get(x_5, 2); +x_28 = lean_ctor_get(x_5, 3); lean_inc(x_28); lean_dec_ref(x_5); x_29 = l_Lean_IR_EmitLLVM_emitSProj(x_1, x_2, x_3, x_4, x_26, x_27, x_28, x_6, x_7); @@ -36545,11 +36571,11 @@ case 4: 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_3, 0); lean_inc(x_36); -x_37 = lean_ctor_get(x_3, 1); +x_37 = lean_ctor_get(x_3, 2); lean_inc(x_37); -x_38 = lean_ctor_get(x_3, 2); +x_38 = lean_ctor_get(x_3, 3); lean_inc(x_38); -x_39 = lean_ctor_get(x_3, 3); +x_39 = lean_ctor_get(x_3, 4); lean_inc(x_39); lean_dec_ref(x_3); x_40 = l_Lean_IR_EmitLLVM_emitUSet(x_1, x_2, x_36, x_37, x_38, x_4, x_5); @@ -36588,15 +36614,15 @@ case 5: lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; x_43 = lean_ctor_get(x_3, 0); lean_inc(x_43); -x_44 = lean_ctor_get(x_3, 1); +x_44 = lean_ctor_get(x_3, 2); lean_inc(x_44); -x_45 = lean_ctor_get(x_3, 2); +x_45 = lean_ctor_get(x_3, 3); lean_inc(x_45); -x_46 = lean_ctor_get(x_3, 3); +x_46 = lean_ctor_get(x_3, 4); lean_inc(x_46); -x_47 = lean_ctor_get(x_3, 4); +x_47 = lean_ctor_get(x_3, 5); lean_inc(x_47); -x_48 = lean_ctor_get(x_3, 5); +x_48 = lean_ctor_get(x_3, 6); lean_inc(x_48); lean_dec_ref(x_3); x_49 = l_Lean_IR_EmitLLVM_emitSSet(x_1, x_2, x_43, x_44, x_45, x_46, x_47, x_4, x_5); @@ -38351,7 +38377,7 @@ return x_30; } else { -lean_object* x_36; lean_object* x_37; size_t x_38; uint8_t x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_64; lean_object* x_65; size_t x_66; uint8_t x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_80; lean_object* x_81; lean_object* x_82; size_t x_83; uint8_t x_84; lean_object* x_109; uint8_t x_139; +lean_object* x_36; lean_object* x_37; size_t x_38; uint8_t x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_64; size_t x_65; lean_object* x_66; uint8_t x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_80; size_t x_81; lean_object* x_82; lean_object* x_83; uint8_t x_84; lean_object* x_109; uint8_t x_139; lean_free_object(x_30); x_36 = lean_ctor_get(x_32, 0); lean_inc(x_36); @@ -38490,10 +38516,10 @@ goto block_15; block_79: { size_t x_72; size_t x_73; uint8_t x_74; -x_72 = lean_llvm_function_type(x_1, x_66, x_68, x_23); +x_72 = lean_llvm_function_type(x_1, x_65, x_68, x_23); lean_dec_ref(x_68); -x_73 = l_Lean_IR_LLVM_getOrAddFunction(x_1, x_2, x_65, x_72); -lean_dec_ref(x_65); +x_73 = l_Lean_IR_LLVM_getOrAddFunction(x_1, x_2, x_66, x_72); +lean_dec_ref(x_66); x_74 = lean_nat_dec_eq(x_64, x_37); if (x_74 == 0) { @@ -38571,8 +38597,8 @@ lean_free_object(x_87); x_93 = lean_ctor_get(x_89, 0); lean_inc(x_93); lean_dec_ref(x_89); -x_65 = x_82; -x_66 = x_83; +x_65 = x_81; +x_66 = x_82; x_67 = x_84; x_68 = x_93; x_69 = x_5; @@ -38621,8 +38647,8 @@ lean_object* x_99; x_99 = lean_ctor_get(x_94, 0); lean_inc(x_99); lean_dec_ref(x_94); -x_65 = x_82; -x_66 = x_83; +x_65 = x_81; +x_66 = x_82; x_67 = x_84; x_68 = x_99; x_69 = x_5; @@ -38667,8 +38693,8 @@ x_104 = lean_llvm_pointer_type(x_1, x_103); x_105 = l_Lean_IR_EmitLLVM_callLeanBox___redArg___closed__1; x_106 = lean_box_usize(x_104); x_107 = lean_array_push(x_105, x_106); -x_65 = x_82; -x_66 = x_83; +x_65 = x_81; +x_66 = x_82; x_67 = x_84; x_68 = x_107; x_69 = x_5; @@ -38736,9 +38762,9 @@ lean_dec(x_22); x_120 = lean_unbox_usize(x_116); lean_dec(x_116); x_80 = x_117; -x_81 = lean_box(0); +x_81 = x_120; x_82 = x_109; -x_83 = x_120; +x_83 = lean_box(0); x_84 = x_119; goto block_108; } @@ -38750,9 +38776,9 @@ lean_dec(x_22); x_122 = lean_unbox_usize(x_116); lean_dec(x_116); x_80 = x_117; -x_81 = lean_box(0); +x_81 = x_122; x_82 = x_109; -x_83 = x_122; +x_83 = lean_box(0); x_84 = x_121; goto block_108; } @@ -38809,9 +38835,9 @@ lean_dec(x_22); x_132 = lean_unbox_usize(x_128); lean_dec(x_128); x_80 = x_129; -x_81 = lean_box(0); +x_81 = x_132; x_82 = x_109; -x_83 = x_132; +x_83 = lean_box(0); x_84 = x_131; goto block_108; } @@ -38823,9 +38849,9 @@ lean_dec(x_22); x_134 = lean_unbox_usize(x_128); lean_dec(x_128); x_80 = x_129; -x_81 = lean_box(0); +x_81 = x_134; x_82 = x_109; -x_83 = x_134; +x_83 = lean_box(0); x_84 = x_133; goto block_108; } @@ -38898,7 +38924,7 @@ return x_146; } else { -lean_object* x_147; lean_object* x_148; size_t x_149; uint8_t x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_170; lean_object* x_171; size_t x_172; uint8_t x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_186; lean_object* x_187; lean_object* x_188; size_t x_189; uint8_t x_190; lean_object* x_210; uint8_t x_229; +lean_object* x_147; lean_object* x_148; size_t x_149; uint8_t x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_170; size_t x_171; lean_object* x_172; uint8_t x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_186; size_t x_187; lean_object* x_188; lean_object* x_189; uint8_t x_190; lean_object* x_210; uint8_t x_229; x_147 = lean_ctor_get(x_142, 0); lean_inc(x_147); lean_dec_ref(x_142); @@ -38998,10 +39024,10 @@ goto block_15; block_185: { size_t x_178; size_t x_179; uint8_t x_180; -x_178 = lean_llvm_function_type(x_1, x_172, x_174, x_23); +x_178 = lean_llvm_function_type(x_1, x_171, x_174, x_23); lean_dec_ref(x_174); -x_179 = l_Lean_IR_LLVM_getOrAddFunction(x_1, x_2, x_171, x_178); -lean_dec_ref(x_171); +x_179 = l_Lean_IR_LLVM_getOrAddFunction(x_1, x_2, x_172, x_178); +lean_dec_ref(x_172); x_180 = lean_nat_dec_eq(x_170, x_148); if (x_180 == 0) { @@ -39089,8 +39115,8 @@ lean_dec(x_195); x_200 = lean_ctor_get(x_194, 0); lean_inc(x_200); lean_dec_ref(x_194); -x_171 = x_188; -x_172 = x_189; +x_171 = x_187; +x_172 = x_188; x_173 = x_190; x_174 = x_200; x_175 = x_5; @@ -39135,8 +39161,8 @@ x_205 = lean_llvm_pointer_type(x_1, x_204); x_206 = l_Lean_IR_EmitLLVM_callLeanBox___redArg___closed__1; x_207 = lean_box_usize(x_205); x_208 = lean_array_push(x_206, x_207); -x_171 = x_188; -x_172 = x_189; +x_171 = x_187; +x_172 = x_188; x_173 = x_190; x_174 = x_208; x_175 = x_5; @@ -39214,9 +39240,9 @@ lean_dec(x_22); x_222 = lean_unbox_usize(x_218); lean_dec(x_218); x_186 = x_219; -x_187 = lean_box(0); +x_187 = x_222; x_188 = x_210; -x_189 = x_222; +x_189 = lean_box(0); x_190 = x_221; goto block_209; } @@ -39228,9 +39254,9 @@ lean_dec(x_22); x_224 = lean_unbox_usize(x_218); lean_dec(x_218); x_186 = x_219; -x_187 = lean_box(0); +x_187 = x_224; x_188 = x_210; -x_189 = x_224; +x_189 = lean_box(0); x_190 = x_223; goto block_209; } @@ -39363,7 +39389,7 @@ return x_247; } else { -lean_object* x_248; lean_object* x_249; size_t x_250; uint8_t x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_271; lean_object* x_272; size_t x_273; uint8_t x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_287; lean_object* x_288; lean_object* x_289; size_t x_290; uint8_t x_291; lean_object* x_311; uint8_t x_330; +lean_object* x_248; lean_object* x_249; size_t x_250; uint8_t x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_271; size_t x_272; lean_object* x_273; uint8_t x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_287; size_t x_288; lean_object* x_289; lean_object* x_290; uint8_t x_291; lean_object* x_311; uint8_t x_330; lean_dec(x_243); x_248 = lean_ctor_get(x_242, 0); lean_inc(x_248); @@ -39464,10 +39490,10 @@ goto block_15; block_286: { size_t x_279; size_t x_280; uint8_t x_281; -x_279 = lean_llvm_function_type(x_1, x_273, x_275, x_23); +x_279 = lean_llvm_function_type(x_1, x_272, x_275, x_23); lean_dec_ref(x_275); -x_280 = l_Lean_IR_LLVM_getOrAddFunction(x_1, x_2, x_272, x_279); -lean_dec_ref(x_272); +x_280 = l_Lean_IR_LLVM_getOrAddFunction(x_1, x_2, x_273, x_279); +lean_dec_ref(x_273); x_281 = lean_nat_dec_eq(x_271, x_249); if (x_281 == 0) { @@ -39555,8 +39581,8 @@ lean_dec(x_296); x_301 = lean_ctor_get(x_295, 0); lean_inc(x_301); lean_dec_ref(x_295); -x_272 = x_289; -x_273 = x_290; +x_272 = x_288; +x_273 = x_289; x_274 = x_291; x_275 = x_301; x_276 = x_5; @@ -39601,8 +39627,8 @@ x_306 = lean_llvm_pointer_type(x_1, x_305); x_307 = l_Lean_IR_EmitLLVM_callLeanBox___redArg___closed__1; x_308 = lean_box_usize(x_306); x_309 = lean_array_push(x_307, x_308); -x_272 = x_289; -x_273 = x_290; +x_272 = x_288; +x_273 = x_289; x_274 = x_291; x_275 = x_309; x_276 = x_5; @@ -39680,9 +39706,9 @@ lean_dec(x_22); x_323 = lean_unbox_usize(x_319); lean_dec(x_319); x_287 = x_320; -x_288 = lean_box(0); +x_288 = x_323; x_289 = x_311; -x_290 = x_323; +x_290 = lean_box(0); x_291 = x_322; goto block_310; } @@ -39694,9 +39720,9 @@ lean_dec(x_22); x_325 = lean_unbox_usize(x_319); lean_dec(x_319); x_287 = x_320; -x_288 = lean_box(0); +x_288 = x_325; x_289 = x_311; -x_290 = x_325; +x_290 = lean_box(0); x_291 = x_324; goto block_310; } @@ -39898,7 +39924,7 @@ return x_361; } else { -lean_object* x_362; lean_object* x_363; size_t x_364; uint8_t x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_385; lean_object* x_386; size_t x_387; uint8_t x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_401; lean_object* x_402; lean_object* x_403; size_t x_404; uint8_t x_405; lean_object* x_425; uint8_t x_444; +lean_object* x_362; lean_object* x_363; size_t x_364; uint8_t x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_385; size_t x_386; lean_object* x_387; uint8_t x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_401; size_t x_402; lean_object* x_403; lean_object* x_404; uint8_t x_405; lean_object* x_425; uint8_t x_444; lean_dec(x_357); x_362 = lean_ctor_get(x_356, 0); lean_inc(x_362); @@ -39999,10 +40025,10 @@ goto block_15; block_400: { size_t x_393; size_t x_394; uint8_t x_395; -x_393 = lean_llvm_function_type(x_1, x_387, x_389, x_343); +x_393 = lean_llvm_function_type(x_1, x_386, x_389, x_343); lean_dec_ref(x_389); -x_394 = l_Lean_IR_LLVM_getOrAddFunction(x_1, x_2, x_386, x_393); -lean_dec_ref(x_386); +x_394 = l_Lean_IR_LLVM_getOrAddFunction(x_1, x_2, x_387, x_393); +lean_dec_ref(x_387); x_395 = lean_nat_dec_eq(x_385, x_363); if (x_395 == 0) { @@ -40090,8 +40116,8 @@ lean_dec(x_410); x_415 = lean_ctor_get(x_409, 0); lean_inc(x_415); lean_dec_ref(x_409); -x_386 = x_403; -x_387 = x_404; +x_386 = x_402; +x_387 = x_403; x_388 = x_405; x_389 = x_415; x_390 = x_5; @@ -40136,8 +40162,8 @@ x_420 = lean_llvm_pointer_type(x_1, x_419); x_421 = l_Lean_IR_EmitLLVM_callLeanBox___redArg___closed__1; x_422 = lean_box_usize(x_420); x_423 = lean_array_push(x_421, x_422); -x_386 = x_403; -x_387 = x_404; +x_386 = x_402; +x_387 = x_403; x_388 = x_405; x_389 = x_423; x_390 = x_5; @@ -40215,9 +40241,9 @@ lean_dec(x_342); x_437 = lean_unbox_usize(x_433); lean_dec(x_433); x_401 = x_434; -x_402 = lean_box(0); +x_402 = x_437; x_403 = x_425; -x_404 = x_437; +x_404 = lean_box(0); x_405 = x_436; goto block_424; } @@ -40229,9 +40255,9 @@ lean_dec(x_342); x_439 = lean_unbox_usize(x_433); lean_dec(x_433); x_401 = x_434; -x_402 = lean_box(0); +x_402 = x_439; x_403 = x_425; -x_404 = x_439; +x_404 = lean_box(0); x_405 = x_438; goto block_424; } @@ -61365,7 +61391,7 @@ return x_15; LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitMainFn(size_t x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_7; size_t x_8; lean_object* x_9; size_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; size_t x_15; lean_object* x_16; size_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_75; lean_object* x_76; lean_object* x_77; +size_t x_7; lean_object* x_8; lean_object* x_9; size_t x_10; lean_object* x_11; size_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; size_t x_18; lean_object* x_19; lean_object* x_75; lean_object* x_76; lean_object* x_77; x_75 = l_Lean_IR_EmitLLVM_toCName___redArg___closed__0; x_76 = l_Lean_IR_EmitLLVM_toCName___redArg___closed__1; x_77 = l_Lean_IR_EmitLLVM_getDecl___redArg(x_76, x_5); @@ -61429,7 +61455,7 @@ if (lean_is_exclusive(x_78)) { } if (lean_obj_tag(x_85) == 0) { -lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; size_t x_92; size_t x_93; lean_object* x_94; size_t x_95; lean_object* x_96; size_t x_97; size_t x_98; size_t x_99; lean_object* x_100; size_t x_101; lean_object* x_102; uint8_t x_103; size_t x_104; uint8_t x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; uint8_t x_344; lean_object* x_579; uint8_t x_580; +lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; size_t x_93; size_t x_94; lean_object* x_95; size_t x_96; size_t x_97; size_t x_98; lean_object* x_99; size_t x_100; uint8_t x_101; lean_object* x_102; size_t x_103; uint8_t x_104; size_t x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; uint8_t x_344; lean_object* x_579; uint8_t x_580; x_87 = lean_ctor_get(x_85, 1); lean_inc_ref(x_87); lean_dec_ref(x_85); @@ -61469,10 +61495,10 @@ lean_dec_ref(x_111); lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec(x_94); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); +lean_dec(x_91); lean_dec_ref(x_89); return x_110; } @@ -61496,10 +61522,10 @@ uint8_t x_115; lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec(x_94); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); +lean_dec(x_91); lean_dec_ref(x_89); x_115 = !lean_is_exclusive(x_114); if (x_115 == 0) @@ -61532,7 +61558,7 @@ lean_dec_ref(x_119); x_121 = lean_ctor_get(x_120, 0); lean_inc(x_121); lean_dec(x_120); -x_122 = l_Lean_IR_EmitLLVM_constInt8___redArg(x_1, x_96); +x_122 = l_Lean_IR_EmitLLVM_constInt8___redArg(x_1, x_92); x_123 = lean_ctor_get(x_122, 0); lean_inc(x_123); lean_dec_ref(x_122); @@ -61557,7 +61583,7 @@ x_132 = lean_box(0); lean_inc_ref(x_102); x_133 = l_Lean_PersistentEnvExtension_getState___redArg(x_131, x_128, x_102, x_130, x_132); lean_dec(x_130); -x_134 = l_Lean_Name_toString(x_127, x_105); +x_134 = l_Lean_Name_toString(x_127, x_104); x_135 = l_Lean_IR_EmitLLVM_emitMainFn___closed__1; x_136 = lean_string_append(x_134, x_135); x_137 = lean_unbox_usize(x_124); @@ -61580,10 +61606,10 @@ uint8_t x_142; lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec(x_94); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); +lean_dec(x_91); lean_dec_ref(x_89); x_142 = !lean_is_exclusive(x_141); if (x_142 == 0) @@ -61611,7 +61637,7 @@ lean_inc(x_145); lean_dec_ref(x_141); x_146 = lean_unbox_usize(x_145); lean_dec(x_145); -x_147 = lean_llvm_build_store(x_1, x_3, x_146, x_101); +x_147 = lean_llvm_build_store(x_1, x_3, x_146, x_105); x_148 = l_LLVM_constTrue(x_1); x_149 = l_Lean_IR_EmitLLVM_callLeanSetPanicMessages___redArg(x_1, x_3, x_148, x_107); if (lean_obj_tag(x_149) == 0) @@ -61625,10 +61651,10 @@ lean_dec_ref(x_150); lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec(x_94); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); +lean_dec(x_91); lean_dec_ref(x_89); return x_149; } @@ -61649,10 +61675,10 @@ lean_dec_ref(x_152); lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec(x_94); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); +lean_dec(x_91); lean_dec_ref(x_89); return x_151; } @@ -61662,7 +61688,7 @@ lean_object* x_153; size_t x_154; lean_object* x_155; lean_object* x_156; lean_dec_ref(x_152); lean_dec_ref(x_151); x_153 = l_Lean_IR_EmitLLVM_emitMainFn___closed__2; -x_154 = lean_llvm_build_load2(x_1, x_3, x_104, x_101, x_153); +x_154 = lean_llvm_build_load2(x_1, x_3, x_103, x_105, x_153); x_155 = l_Lean_IR_EmitLLVM_emitMainFn___closed__3; x_156 = l_Lean_IR_EmitLLVM_callLeanIOResultIsOk___redArg(x_1, x_3, x_154, x_155, x_107); if (lean_obj_tag(x_156) == 0) @@ -61679,10 +61705,10 @@ uint8_t x_159; lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec(x_94); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); +lean_dec(x_91); lean_dec_ref(x_89); x_159 = !lean_is_exclusive(x_158); if (x_159 == 0) @@ -61710,24 +61736,24 @@ lean_inc(x_162); lean_dec_ref(x_158); x_163 = lean_box_usize(x_1); x_164 = lean_box_usize(x_154); -x_165 = lean_box_usize(x_92); -x_166 = lean_box_usize(x_98); -x_167 = lean_box_usize(x_93); -x_168 = lean_box_usize(x_95); +x_165 = lean_box_usize(x_96); +x_166 = lean_box_usize(x_94); +x_167 = lean_box_usize(x_100); +x_168 = lean_box_usize(x_98); x_169 = lean_box_usize(x_97); -lean_inc(x_96); -lean_inc(x_100); +lean_inc(x_92); +lean_inc(x_95); x_170 = lean_alloc_closure((void*)(l_Lean_IR_EmitLLVM_emitMainFn___lam__3___boxed), 17, 13); lean_closure_set(x_170, 0, x_163); lean_closure_set(x_170, 1, x_164); lean_closure_set(x_170, 2, x_90); -lean_closure_set(x_170, 3, x_94); +lean_closure_set(x_170, 3, x_91); lean_closure_set(x_170, 4, x_153); lean_closure_set(x_170, 5, x_165); -lean_closure_set(x_170, 6, x_100); +lean_closure_set(x_170, 6, x_95); lean_closure_set(x_170, 7, x_166); lean_closure_set(x_170, 8, x_167); -lean_closure_set(x_170, 9, x_96); +lean_closure_set(x_170, 9, x_92); lean_closure_set(x_170, 10, x_89); lean_closure_set(x_170, 11, x_168); lean_closure_set(x_170, 12, x_169); @@ -61748,9 +61774,9 @@ lean_dec_ref(x_174); lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); return x_173; } else @@ -61759,33 +61785,33 @@ lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_dec_ref(x_174); lean_dec_ref(x_173); x_175 = lean_box_usize(x_1); -x_176 = lean_box_usize(x_99); -x_177 = lean_box_usize(x_92); +x_176 = lean_box_usize(x_93); +x_177 = lean_box_usize(x_96); x_178 = lean_alloc_closure((void*)(l_Lean_IR_EmitLLVM_emitMainFn___lam__4___boxed), 10, 6); lean_closure_set(x_178, 0, x_175); lean_closure_set(x_178, 1, x_176); lean_closure_set(x_178, 2, x_177); lean_closure_set(x_178, 3, x_153); -lean_closure_set(x_178, 4, x_91); -lean_closure_set(x_178, 5, x_96); -x_179 = l_Lean_Environment_find_x3f(x_102, x_76, x_103); +lean_closure_set(x_178, 4, x_99); +lean_closure_set(x_178, 5, x_92); +x_179 = l_Lean_Environment_find_x3f(x_102, x_76, x_101); if (lean_obj_tag(x_179) == 0) { lean_object* x_180; lean_object* x_181; x_180 = l_Lean_IR_EmitLLVM_emitMainFn___closed__8; x_181 = l_panic___at___00Lean_IR_EmitLLVM_emitMainFn_spec__1(x_180); -x_7 = x_153; -x_8 = x_92; -x_9 = x_178; -x_10 = x_99; -x_11 = x_100; -x_12 = x_155; -x_13 = lean_box(0); -x_14 = x_153; -x_15 = x_101; -x_16 = x_107; -x_17 = x_104; -x_18 = x_106; +x_7 = x_93; +x_8 = x_178; +x_9 = x_153; +x_10 = x_96; +x_11 = x_95; +x_12 = x_103; +x_13 = x_153; +x_14 = x_107; +x_15 = lean_box(0); +x_16 = x_106; +x_17 = x_155; +x_18 = x_105; x_19 = x_181; goto block_74; } @@ -61795,18 +61821,18 @@ lean_object* x_182; x_182 = lean_ctor_get(x_179, 0); lean_inc(x_182); lean_dec_ref(x_179); -x_7 = x_153; -x_8 = x_92; -x_9 = x_178; -x_10 = x_99; -x_11 = x_100; -x_12 = x_155; -x_13 = lean_box(0); -x_14 = x_153; -x_15 = x_101; -x_16 = x_107; -x_17 = x_104; -x_18 = x_106; +x_7 = x_93; +x_8 = x_178; +x_9 = x_153; +x_10 = x_96; +x_11 = x_95; +x_12 = x_103; +x_13 = x_153; +x_14 = x_107; +x_15 = lean_box(0); +x_16 = x_106; +x_17 = x_155; +x_18 = x_105; x_19 = x_182; goto block_74; } @@ -61817,9 +61843,9 @@ else lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); return x_173; } } @@ -61836,10 +61862,10 @@ lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec(x_94); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); +lean_dec(x_91); lean_dec_ref(x_89); x_184 = lean_ctor_get(x_183, 0); lean_inc(x_184); @@ -61868,24 +61894,24 @@ lean_inc(x_188); lean_dec_ref(x_183); x_189 = lean_box_usize(x_1); x_190 = lean_box_usize(x_154); -x_191 = lean_box_usize(x_92); -x_192 = lean_box_usize(x_98); -x_193 = lean_box_usize(x_93); -x_194 = lean_box_usize(x_95); +x_191 = lean_box_usize(x_96); +x_192 = lean_box_usize(x_94); +x_193 = lean_box_usize(x_100); +x_194 = lean_box_usize(x_98); x_195 = lean_box_usize(x_97); -lean_inc(x_96); -lean_inc(x_100); +lean_inc(x_92); +lean_inc(x_95); x_196 = lean_alloc_closure((void*)(l_Lean_IR_EmitLLVM_emitMainFn___lam__3___boxed), 17, 13); lean_closure_set(x_196, 0, x_189); lean_closure_set(x_196, 1, x_190); lean_closure_set(x_196, 2, x_90); -lean_closure_set(x_196, 3, x_94); +lean_closure_set(x_196, 3, x_91); lean_closure_set(x_196, 4, x_153); lean_closure_set(x_196, 5, x_191); -lean_closure_set(x_196, 6, x_100); +lean_closure_set(x_196, 6, x_95); lean_closure_set(x_196, 7, x_192); lean_closure_set(x_196, 8, x_193); -lean_closure_set(x_196, 9, x_96); +lean_closure_set(x_196, 9, x_92); lean_closure_set(x_196, 10, x_89); lean_closure_set(x_196, 11, x_194); lean_closure_set(x_196, 12, x_195); @@ -61906,9 +61932,9 @@ lean_dec_ref(x_200); lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); return x_199; } else @@ -61917,33 +61943,33 @@ lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_dec_ref(x_200); lean_dec_ref(x_199); x_201 = lean_box_usize(x_1); -x_202 = lean_box_usize(x_99); -x_203 = lean_box_usize(x_92); +x_202 = lean_box_usize(x_93); +x_203 = lean_box_usize(x_96); x_204 = lean_alloc_closure((void*)(l_Lean_IR_EmitLLVM_emitMainFn___lam__4___boxed), 10, 6); lean_closure_set(x_204, 0, x_201); lean_closure_set(x_204, 1, x_202); lean_closure_set(x_204, 2, x_203); lean_closure_set(x_204, 3, x_153); -lean_closure_set(x_204, 4, x_91); -lean_closure_set(x_204, 5, x_96); -x_205 = l_Lean_Environment_find_x3f(x_102, x_76, x_103); +lean_closure_set(x_204, 4, x_99); +lean_closure_set(x_204, 5, x_92); +x_205 = l_Lean_Environment_find_x3f(x_102, x_76, x_101); if (lean_obj_tag(x_205) == 0) { lean_object* x_206; lean_object* x_207; x_206 = l_Lean_IR_EmitLLVM_emitMainFn___closed__8; x_207 = l_panic___at___00Lean_IR_EmitLLVM_emitMainFn_spec__1(x_206); -x_7 = x_153; -x_8 = x_92; -x_9 = x_204; -x_10 = x_99; -x_11 = x_100; -x_12 = x_155; -x_13 = lean_box(0); -x_14 = x_153; -x_15 = x_101; -x_16 = x_107; -x_17 = x_104; -x_18 = x_106; +x_7 = x_93; +x_8 = x_204; +x_9 = x_153; +x_10 = x_96; +x_11 = x_95; +x_12 = x_103; +x_13 = x_153; +x_14 = x_107; +x_15 = lean_box(0); +x_16 = x_106; +x_17 = x_155; +x_18 = x_105; x_19 = x_207; goto block_74; } @@ -61953,18 +61979,18 @@ lean_object* x_208; x_208 = lean_ctor_get(x_205, 0); lean_inc(x_208); lean_dec_ref(x_205); -x_7 = x_153; -x_8 = x_92; -x_9 = x_204; -x_10 = x_99; -x_11 = x_100; -x_12 = x_155; -x_13 = lean_box(0); -x_14 = x_153; -x_15 = x_101; -x_16 = x_107; -x_17 = x_104; -x_18 = x_106; +x_7 = x_93; +x_8 = x_204; +x_9 = x_153; +x_10 = x_96; +x_11 = x_95; +x_12 = x_103; +x_13 = x_153; +x_14 = x_107; +x_15 = lean_box(0); +x_16 = x_106; +x_17 = x_155; +x_18 = x_105; x_19 = x_208; goto block_74; } @@ -61975,9 +62001,9 @@ else lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); return x_199; } } @@ -61989,10 +62015,10 @@ uint8_t x_209; lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec(x_94); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); +lean_dec(x_91); lean_dec_ref(x_89); x_209 = !lean_is_exclusive(x_156); if (x_209 == 0) @@ -62017,10 +62043,10 @@ else lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec(x_94); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); +lean_dec(x_91); lean_dec_ref(x_89); return x_151; } @@ -62031,10 +62057,10 @@ else lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec(x_94); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); +lean_dec(x_91); lean_dec_ref(x_89); return x_149; } @@ -62052,10 +62078,10 @@ lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec(x_94); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); +lean_dec(x_91); lean_dec_ref(x_89); x_213 = lean_ctor_get(x_212, 0); lean_inc(x_213); @@ -62084,7 +62110,7 @@ lean_inc(x_217); lean_dec_ref(x_212); x_218 = lean_unbox_usize(x_217); lean_dec(x_217); -x_219 = lean_llvm_build_store(x_1, x_3, x_218, x_101); +x_219 = lean_llvm_build_store(x_1, x_3, x_218, x_105); x_220 = l_LLVM_constTrue(x_1); x_221 = l_Lean_IR_EmitLLVM_callLeanSetPanicMessages___redArg(x_1, x_3, x_220, x_107); if (lean_obj_tag(x_221) == 0) @@ -62098,10 +62124,10 @@ lean_dec_ref(x_222); lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec(x_94); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); +lean_dec(x_91); lean_dec_ref(x_89); return x_221; } @@ -62122,10 +62148,10 @@ lean_dec_ref(x_224); lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec(x_94); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); +lean_dec(x_91); lean_dec_ref(x_89); return x_223; } @@ -62135,7 +62161,7 @@ lean_object* x_225; size_t x_226; lean_object* x_227; lean_object* x_228; lean_dec_ref(x_224); lean_dec_ref(x_223); x_225 = l_Lean_IR_EmitLLVM_emitMainFn___closed__2; -x_226 = lean_llvm_build_load2(x_1, x_3, x_104, x_101, x_225); +x_226 = lean_llvm_build_load2(x_1, x_3, x_103, x_105, x_225); x_227 = l_Lean_IR_EmitLLVM_emitMainFn___closed__3; x_228 = l_Lean_IR_EmitLLVM_callLeanIOResultIsOk___redArg(x_1, x_3, x_226, x_227, x_107); if (lean_obj_tag(x_228) == 0) @@ -62156,10 +62182,10 @@ lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec(x_94); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); +lean_dec(x_91); lean_dec_ref(x_89); x_231 = lean_ctor_get(x_229, 0); lean_inc(x_231); @@ -62193,24 +62219,24 @@ lean_inc(x_235); lean_dec_ref(x_229); x_236 = lean_box_usize(x_1); x_237 = lean_box_usize(x_226); -x_238 = lean_box_usize(x_92); -x_239 = lean_box_usize(x_98); -x_240 = lean_box_usize(x_93); -x_241 = lean_box_usize(x_95); +x_238 = lean_box_usize(x_96); +x_239 = lean_box_usize(x_94); +x_240 = lean_box_usize(x_100); +x_241 = lean_box_usize(x_98); x_242 = lean_box_usize(x_97); -lean_inc(x_96); -lean_inc(x_100); +lean_inc(x_92); +lean_inc(x_95); x_243 = lean_alloc_closure((void*)(l_Lean_IR_EmitLLVM_emitMainFn___lam__3___boxed), 17, 13); lean_closure_set(x_243, 0, x_236); lean_closure_set(x_243, 1, x_237); lean_closure_set(x_243, 2, x_90); -lean_closure_set(x_243, 3, x_94); +lean_closure_set(x_243, 3, x_91); lean_closure_set(x_243, 4, x_225); lean_closure_set(x_243, 5, x_238); -lean_closure_set(x_243, 6, x_100); +lean_closure_set(x_243, 6, x_95); lean_closure_set(x_243, 7, x_239); lean_closure_set(x_243, 8, x_240); -lean_closure_set(x_243, 9, x_96); +lean_closure_set(x_243, 9, x_92); lean_closure_set(x_243, 10, x_89); lean_closure_set(x_243, 11, x_241); lean_closure_set(x_243, 12, x_242); @@ -62231,9 +62257,9 @@ lean_dec_ref(x_247); lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); return x_246; } else @@ -62242,33 +62268,33 @@ lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_dec_ref(x_247); lean_dec_ref(x_246); x_248 = lean_box_usize(x_1); -x_249 = lean_box_usize(x_99); -x_250 = lean_box_usize(x_92); +x_249 = lean_box_usize(x_93); +x_250 = lean_box_usize(x_96); x_251 = lean_alloc_closure((void*)(l_Lean_IR_EmitLLVM_emitMainFn___lam__4___boxed), 10, 6); lean_closure_set(x_251, 0, x_248); lean_closure_set(x_251, 1, x_249); lean_closure_set(x_251, 2, x_250); lean_closure_set(x_251, 3, x_225); -lean_closure_set(x_251, 4, x_91); -lean_closure_set(x_251, 5, x_96); -x_252 = l_Lean_Environment_find_x3f(x_102, x_76, x_103); +lean_closure_set(x_251, 4, x_99); +lean_closure_set(x_251, 5, x_92); +x_252 = l_Lean_Environment_find_x3f(x_102, x_76, x_101); if (lean_obj_tag(x_252) == 0) { lean_object* x_253; lean_object* x_254; x_253 = l_Lean_IR_EmitLLVM_emitMainFn___closed__8; x_254 = l_panic___at___00Lean_IR_EmitLLVM_emitMainFn_spec__1(x_253); -x_7 = x_225; -x_8 = x_92; -x_9 = x_251; -x_10 = x_99; -x_11 = x_100; -x_12 = x_227; -x_13 = lean_box(0); -x_14 = x_225; -x_15 = x_101; -x_16 = x_107; -x_17 = x_104; -x_18 = x_106; +x_7 = x_93; +x_8 = x_251; +x_9 = x_225; +x_10 = x_96; +x_11 = x_95; +x_12 = x_103; +x_13 = x_225; +x_14 = x_107; +x_15 = lean_box(0); +x_16 = x_106; +x_17 = x_227; +x_18 = x_105; x_19 = x_254; goto block_74; } @@ -62278,18 +62304,18 @@ lean_object* x_255; x_255 = lean_ctor_get(x_252, 0); lean_inc(x_255); lean_dec_ref(x_252); -x_7 = x_225; -x_8 = x_92; -x_9 = x_251; -x_10 = x_99; -x_11 = x_100; -x_12 = x_227; -x_13 = lean_box(0); -x_14 = x_225; -x_15 = x_101; -x_16 = x_107; -x_17 = x_104; -x_18 = x_106; +x_7 = x_93; +x_8 = x_251; +x_9 = x_225; +x_10 = x_96; +x_11 = x_95; +x_12 = x_103; +x_13 = x_225; +x_14 = x_107; +x_15 = lean_box(0); +x_16 = x_106; +x_17 = x_227; +x_18 = x_105; x_19 = x_255; goto block_74; } @@ -62300,9 +62326,9 @@ else lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); return x_246; } } @@ -62313,10 +62339,10 @@ lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec(x_94); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); +lean_dec(x_91); lean_dec_ref(x_89); x_256 = lean_ctor_get(x_228, 0); lean_inc(x_256); @@ -62342,10 +62368,10 @@ else lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec(x_94); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); +lean_dec(x_91); lean_dec_ref(x_89); return x_223; } @@ -62356,10 +62382,10 @@ else lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec(x_94); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); +lean_dec(x_91); lean_dec_ref(x_89); return x_221; } @@ -62372,10 +62398,10 @@ uint8_t x_259; lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec(x_94); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); +lean_dec(x_91); lean_dec_ref(x_89); x_259 = !lean_is_exclusive(x_139); if (x_259 == 0) @@ -62407,10 +62433,10 @@ lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec(x_94); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); +lean_dec(x_91); lean_dec_ref(x_89); x_263 = lean_ctor_get(x_262, 0); lean_inc(x_263); @@ -62444,7 +62470,7 @@ lean_dec_ref(x_268); x_270 = lean_ctor_get(x_269, 0); lean_inc(x_270); lean_dec(x_269); -x_271 = l_Lean_IR_EmitLLVM_constInt8___redArg(x_1, x_96); +x_271 = l_Lean_IR_EmitLLVM_constInt8___redArg(x_1, x_92); x_272 = lean_ctor_get(x_271, 0); lean_inc(x_272); lean_dec_ref(x_271); @@ -62469,7 +62495,7 @@ x_281 = lean_box(0); lean_inc_ref(x_102); x_282 = l_Lean_PersistentEnvExtension_getState___redArg(x_280, x_277, x_102, x_279, x_281); lean_dec(x_279); -x_283 = l_Lean_Name_toString(x_276, x_105); +x_283 = l_Lean_Name_toString(x_276, x_104); x_284 = l_Lean_IR_EmitLLVM_emitMainFn___closed__1; x_285 = lean_string_append(x_283, x_284); x_286 = lean_unbox_usize(x_273); @@ -62496,10 +62522,10 @@ lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec(x_94); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); +lean_dec(x_91); lean_dec_ref(x_89); x_291 = lean_ctor_get(x_289, 0); lean_inc(x_291); @@ -62533,7 +62559,7 @@ lean_inc(x_295); lean_dec_ref(x_289); x_296 = lean_unbox_usize(x_295); lean_dec(x_295); -x_297 = lean_llvm_build_store(x_1, x_3, x_296, x_101); +x_297 = lean_llvm_build_store(x_1, x_3, x_296, x_105); x_298 = l_LLVM_constTrue(x_1); x_299 = l_Lean_IR_EmitLLVM_callLeanSetPanicMessages___redArg(x_1, x_3, x_298, x_107); if (lean_obj_tag(x_299) == 0) @@ -62547,10 +62573,10 @@ lean_dec_ref(x_300); lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec(x_94); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); +lean_dec(x_91); lean_dec_ref(x_89); return x_299; } @@ -62571,10 +62597,10 @@ lean_dec_ref(x_302); lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec(x_94); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); +lean_dec(x_91); lean_dec_ref(x_89); return x_301; } @@ -62584,7 +62610,7 @@ lean_object* x_303; size_t x_304; lean_object* x_305; lean_object* x_306; lean_dec_ref(x_302); lean_dec_ref(x_301); x_303 = l_Lean_IR_EmitLLVM_emitMainFn___closed__2; -x_304 = lean_llvm_build_load2(x_1, x_3, x_104, x_101, x_303); +x_304 = lean_llvm_build_load2(x_1, x_3, x_103, x_105, x_303); x_305 = l_Lean_IR_EmitLLVM_emitMainFn___closed__3; x_306 = l_Lean_IR_EmitLLVM_callLeanIOResultIsOk___redArg(x_1, x_3, x_304, x_305, x_107); if (lean_obj_tag(x_306) == 0) @@ -62605,10 +62631,10 @@ lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec(x_94); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); +lean_dec(x_91); lean_dec_ref(x_89); x_309 = lean_ctor_get(x_307, 0); lean_inc(x_309); @@ -62642,24 +62668,24 @@ lean_inc(x_313); lean_dec_ref(x_307); x_314 = lean_box_usize(x_1); x_315 = lean_box_usize(x_304); -x_316 = lean_box_usize(x_92); -x_317 = lean_box_usize(x_98); -x_318 = lean_box_usize(x_93); -x_319 = lean_box_usize(x_95); +x_316 = lean_box_usize(x_96); +x_317 = lean_box_usize(x_94); +x_318 = lean_box_usize(x_100); +x_319 = lean_box_usize(x_98); x_320 = lean_box_usize(x_97); -lean_inc(x_96); -lean_inc(x_100); +lean_inc(x_92); +lean_inc(x_95); x_321 = lean_alloc_closure((void*)(l_Lean_IR_EmitLLVM_emitMainFn___lam__3___boxed), 17, 13); lean_closure_set(x_321, 0, x_314); lean_closure_set(x_321, 1, x_315); lean_closure_set(x_321, 2, x_90); -lean_closure_set(x_321, 3, x_94); +lean_closure_set(x_321, 3, x_91); lean_closure_set(x_321, 4, x_303); lean_closure_set(x_321, 5, x_316); -lean_closure_set(x_321, 6, x_100); +lean_closure_set(x_321, 6, x_95); lean_closure_set(x_321, 7, x_317); lean_closure_set(x_321, 8, x_318); -lean_closure_set(x_321, 9, x_96); +lean_closure_set(x_321, 9, x_92); lean_closure_set(x_321, 10, x_89); lean_closure_set(x_321, 11, x_319); lean_closure_set(x_321, 12, x_320); @@ -62680,9 +62706,9 @@ lean_dec_ref(x_325); lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); return x_324; } else @@ -62691,33 +62717,33 @@ lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_dec_ref(x_325); lean_dec_ref(x_324); x_326 = lean_box_usize(x_1); -x_327 = lean_box_usize(x_99); -x_328 = lean_box_usize(x_92); +x_327 = lean_box_usize(x_93); +x_328 = lean_box_usize(x_96); x_329 = lean_alloc_closure((void*)(l_Lean_IR_EmitLLVM_emitMainFn___lam__4___boxed), 10, 6); lean_closure_set(x_329, 0, x_326); lean_closure_set(x_329, 1, x_327); lean_closure_set(x_329, 2, x_328); lean_closure_set(x_329, 3, x_303); -lean_closure_set(x_329, 4, x_91); -lean_closure_set(x_329, 5, x_96); -x_330 = l_Lean_Environment_find_x3f(x_102, x_76, x_103); +lean_closure_set(x_329, 4, x_99); +lean_closure_set(x_329, 5, x_92); +x_330 = l_Lean_Environment_find_x3f(x_102, x_76, x_101); if (lean_obj_tag(x_330) == 0) { lean_object* x_331; lean_object* x_332; x_331 = l_Lean_IR_EmitLLVM_emitMainFn___closed__8; x_332 = l_panic___at___00Lean_IR_EmitLLVM_emitMainFn_spec__1(x_331); -x_7 = x_303; -x_8 = x_92; -x_9 = x_329; -x_10 = x_99; -x_11 = x_100; -x_12 = x_305; -x_13 = lean_box(0); -x_14 = x_303; -x_15 = x_101; -x_16 = x_107; -x_17 = x_104; -x_18 = x_106; +x_7 = x_93; +x_8 = x_329; +x_9 = x_303; +x_10 = x_96; +x_11 = x_95; +x_12 = x_103; +x_13 = x_303; +x_14 = x_107; +x_15 = lean_box(0); +x_16 = x_106; +x_17 = x_305; +x_18 = x_105; x_19 = x_332; goto block_74; } @@ -62727,18 +62753,18 @@ lean_object* x_333; x_333 = lean_ctor_get(x_330, 0); lean_inc(x_333); lean_dec_ref(x_330); -x_7 = x_303; -x_8 = x_92; -x_9 = x_329; -x_10 = x_99; -x_11 = x_100; -x_12 = x_305; -x_13 = lean_box(0); -x_14 = x_303; -x_15 = x_101; -x_16 = x_107; -x_17 = x_104; -x_18 = x_106; +x_7 = x_93; +x_8 = x_329; +x_9 = x_303; +x_10 = x_96; +x_11 = x_95; +x_12 = x_103; +x_13 = x_303; +x_14 = x_107; +x_15 = lean_box(0); +x_16 = x_106; +x_17 = x_305; +x_18 = x_105; x_19 = x_333; goto block_74; } @@ -62749,9 +62775,9 @@ else lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); return x_324; } } @@ -62762,10 +62788,10 @@ lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec(x_94); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); +lean_dec(x_91); lean_dec_ref(x_89); x_334 = lean_ctor_get(x_306, 0); lean_inc(x_334); @@ -62791,10 +62817,10 @@ else lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec(x_94); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); +lean_dec(x_91); lean_dec_ref(x_89); return x_301; } @@ -62805,10 +62831,10 @@ else lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec(x_94); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); +lean_dec(x_91); lean_dec_ref(x_89); return x_299; } @@ -62820,10 +62846,10 @@ lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec(x_94); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); +lean_dec(x_91); lean_dec_ref(x_89); x_337 = lean_ctor_get(x_288, 0); lean_inc(x_337); @@ -62851,10 +62877,10 @@ uint8_t x_340; lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec(x_94); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); +lean_dec(x_91); lean_dec_ref(x_89); x_340 = !lean_is_exclusive(x_112); if (x_340 == 0) @@ -62879,10 +62905,10 @@ else lean_dec_ref(x_107); lean_dec(x_106); lean_dec_ref(x_102); -lean_dec(x_100); -lean_dec(x_96); -lean_dec(x_94); -lean_dec_ref(x_91); +lean_dec_ref(x_99); +lean_dec(x_95); +lean_dec(x_92); +lean_dec(x_91); lean_dec_ref(x_89); return x_110; } @@ -63148,33 +63174,33 @@ else size_t x_425; size_t x_426; size_t x_427; size_t x_428; size_t x_429; size_t x_430; size_t x_431; lean_dec_ref(x_424); lean_dec_ref(x_423); -x_425 = lean_unbox_usize(x_404); -x_426 = lean_unbox_usize(x_418); -lean_dec(x_418); -x_427 = lean_unbox_usize(x_381); -lean_dec(x_381); -x_428 = lean_unbox_usize(x_391); +x_425 = lean_unbox_usize(x_394); +x_426 = lean_unbox_usize(x_391); lean_dec(x_391); -x_429 = lean_unbox_usize(x_394); -x_430 = lean_unbox_usize(x_404); -lean_dec(x_404); -x_431 = lean_unbox_usize(x_394); +x_427 = lean_unbox_usize(x_404); +x_428 = lean_unbox_usize(x_381); +lean_dec(x_381); +x_429 = lean_unbox_usize(x_418); +lean_dec(x_418); +x_430 = lean_unbox_usize(x_394); lean_dec(x_394); -x_91 = x_410; -x_92 = x_425; -x_93 = x_406; -x_94 = x_366; -x_95 = x_426; -x_96 = x_422; -x_97 = x_427; -x_98 = x_428; -x_99 = x_429; -x_100 = x_421; -x_101 = x_430; +x_431 = lean_unbox_usize(x_404); +lean_dec(x_404); +x_91 = x_366; +x_92 = x_422; +x_93 = x_425; +x_94 = x_426; +x_95 = x_421; +x_96 = x_427; +x_97 = x_428; +x_98 = x_429; +x_99 = x_410; +x_100 = x_406; +x_101 = x_371; x_102 = x_354; -x_103 = x_371; -x_104 = x_431; -x_105 = x_344; +x_103 = x_430; +x_104 = x_344; +x_105 = x_431; x_106 = x_4; x_107 = x_5; x_108 = lean_box(0); @@ -63223,33 +63249,33 @@ else size_t x_434; size_t x_435; size_t x_436; size_t x_437; size_t x_438; size_t x_439; size_t x_440; lean_dec_ref(x_433); lean_dec_ref(x_432); -x_434 = lean_unbox_usize(x_404); -x_435 = lean_unbox_usize(x_418); -lean_dec(x_418); -x_436 = lean_unbox_usize(x_381); -lean_dec(x_381); -x_437 = lean_unbox_usize(x_391); +x_434 = lean_unbox_usize(x_394); +x_435 = lean_unbox_usize(x_391); lean_dec(x_391); -x_438 = lean_unbox_usize(x_394); -x_439 = lean_unbox_usize(x_404); -lean_dec(x_404); -x_440 = lean_unbox_usize(x_394); +x_436 = lean_unbox_usize(x_404); +x_437 = lean_unbox_usize(x_381); +lean_dec(x_381); +x_438 = lean_unbox_usize(x_418); +lean_dec(x_418); +x_439 = lean_unbox_usize(x_394); lean_dec(x_394); -x_91 = x_410; -x_92 = x_434; -x_93 = x_406; -x_94 = x_366; -x_95 = x_435; -x_96 = x_422; -x_97 = x_436; -x_98 = x_437; -x_99 = x_438; -x_100 = x_421; -x_101 = x_439; +x_440 = lean_unbox_usize(x_404); +lean_dec(x_404); +x_91 = x_366; +x_92 = x_422; +x_93 = x_434; +x_94 = x_435; +x_95 = x_421; +x_96 = x_436; +x_97 = x_437; +x_98 = x_438; +x_99 = x_410; +x_100 = x_406; +x_101 = x_371; x_102 = x_354; -x_103 = x_371; -x_104 = x_440; -x_105 = x_344; +x_103 = x_439; +x_104 = x_344; +x_105 = x_440; x_106 = x_4; x_107 = x_5; x_108 = lean_box(0); @@ -63346,33 +63372,33 @@ else size_t x_453; size_t x_454; size_t x_455; size_t x_456; size_t x_457; size_t x_458; size_t x_459; lean_dec_ref(x_452); lean_dec_ref(x_451); -x_453 = lean_unbox_usize(x_404); -x_454 = lean_unbox_usize(x_446); -lean_dec(x_446); -x_455 = lean_unbox_usize(x_381); -lean_dec(x_381); -x_456 = lean_unbox_usize(x_391); +x_453 = lean_unbox_usize(x_394); +x_454 = lean_unbox_usize(x_391); lean_dec(x_391); -x_457 = lean_unbox_usize(x_394); -x_458 = lean_unbox_usize(x_404); -lean_dec(x_404); -x_459 = lean_unbox_usize(x_394); +x_455 = lean_unbox_usize(x_404); +x_456 = lean_unbox_usize(x_381); +lean_dec(x_381); +x_457 = lean_unbox_usize(x_446); +lean_dec(x_446); +x_458 = lean_unbox_usize(x_394); lean_dec(x_394); -x_91 = x_410; -x_92 = x_453; -x_93 = x_406; -x_94 = x_366; -x_95 = x_454; -x_96 = x_450; -x_97 = x_455; -x_98 = x_456; -x_99 = x_457; -x_100 = x_449; -x_101 = x_458; +x_459 = lean_unbox_usize(x_404); +lean_dec(x_404); +x_91 = x_366; +x_92 = x_450; +x_93 = x_453; +x_94 = x_454; +x_95 = x_449; +x_96 = x_455; +x_97 = x_456; +x_98 = x_457; +x_99 = x_410; +x_100 = x_406; +x_101 = x_371; x_102 = x_354; -x_103 = x_371; -x_104 = x_459; -x_105 = x_344; +x_103 = x_458; +x_104 = x_344; +x_105 = x_459; x_106 = x_4; x_107 = x_5; x_108 = lean_box(0); @@ -63421,33 +63447,33 @@ else size_t x_462; size_t x_463; size_t x_464; size_t x_465; size_t x_466; size_t x_467; size_t x_468; lean_dec_ref(x_461); lean_dec_ref(x_460); -x_462 = lean_unbox_usize(x_404); -x_463 = lean_unbox_usize(x_446); -lean_dec(x_446); -x_464 = lean_unbox_usize(x_381); -lean_dec(x_381); -x_465 = lean_unbox_usize(x_391); +x_462 = lean_unbox_usize(x_394); +x_463 = lean_unbox_usize(x_391); lean_dec(x_391); -x_466 = lean_unbox_usize(x_394); -x_467 = lean_unbox_usize(x_404); -lean_dec(x_404); -x_468 = lean_unbox_usize(x_394); +x_464 = lean_unbox_usize(x_404); +x_465 = lean_unbox_usize(x_381); +lean_dec(x_381); +x_466 = lean_unbox_usize(x_446); +lean_dec(x_446); +x_467 = lean_unbox_usize(x_394); lean_dec(x_394); -x_91 = x_410; -x_92 = x_462; -x_93 = x_406; -x_94 = x_366; -x_95 = x_463; -x_96 = x_450; -x_97 = x_464; -x_98 = x_465; -x_99 = x_466; -x_100 = x_449; -x_101 = x_467; +x_468 = lean_unbox_usize(x_404); +lean_dec(x_404); +x_91 = x_366; +x_92 = x_450; +x_93 = x_462; +x_94 = x_463; +x_95 = x_449; +x_96 = x_464; +x_97 = x_465; +x_98 = x_466; +x_99 = x_410; +x_100 = x_406; +x_101 = x_371; x_102 = x_354; -x_103 = x_371; -x_104 = x_468; -x_105 = x_344; +x_103 = x_467; +x_104 = x_344; +x_105 = x_468; x_106 = x_4; x_107 = x_5; x_108 = lean_box(0); @@ -63634,33 +63660,33 @@ else size_t x_499; size_t x_500; size_t x_501; size_t x_502; size_t x_503; size_t x_504; size_t x_505; lean_dec_ref(x_498); lean_dec_ref(x_497); -x_499 = lean_unbox_usize(x_477); -x_500 = lean_unbox_usize(x_492); -lean_dec(x_492); -x_501 = lean_unbox_usize(x_381); -lean_dec(x_381); -x_502 = lean_unbox_usize(x_391); +x_499 = lean_unbox_usize(x_394); +x_500 = lean_unbox_usize(x_391); lean_dec(x_391); -x_503 = lean_unbox_usize(x_394); -x_504 = lean_unbox_usize(x_477); -lean_dec(x_477); -x_505 = lean_unbox_usize(x_394); +x_501 = lean_unbox_usize(x_477); +x_502 = lean_unbox_usize(x_381); +lean_dec(x_381); +x_503 = lean_unbox_usize(x_492); +lean_dec(x_492); +x_504 = lean_unbox_usize(x_394); lean_dec(x_394); -x_91 = x_483; -x_92 = x_499; -x_93 = x_479; -x_94 = x_366; -x_95 = x_500; -x_96 = x_496; -x_97 = x_501; -x_98 = x_502; -x_99 = x_503; -x_100 = x_495; -x_101 = x_504; +x_505 = lean_unbox_usize(x_477); +lean_dec(x_477); +x_91 = x_366; +x_92 = x_496; +x_93 = x_499; +x_94 = x_500; +x_95 = x_495; +x_96 = x_501; +x_97 = x_502; +x_98 = x_503; +x_99 = x_483; +x_100 = x_479; +x_101 = x_371; x_102 = x_354; -x_103 = x_371; -x_104 = x_505; -x_105 = x_344; +x_103 = x_504; +x_104 = x_344; +x_105 = x_505; x_106 = x_4; x_107 = x_5; x_108 = lean_box(0); @@ -63709,33 +63735,33 @@ else size_t x_508; size_t x_509; size_t x_510; size_t x_511; size_t x_512; size_t x_513; size_t x_514; lean_dec_ref(x_507); lean_dec_ref(x_506); -x_508 = lean_unbox_usize(x_477); -x_509 = lean_unbox_usize(x_492); -lean_dec(x_492); -x_510 = lean_unbox_usize(x_381); -lean_dec(x_381); -x_511 = lean_unbox_usize(x_391); +x_508 = lean_unbox_usize(x_394); +x_509 = lean_unbox_usize(x_391); lean_dec(x_391); -x_512 = lean_unbox_usize(x_394); -x_513 = lean_unbox_usize(x_477); -lean_dec(x_477); -x_514 = lean_unbox_usize(x_394); +x_510 = lean_unbox_usize(x_477); +x_511 = lean_unbox_usize(x_381); +lean_dec(x_381); +x_512 = lean_unbox_usize(x_492); +lean_dec(x_492); +x_513 = lean_unbox_usize(x_394); lean_dec(x_394); -x_91 = x_483; -x_92 = x_508; -x_93 = x_479; -x_94 = x_366; -x_95 = x_509; -x_96 = x_496; -x_97 = x_510; -x_98 = x_511; -x_99 = x_512; -x_100 = x_495; -x_101 = x_513; +x_514 = lean_unbox_usize(x_477); +lean_dec(x_477); +x_91 = x_366; +x_92 = x_496; +x_93 = x_508; +x_94 = x_509; +x_95 = x_495; +x_96 = x_510; +x_97 = x_511; +x_98 = x_512; +x_99 = x_483; +x_100 = x_479; +x_101 = x_371; x_102 = x_354; -x_103 = x_371; -x_104 = x_514; -x_105 = x_344; +x_103 = x_513; +x_104 = x_344; +x_105 = x_514; x_106 = x_4; x_107 = x_5; x_108 = lean_box(0); @@ -63982,33 +64008,33 @@ else size_t x_559; size_t x_560; size_t x_561; size_t x_562; size_t x_563; size_t x_564; size_t x_565; lean_dec_ref(x_558); lean_dec_ref(x_557); -x_559 = lean_unbox_usize(x_537); -x_560 = lean_unbox_usize(x_552); -lean_dec(x_552); -x_561 = lean_unbox_usize(x_381); -lean_dec(x_381); -x_562 = lean_unbox_usize(x_523); +x_559 = lean_unbox_usize(x_526); +x_560 = lean_unbox_usize(x_523); lean_dec(x_523); -x_563 = lean_unbox_usize(x_526); -x_564 = lean_unbox_usize(x_537); -lean_dec(x_537); -x_565 = lean_unbox_usize(x_526); +x_561 = lean_unbox_usize(x_537); +x_562 = lean_unbox_usize(x_381); +lean_dec(x_381); +x_563 = lean_unbox_usize(x_552); +lean_dec(x_552); +x_564 = lean_unbox_usize(x_526); lean_dec(x_526); -x_91 = x_543; -x_92 = x_559; -x_93 = x_539; -x_94 = x_366; -x_95 = x_560; -x_96 = x_556; -x_97 = x_561; -x_98 = x_562; -x_99 = x_563; -x_100 = x_555; -x_101 = x_564; +x_565 = lean_unbox_usize(x_537); +lean_dec(x_537); +x_91 = x_366; +x_92 = x_556; +x_93 = x_559; +x_94 = x_560; +x_95 = x_555; +x_96 = x_561; +x_97 = x_562; +x_98 = x_563; +x_99 = x_543; +x_100 = x_539; +x_101 = x_371; x_102 = x_354; -x_103 = x_371; -x_104 = x_565; -x_105 = x_344; +x_103 = x_564; +x_104 = x_344; +x_105 = x_565; x_106 = x_4; x_107 = x_5; x_108 = lean_box(0); @@ -64057,33 +64083,33 @@ else size_t x_568; size_t x_569; size_t x_570; size_t x_571; size_t x_572; size_t x_573; size_t x_574; lean_dec_ref(x_567); lean_dec_ref(x_566); -x_568 = lean_unbox_usize(x_537); -x_569 = lean_unbox_usize(x_552); -lean_dec(x_552); -x_570 = lean_unbox_usize(x_381); -lean_dec(x_381); -x_571 = lean_unbox_usize(x_523); +x_568 = lean_unbox_usize(x_526); +x_569 = lean_unbox_usize(x_523); lean_dec(x_523); -x_572 = lean_unbox_usize(x_526); -x_573 = lean_unbox_usize(x_537); -lean_dec(x_537); -x_574 = lean_unbox_usize(x_526); +x_570 = lean_unbox_usize(x_537); +x_571 = lean_unbox_usize(x_381); +lean_dec(x_381); +x_572 = lean_unbox_usize(x_552); +lean_dec(x_552); +x_573 = lean_unbox_usize(x_526); lean_dec(x_526); -x_91 = x_543; -x_92 = x_568; -x_93 = x_539; -x_94 = x_366; -x_95 = x_569; -x_96 = x_556; -x_97 = x_570; -x_98 = x_571; -x_99 = x_572; -x_100 = x_555; -x_101 = x_573; +x_574 = lean_unbox_usize(x_537); +lean_dec(x_537); +x_91 = x_366; +x_92 = x_556; +x_93 = x_568; +x_94 = x_569; +x_95 = x_555; +x_96 = x_570; +x_97 = x_571; +x_98 = x_572; +x_99 = x_543; +x_100 = x_539; +x_101 = x_371; x_102 = x_354; -x_103 = x_371; -x_104 = x_574; -x_105 = x_344; +x_103 = x_573; +x_104 = x_344; +x_105 = x_574; x_106 = x_4; x_107 = x_5; x_108 = lean_box(0); @@ -64182,7 +64208,7 @@ return x_587; block_74: { lean_object* x_20; -x_20 = l_Lean_IR_EmitLLVM_callLeanFinalizeTaskManager___redArg(x_1, x_3, x_16); +x_20 = l_Lean_IR_EmitLLVM_callLeanFinalizeTaskManager___redArg(x_1, x_3, x_14); if (lean_obj_tag(x_20) == 0) { lean_object* x_21; @@ -64192,13 +64218,13 @@ if (lean_obj_tag(x_21) == 0) { lean_dec_ref(x_21); lean_dec_ref(x_19); -lean_dec(x_18); -lean_dec_ref(x_16); +lean_dec_ref(x_17); +lean_dec(x_16); lean_dec_ref(x_14); -lean_dec_ref(x_12); +lean_dec_ref(x_13); lean_dec(x_11); lean_dec_ref(x_9); -lean_dec_ref(x_7); +lean_dec_ref(x_8); return x_20; } else @@ -64206,8 +64232,8 @@ else size_t x_22; lean_object* x_23; lean_dec_ref(x_21); lean_dec_ref(x_20); -x_22 = lean_llvm_build_load2(x_1, x_3, x_17, x_15, x_14); -x_23 = l_Lean_IR_EmitLLVM_callLeanIOResultIsOk___redArg(x_1, x_3, x_22, x_12, x_16); +x_22 = lean_llvm_build_load2(x_1, x_3, x_12, x_18, x_13); +x_23 = l_Lean_IR_EmitLLVM_callLeanIOResultIsOk___redArg(x_1, x_3, x_22, x_17, x_14); if (lean_obj_tag(x_23) == 0) { uint8_t x_24; @@ -64220,11 +64246,11 @@ if (lean_obj_tag(x_25) == 0) { uint8_t x_26; lean_dec_ref(x_19); -lean_dec(x_18); -lean_dec_ref(x_16); +lean_dec(x_16); +lean_dec_ref(x_14); lean_dec(x_11); lean_dec_ref(x_9); -lean_dec_ref(x_7); +lean_dec_ref(x_8); x_26 = !lean_is_exclusive(x_25); if (x_26 == 0) { @@ -64257,8 +64283,8 @@ x_32 = l_Lean_Expr_appArg_x21(x_31); lean_dec_ref(x_31); x_33 = lean_box_usize(x_1); x_34 = lean_box_usize(x_22); -x_35 = lean_box_usize(x_10); -x_36 = lean_box_usize(x_8); +x_35 = lean_box_usize(x_7); +x_36 = lean_box_usize(x_10); x_37 = lean_alloc_closure((void*)(l_Lean_IR_EmitLLVM_emitMainFn___lam__5___boxed), 11, 7); lean_closure_set(x_37, 0, x_32); lean_closure_set(x_37, 1, x_33); @@ -64266,11 +64292,11 @@ lean_closure_set(x_37, 2, x_34); lean_closure_set(x_37, 3, x_11); lean_closure_set(x_37, 4, x_35); lean_closure_set(x_37, 5, x_36); -lean_closure_set(x_37, 6, x_7); +lean_closure_set(x_37, 6, x_9); x_38 = l_Lean_IR_EmitLLVM_emitMainFn___closed__0; x_39 = lean_unbox_usize(x_29); lean_dec(x_29); -x_40 = l_Lean_IR_EmitLLVM_buildIfThenElse__(x_1, x_3, x_38, x_39, x_37, x_9, x_18, x_16); +x_40 = l_Lean_IR_EmitLLVM_buildIfThenElse__(x_1, x_3, x_38, x_39, x_37, x_8, x_16, x_14); if (lean_obj_tag(x_40) == 0) { lean_object* x_41; @@ -64324,11 +64350,11 @@ if (lean_obj_tag(x_49) == 0) { lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_dec_ref(x_19); -lean_dec(x_18); -lean_dec_ref(x_16); +lean_dec(x_16); +lean_dec_ref(x_14); lean_dec(x_11); lean_dec_ref(x_9); -lean_dec_ref(x_7); +lean_dec_ref(x_8); x_50 = lean_ctor_get(x_49, 0); lean_inc(x_50); if (lean_is_exclusive(x_49)) { @@ -64362,8 +64388,8 @@ x_57 = l_Lean_Expr_appArg_x21(x_56); lean_dec_ref(x_56); x_58 = lean_box_usize(x_1); x_59 = lean_box_usize(x_22); -x_60 = lean_box_usize(x_10); -x_61 = lean_box_usize(x_8); +x_60 = lean_box_usize(x_7); +x_61 = lean_box_usize(x_10); x_62 = lean_alloc_closure((void*)(l_Lean_IR_EmitLLVM_emitMainFn___lam__5___boxed), 11, 7); lean_closure_set(x_62, 0, x_57); lean_closure_set(x_62, 1, x_58); @@ -64371,11 +64397,11 @@ lean_closure_set(x_62, 2, x_59); lean_closure_set(x_62, 3, x_11); lean_closure_set(x_62, 4, x_60); lean_closure_set(x_62, 5, x_61); -lean_closure_set(x_62, 6, x_7); +lean_closure_set(x_62, 6, x_9); x_63 = l_Lean_IR_EmitLLVM_emitMainFn___closed__0; x_64 = lean_unbox_usize(x_54); lean_dec(x_54); -x_65 = l_Lean_IR_EmitLLVM_buildIfThenElse__(x_1, x_3, x_63, x_64, x_62, x_9, x_18, x_16); +x_65 = l_Lean_IR_EmitLLVM_buildIfThenElse__(x_1, x_3, x_63, x_64, x_62, x_8, x_16, x_14); if (lean_obj_tag(x_65) == 0) { lean_object* x_66; @@ -64419,11 +64445,11 @@ else { uint8_t x_71; lean_dec_ref(x_19); -lean_dec(x_18); -lean_dec_ref(x_16); +lean_dec(x_16); +lean_dec_ref(x_14); lean_dec(x_11); lean_dec_ref(x_9); -lean_dec_ref(x_7); +lean_dec_ref(x_8); x_71 = !lean_is_exclusive(x_23); if (x_71 == 0) { @@ -64445,13 +64471,13 @@ return x_73; else { lean_dec_ref(x_19); -lean_dec(x_18); -lean_dec_ref(x_16); +lean_dec_ref(x_17); +lean_dec(x_16); lean_dec_ref(x_14); -lean_dec_ref(x_12); +lean_dec_ref(x_13); lean_dec(x_11); lean_dec_ref(x_9); -lean_dec_ref(x_7); +lean_dec_ref(x_8); return x_20; } } diff --git a/stage0/stdlib/Lean/Compiler/IR/EmitUtil.c b/stage0/stdlib/Lean/Compiler/IR/EmitUtil.c index 64d242a235ab..08324435a91e 100644 --- a/stage0/stdlib/Lean/Compiler/IR/EmitUtil.c +++ b/stage0/stdlib/Lean/Compiler/IR/EmitUtil.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Compiler.IR.EmitUtil -// Imports: public import Lean.Compiler.InitAttr public import Lean.Compiler.IR.CompilerM +// Imports: public import Lean.Compiler.InitAttr public import Lean.Compiler.IR.CompilerM import Lean.Compiler.NameMangling #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -13,103 +13,200 @@ #ifdef __cplusplus extern "C" { #endif +LEAN_EXPORT lean_object* l_Lean_IR_instInhabitedStructTypeInfo; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectMaps_collectParams_spec__0_spec__0___redArg___boxed(lean_object*, lean_object*); +lean_object* lean_mk_empty_array_with_capacity(lean_object*); +LEAN_EXPORT lean_object* l_List_forM___at___00Lean_IR_collectStructTypes_spec__0(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_IR_instBEqJoinPointId_beq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_collectStructTypes(lean_object*); +uint64_t lean_uint64_of_nat(lean_object*); +uint64_t lean_uint64_mix_hash(uint64_t, uint64_t); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_CollectStructTypes_registerType_spec__0___boxed(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_IR_instBEqVarId_beq(lean_object*, lean_object*); size_t lean_uint64_to_usize(uint64_t); LEAN_EXPORT lean_object* l___private_Std_Data_DHashMap_Internal_Defs_0__Std_DHashMap_Internal_Raw_u2080_expand_go___at___00Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectMaps_collectFnBody_spec__0_spec__1_spec__2___redArg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectMaps_collectParams_spec__0_spec__0___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_CollectMaps_collectParams(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_CollectMaps_collectJP___closed__0; +lean_object* lean_array_push(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_panic___at___00Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0_spec__1___redArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_CollectStructTypes_registerType_spec__0_spec__0___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_CollectStructTypes_registerType_spec__0_spec__0___redArg(lean_object*, lean_object*); +static lean_object* l_Lean_IR_collectStructTypes___closed__1; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_CollectStructTypes_collectFnBody_spec__0(lean_object*, size_t, size_t, lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_IR_instBEqIRTypeApprox___lam__0___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectMaps_collectFnBody_spec__0_spec__0___redArg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectMaps_collectParams_spec__0_spec__2___redArg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_IR_instHashableVarId_hash___boxed(lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0___redArg(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_IR_FnBody_isTerminal(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_EmitUtil_0__Lean_IR_CollectStructTypes_addReshape_match__3_splitter___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_IR_instBEqIRTypeApprox___lam__0(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___00__private_Std_Data_DHashMap_Internal_Defs_0__Std_DHashMap_Internal_Raw_u2080_expand_go___at___00Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectMaps_collectFnBody_spec__0_spec__1_spec__2_spec__4___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___00__private_Std_Data_DHashMap_Internal_Defs_0__Std_DHashMap_Internal_Raw_u2080_expand_go___at___00Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectMaps_collectFnBody_spec__0_spec__1_spec__2_spec__4(lean_object*, lean_object*, lean_object*); +lean_object* lean_array_fget_borrowed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_CollectUsedDecls_collectFnBody_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_IR_instBEqJoinPointId_beq___boxed(lean_object*, lean_object*); +static lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___redArg___closed__3; +static lean_object* l_Lean_IR_collectStructTypes___closed__3; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_CollectMaps_collectFnBody_spec__1(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___00__private_Std_Data_DHashMap_Internal_Defs_0__Std_DHashMap_Internal_Raw_u2080_expand_go___at___00Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectMaps_collectParams_spec__0_spec__1_spec__2_spec__4___redArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__3___redArg(lean_object*); uint8_t l_Lean_Name_isPrefixOf(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectMaps_collectFnBody_spec__0_spec__0___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_CollectStructTypes_registerType(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_CollectStructTypes_collectParams(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__2___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Std_Data_DHashMap_Internal_Defs_0__Std_DHashMap_Internal_Raw_u2080_expand_go___at___00Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectMaps_collectParams_spec__0_spec__1_spec__2___redArg(lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); static lean_object* l_Lean_IR_mkVarJPMaps___closed__1; +static lean_object* l_Lean_IR_CollectStructTypes_registerType___closed__0; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___00__private_Std_Data_DHashMap_Internal_Defs_0__Std_DHashMap_Internal_Raw_u2080_expand_go___at___00Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__3_spec__4_spec__6___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_CollectMaps_collectDecl(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectMaps_collectFnBody_spec__0___redArg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_IR_usesModuleFrom(lean_object*, lean_object*); +uint8_t l_Lean_IR_IRType_compatibleWith(lean_object*, lean_object*, uint8_t); lean_object* l_Std_DHashMap_Internal_Raw_u2080_insert___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_CollectUsedDecls_collect___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_get_init_fn_name_for(lean_object*, lean_object*); lean_object* l_Lean_IR_instBEqVarId_beq___boxed(lean_object*, lean_object*); uint64_t lean_uint64_shift_right(uint64_t, uint64_t); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_div(lean_object*, lean_object*); +lean_object* l_Array_empty(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_EmitUtil_0__Lean_IR_CollectStructTypes_addReshape_match__1_splitter(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_IRType_hashApprox___boxed(lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectMaps_collectParams_spec__0___redArg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Array_contains___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__1_spec__2(lean_object*, lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Array_contains___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__1_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_CollectUsedDecls_collect___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_isTailCallTo___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_CollectStructTypes_registerType_spec__0(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_CollectStructTypes_registerType_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_IR_FnBody_body(lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_CollectStructTypes_registerType_spec__0_spec__0___redArg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_CollectStructTypes_addReshapeEntry___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_CollectMaps_collectVar(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_CollectStructTypes_collectParams_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_EmitUtil_0__Lean_IR_CollectStructTypes_addReshape_match__3_splitter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___redArg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_mkVarJPMaps___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_CollectUsedDecls_collectDecl(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectMaps_collectParams_spec__0_spec__1___redArg(lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_CollectMaps_collectFnBody(lean_object*, lean_object*); +static lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___redArg___closed__1; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__3(lean_object*, lean_object*); +LEAN_EXPORT uint64_t l_Lean_IR_IRType_hashApprox(lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_CollectStructTypes_collectFnBody(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectMaps_collectFnBody_spec__0_spec__2(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_IR_instHashableJoinPointId_hash___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_CollectStructTypes_collectDecl(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_CollectStructTypes_collectParams_spec__0(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_CollectMaps_collectFnBody_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_NameSet_insert(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__2(lean_object*, lean_object*, lean_object*); +static uint64_t l_Lean_IR_IRType_hashApprox___closed__0; +LEAN_EXPORT lean_object* l_Array_contains___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_CollectStructTypes_addReshape(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_CollectMaps_collectParams_spec__1(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectMaps_collectParams_spec__0(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectMaps_collectFnBody_spec__0_spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_CollectStructTypes_collectParams___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectMaps_collectFnBody_spec__0_spec__0(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_CollectMaps_collectParams_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__4___redArg(lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); +static lean_object* l_Lean_IR_collectStructTypes___closed__2; +static lean_object* l_Lean_IR_collectStructTypes___closed__0; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectMaps_collectParams_spec__0_spec__0___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_EmitUtil_0__Lean_IR_CollectStructTypes_addReshape_match__1_splitter___redArg(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_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_CollectStructTypes_addReshape_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_CollectStructTypes_collectFnBody_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectMaps_collectFnBody_spec__0(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___redArg___closed__0; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_CollectStructTypes_registerType_spec__0_spec__0(lean_object*, 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_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_CollectStructTypes_registerType_spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectMaps_collectParams_spec__0_spec__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_usesModuleFrom___lam__0___boxed(lean_object*, lean_object*); +static lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___redArg___closed__2; lean_object* l_Lean_IR_Alt_body(lean_object*); uint64_t lean_uint64_xor(uint64_t, uint64_t); +lean_object* l_Array_toSubarray___redArg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__2___redArg___boxed(lean_object*, lean_object*); +lean_object* lean_panic_fn(lean_object*, lean_object*); +static lean_object* l_Lean_IR_instHashableIRTypeApprox___closed__0; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1___redArg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_CollectMaps_collectVar___closed__1; LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectMaps_collectParams_spec__0_spec__0(lean_object*, lean_object*, 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_EXPORT lean_object* l_panic___at___00Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0_spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___00__private_Std_Data_DHashMap_Internal_Defs_0__Std_DHashMap_Internal_Raw_u2080_expand_go___at___00Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectMaps_collectParams_spec__0_spec__1_spec__2_spec__4(lean_object*, lean_object*, lean_object*); lean_object* lean_nat_mul(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_IR_usesModuleFrom___lam__0(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Data_DHashMap_Internal_Defs_0__Std_DHashMap_Internal_Raw_u2080_expand_go___at___00Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__3_spec__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_CollectMaps_collectJP(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_CollectMaps_collectJP___closed__1; +LEAN_EXPORT lean_object* l_Lean_IR_CollectStructTypes_addReshapeEntry(lean_object*, lean_object*, lean_object*); size_t lean_usize_sub(size_t, size_t); +uint8_t l_Lean_IR_IRType_isStruct(lean_object*); uint64_t l_Lean_IR_instHashableJoinPointId_hash(lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_CollectStructTypes_registerType_spec__0___redArg(lean_object*, lean_object*); static lean_object* l_Lean_IR_mkVarJPMaps___closed__2; +LEAN_EXPORT lean_object* l___private_Std_Data_DHashMap_Internal_Defs_0__Std_DHashMap_Internal_Raw_u2080_expand_go___at___00Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__3_spec__4___redArg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_instBEqIRTypeApprox; size_t lean_usize_add(size_t, size_t); +lean_object* l_mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectMaps_collectFnBody_spec__0_spec__1___redArg(lean_object*); +size_t lean_array_size(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_CollectStructTypes_addReshape_spec__0(lean_object*, size_t, size_t, lean_object*, lean_object*); +static lean_object* l_Lean_IR_instInhabitedStructTypeInfo_default___closed__0; +LEAN_EXPORT lean_object* l_Lean_IR_instInhabitedStructTypeInfo_default; +LEAN_EXPORT lean_object* l_Lean_IR_instHashableIRTypeApprox; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_CollectUsedDecls_collectFnBody_spec__0(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__2___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_IR_IRType_normalizeObject(lean_object*); lean_object* lean_array_get_size(lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___00__private_Std_Data_DHashMap_Internal_Defs_0__Std_DHashMap_Internal_Raw_u2080_expand_go___at___00Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__3_spec__4_spec__6(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectMaps_collectFnBody_spec__0_spec__2___redArg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_CollectUsedDecls_collectFnBody(lean_object*, lean_object*, lean_object*); +uint8_t lean_usize_dec_lt(size_t, size_t); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_IRType_hashApprox_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint64_t l_Lean_IR_instHashableVarId_hash(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_collectUsedDecls(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_instInhabitedStructTypeInfo_default___closed__1; uint8_t l_List_any___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_CollectUsedDecls_collect(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectMaps_collectParams_spec__0_spec__1(lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_CollectStructTypes_registerType_spec__0___redArg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Std_Data_DHashMap_Internal_Defs_0__Std_DHashMap_Internal_Raw_u2080_expand_go___at___00Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectMaps_collectFnBody_spec__0_spec__1_spec__2(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_CollectStructTypes_addReshape___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_CollectMaps_collectVar___closed__0; +LEAN_EXPORT uint8_t l_Array_contains___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +LEAN_EXPORT uint64_t l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_IRType_hashApprox_spec__0(lean_object*, size_t, size_t, uint64_t); LEAN_EXPORT lean_object* l___private_Std_Data_DHashMap_Internal_Defs_0__Std_DHashMap_Internal_Raw_u2080_expand_go___at___00Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectMaps_collectParams_spec__0_spec__1_spec__2(lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_land(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_IR_CollectUsedDecls_collectFnBody___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_usesModuleFrom___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectMaps_collectFnBody_spec__0_spec__0___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_CollectMaps_collectParams___boxed(lean_object*, lean_object*); @@ -227,6 +324,2290 @@ x_4 = lean_box(x_3); return x_4; } } +static uint64_t _init_l_Lean_IR_IRType_hashApprox___closed__0() { +_start: +{ +uint64_t x_1; uint64_t x_2; uint64_t x_3; +x_1 = 7; +x_2 = 43; +x_3 = lean_uint64_mix_hash(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT uint64_t l_Lean_IR_IRType_hashApprox(lean_object* x_1) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 0: +{ +uint64_t x_2; +x_2 = 11; +return x_2; +} +case 1: +{ +uint64_t x_3; +x_3 = 13; +return x_3; +} +case 2: +{ +uint64_t x_4; +x_4 = 17; +return x_4; +} +case 3: +{ +uint64_t x_5; +x_5 = 19; +return x_5; +} +case 4: +{ +uint64_t x_6; +x_6 = 23; +return x_6; +} +case 5: +{ +uint64_t x_7; +x_7 = 29; +return x_7; +} +case 9: +{ +uint64_t x_8; +x_8 = 31; +return x_8; +} +case 10: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint64_t x_12; uint64_t x_13; uint64_t x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_9 = lean_ctor_get(x_1, 1); +x_10 = lean_ctor_get(x_1, 2); +x_11 = lean_ctor_get(x_1, 3); +x_12 = 41; +x_20 = 7; +x_21 = lean_unsigned_to_nat(0u); +x_22 = lean_array_get_size(x_9); +x_23 = lean_nat_dec_lt(x_21, x_22); +if (x_23 == 0) +{ +x_13 = x_20; +goto block_19; +} +else +{ +uint8_t x_24; +x_24 = lean_nat_dec_le(x_22, x_22); +if (x_24 == 0) +{ +if (x_23 == 0) +{ +x_13 = x_20; +goto block_19; +} +else +{ +size_t x_25; size_t x_26; uint64_t x_27; +x_25 = 0; +x_26 = lean_usize_of_nat(x_22); +x_27 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_IRType_hashApprox_spec__0(x_9, x_25, x_26, x_20); +x_13 = x_27; +goto block_19; +} +} +else +{ +size_t x_28; size_t x_29; uint64_t x_30; +x_28 = 0; +x_29 = lean_usize_of_nat(x_22); +x_30 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_IRType_hashApprox_spec__0(x_9, x_28, x_29, x_20); +x_13 = x_30; +goto block_19; +} +} +block_19: +{ +uint64_t x_14; uint64_t x_15; uint64_t x_16; uint64_t x_17; uint64_t x_18; +x_14 = lean_uint64_mix_hash(x_12, x_13); +x_15 = lean_uint64_of_nat(x_10); +x_16 = lean_uint64_mix_hash(x_14, x_15); +x_17 = lean_uint64_of_nat(x_11); +x_18 = lean_uint64_mix_hash(x_16, x_17); +return x_18; +} +} +case 11: +{ +lean_object* x_31; uint64_t x_32; uint64_t x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_31 = lean_ctor_get(x_1, 1); +x_32 = 43; +x_33 = 7; +x_34 = lean_unsigned_to_nat(0u); +x_35 = lean_array_get_size(x_31); +x_36 = lean_nat_dec_lt(x_34, x_35); +if (x_36 == 0) +{ +uint64_t x_37; +x_37 = l_Lean_IR_IRType_hashApprox___closed__0; +return x_37; +} +else +{ +uint8_t x_38; +x_38 = lean_nat_dec_le(x_35, x_35); +if (x_38 == 0) +{ +if (x_36 == 0) +{ +uint64_t x_39; +x_39 = l_Lean_IR_IRType_hashApprox___closed__0; +return x_39; +} +else +{ +size_t x_40; size_t x_41; uint64_t x_42; uint64_t x_43; +x_40 = 0; +x_41 = lean_usize_of_nat(x_35); +x_42 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_IRType_hashApprox_spec__0(x_31, x_40, x_41, x_33); +x_43 = lean_uint64_mix_hash(x_32, x_42); +return x_43; +} +} +else +{ +size_t x_44; size_t x_45; uint64_t x_46; uint64_t x_47; +x_44 = 0; +x_45 = lean_usize_of_nat(x_35); +x_46 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_IRType_hashApprox_spec__0(x_31, x_44, x_45, x_33); +x_47 = lean_uint64_mix_hash(x_32, x_46); +return x_47; +} +} +} +default: +{ +uint64_t x_48; +x_48 = 37; +return x_48; +} +} +} +} +LEAN_EXPORT uint64_t l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_IRType_hashApprox_spec__0(lean_object* x_1, size_t x_2, size_t x_3, uint64_t x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_2, x_3); +if (x_5 == 0) +{ +lean_object* x_6; 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_Lean_IR_IRType_hashApprox(x_6); +lean_dec(x_6); +x_8 = lean_uint64_mix_hash(x_4, x_7); +x_9 = 1; +x_10 = lean_usize_add(x_2, x_9); +x_2 = x_10; +x_4 = x_8; +goto _start; +} +else +{ +return x_4; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_IRType_hashApprox_spec__0___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; +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +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___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_IRType_hashApprox_spec__0(x_1, x_5, x_6, x_7); +lean_dec_ref(x_1); +x_9 = lean_box_uint64(x_8); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_IRType_hashApprox___boxed(lean_object* x_1) { +_start: +{ +uint64_t x_2; lean_object* x_3; +x_2 = l_Lean_IR_IRType_hashApprox(x_1); +lean_dec(x_1); +x_3 = lean_box_uint64(x_2); +return x_3; +} +} +LEAN_EXPORT uint8_t l_Lean_IR_instBEqIRTypeApprox___lam__0(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; uint8_t x_4; +x_3 = 1; +x_4 = l_Lean_IR_IRType_compatibleWith(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_instBEqIRTypeApprox___lam__0___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_IR_instBEqIRTypeApprox___lam__0(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_IR_instBEqIRTypeApprox() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_IR_instBEqIRTypeApprox___lam__0___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_instHashableIRTypeApprox___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_IR_IRType_hashApprox___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_instHashableIRTypeApprox() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_IR_instHashableIRTypeApprox___closed__0; +return x_1; +} +} +static lean_object* _init_l_Lean_IR_instInhabitedStructTypeInfo_default___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = l_Array_empty(lean_box(0)); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_instInhabitedStructTypeInfo_default___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_instInhabitedStructTypeInfo_default___closed__0; +x_2 = lean_box(0); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_IR_instInhabitedStructTypeInfo_default() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_IR_instInhabitedStructTypeInfo_default___closed__1; +return x_1; +} +} +static lean_object* _init_l_Lean_IR_instInhabitedStructTypeInfo() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_IR_instInhabitedStructTypeInfo_default; +return x_1; +} +} +LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__2___redArg(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; uint8_t x_7; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = 1; +x_7 = l_Lean_IR_IRType_compatibleWith(x_4, x_1, x_6); +if (x_7 == 0) +{ +x_2 = x_5; +goto _start; +} +else +{ +return x_7; +} +} +} +} +LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = l_Std_DHashMap_Internal_AssocList_contains___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__2___redArg(x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___00__private_Std_Data_DHashMap_Internal_Defs_0__Std_DHashMap_Internal_Raw_u2080_expand_go___at___00Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__3_spec__4_spec__6___redArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +return x_1; +} +else +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint64_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_array_get_size(x_1); +x_7 = l_Lean_IR_IRType_hashApprox(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); +x_16 = 1; +x_17 = lean_usize_sub(x_15, x_16); +x_18 = lean_usize_land(x_14, x_17); +x_19 = lean_array_uget(x_1, x_18); +lean_ctor_set(x_2, 2, x_19); +x_20 = lean_array_uset(x_1, x_18, x_2); +x_1 = x_20; +x_2 = x_5; +goto _start; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint64_t x_26; uint64_t x_27; uint64_t x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; size_t x_33; size_t x_34; size_t x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_22 = lean_ctor_get(x_2, 0); +x_23 = lean_ctor_get(x_2, 1); +x_24 = lean_ctor_get(x_2, 2); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_2); +x_25 = lean_array_get_size(x_1); +x_26 = l_Lean_IR_IRType_hashApprox(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); +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_AssocList_foldlM___at___00__private_Std_Data_DHashMap_Internal_Defs_0__Std_DHashMap_Internal_Raw_u2080_expand_go___at___00Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__3_spec__4_spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_DHashMap_Internal_AssocList_foldlM___at___00__private_Std_Data_DHashMap_Internal_Defs_0__Std_DHashMap_Internal_Raw_u2080_expand_go___at___00Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__3_spec__4_spec__6___redArg(x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_CollectStructTypes_registerType_spec__0_spec__0___redArg(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; uint8_t x_8; +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 = 1; +x_8 = l_Lean_IR_IRType_compatibleWith(x_4, x_1, x_7); +if (x_8 == 0) +{ +x_2 = x_6; +goto _start; +} +else +{ +lean_object* x_10; +lean_inc(x_5); +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_5); +return x_10; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_CollectStructTypes_registerType_spec__0_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_CollectStructTypes_registerType_spec__0_spec__0___redArg(x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__4___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +else +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; uint8_t x_9; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_3, 1); +x_7 = lean_ctor_get(x_3, 2); +x_8 = 1; +x_9 = l_Lean_IR_IRType_compatibleWith(x_5, x_1, x_8); +if (x_9 == 0) +{ +lean_object* x_10; +x_10 = l_Std_DHashMap_Internal_AssocList_replace___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__4___redArg(x_1, x_2, x_7); +lean_ctor_set(x_3, 2, x_10); +return x_3; +} +else +{ +lean_dec(x_6); +lean_dec(x_5); +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; uint8_t x_15; +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 = 1; +x_15 = l_Lean_IR_IRType_compatibleWith(x_11, x_1, x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; +x_16 = l_Std_DHashMap_Internal_AssocList_replace___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__4___redArg(x_1, x_2, x_13); +x_17 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_17, 0, x_11); +lean_ctor_set(x_17, 1, x_12); +lean_ctor_set(x_17, 2, x_16); +return x_17; +} +else +{ +lean_object* x_18; +lean_dec(x_12); +lean_dec(x_11); +x_18 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_18, 0, x_1); +lean_ctor_set(x_18, 1, x_2); +lean_ctor_set(x_18, 2, x_13); +return x_18; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Std_DHashMap_Internal_AssocList_replace___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__4___redArg(x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_CollectStructTypes_registerType_spec__0___redArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint64_t x_5; uint64_t x_6; uint64_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; size_t x_12; size_t x_13; size_t x_14; size_t x_15; size_t x_16; lean_object* x_17; lean_object* x_18; +x_3 = lean_ctor_get(x_1, 1); +x_4 = lean_array_get_size(x_3); +x_5 = l_Lean_IR_IRType_hashApprox(x_2); +x_6 = 32; +x_7 = lean_uint64_shift_right(x_5, x_6); +x_8 = lean_uint64_xor(x_5, x_7); +x_9 = 16; +x_10 = lean_uint64_shift_right(x_8, x_9); +x_11 = lean_uint64_xor(x_8, x_10); +x_12 = lean_uint64_to_usize(x_11); +x_13 = lean_usize_of_nat(x_4); +x_14 = 1; +x_15 = lean_usize_sub(x_13, x_14); +x_16 = lean_usize_land(x_12, x_15); +x_17 = lean_array_uget(x_3, x_16); +x_18 = l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_CollectStructTypes_registerType_spec__0_spec__0___redArg(x_2, x_17); +lean_dec(x_17); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_CollectStructTypes_registerType_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_CollectStructTypes_registerType_spec__0___redArg(x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Std_Data_DHashMap_Internal_Defs_0__Std_DHashMap_Internal_Raw_u2080_expand_go___at___00Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__3_spec__4___redArg(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); +if (x_5 == 0) +{ +lean_dec_ref(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___00__private_Std_Data_DHashMap_Internal_Defs_0__Std_DHashMap_Internal_Raw_u2080_expand_go___at___00Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__3_spec__4_spec__6___redArg(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___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__3___redArg(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); +x_5 = lean_unsigned_to_nat(0u); +x_6 = lean_box(0); +x_7 = lean_mk_array(x_4, x_6); +x_8 = l___private_Std_Data_DHashMap_Internal_Defs_0__Std_DHashMap_Internal_Raw_u2080_expand_go___at___00Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__3_spec__4___redArg(x_5, x_1, x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__3(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__3___redArg(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; uint64_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; size_t x_19; lean_object* x_20; uint8_t x_21; +x_5 = lean_ctor_get(x_1, 0); +x_6 = lean_ctor_get(x_1, 1); +x_7 = lean_array_get_size(x_6); +x_8 = l_Lean_IR_IRType_hashApprox(x_2); +x_9 = 32; +x_10 = lean_uint64_shift_right(x_8, x_9); +x_11 = lean_uint64_xor(x_8, x_10); +x_12 = 16; +x_13 = lean_uint64_shift_right(x_11, x_12); +x_14 = lean_uint64_xor(x_11, x_13); +x_15 = lean_uint64_to_usize(x_14); +x_16 = lean_usize_of_nat(x_7); +x_17 = 1; +x_18 = lean_usize_sub(x_16, x_17); +x_19 = lean_usize_land(x_15, x_18); +x_20 = lean_array_uget(x_6, x_19); +x_21 = l_Std_DHashMap_Internal_AssocList_contains___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__2___redArg(x_2, x_20); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_22 = lean_unsigned_to_nat(1u); +x_23 = lean_nat_add(x_5, x_22); +lean_dec(x_5); +x_24 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_24, 0, x_2); +lean_ctor_set(x_24, 1, x_3); +lean_ctor_set(x_24, 2, x_20); +x_25 = lean_array_uset(x_6, x_19, x_24); +x_26 = lean_unsigned_to_nat(4u); +x_27 = lean_nat_mul(x_23, x_26); +x_28 = lean_unsigned_to_nat(3u); +x_29 = lean_nat_div(x_27, x_28); +lean_dec(x_27); +x_30 = lean_array_get_size(x_25); +x_31 = lean_nat_dec_le(x_29, x_30); +lean_dec(x_29); +if (x_31 == 0) +{ +lean_object* x_32; +x_32 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__3___redArg(x_25); +lean_ctor_set(x_1, 1, x_32); +lean_ctor_set(x_1, 0, x_23); +return x_1; +} +else +{ +lean_ctor_set(x_1, 1, x_25); +lean_ctor_set(x_1, 0, x_23); +return x_1; +} +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_33 = lean_box(0); +x_34 = lean_array_uset(x_6, x_19, x_33); +x_35 = l_Std_DHashMap_Internal_AssocList_replace___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__4___redArg(x_2, x_3, x_20); +x_36 = lean_array_uset(x_34, x_19, x_35); +lean_ctor_set(x_1, 1, x_36); +return x_1; +} +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; uint64_t x_40; uint64_t x_41; uint64_t x_42; uint64_t x_43; uint64_t x_44; uint64_t x_45; uint64_t x_46; size_t x_47; size_t x_48; size_t x_49; size_t x_50; size_t x_51; lean_object* x_52; uint8_t x_53; +x_37 = lean_ctor_get(x_1, 0); +x_38 = lean_ctor_get(x_1, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_1); +x_39 = lean_array_get_size(x_38); +x_40 = l_Lean_IR_IRType_hashApprox(x_2); +x_41 = 32; +x_42 = lean_uint64_shift_right(x_40, x_41); +x_43 = lean_uint64_xor(x_40, x_42); +x_44 = 16; +x_45 = lean_uint64_shift_right(x_43, x_44); +x_46 = lean_uint64_xor(x_43, x_45); +x_47 = lean_uint64_to_usize(x_46); +x_48 = lean_usize_of_nat(x_39); +x_49 = 1; +x_50 = lean_usize_sub(x_48, x_49); +x_51 = lean_usize_land(x_47, x_50); +x_52 = lean_array_uget(x_38, x_51); +x_53 = l_Std_DHashMap_Internal_AssocList_contains___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__2___redArg(x_2, x_52); +if (x_53 == 0) +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; +x_54 = lean_unsigned_to_nat(1u); +x_55 = lean_nat_add(x_37, x_54); +lean_dec(x_37); +x_56 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_56, 0, x_2); +lean_ctor_set(x_56, 1, x_3); +lean_ctor_set(x_56, 2, x_52); +x_57 = lean_array_uset(x_38, x_51, x_56); +x_58 = lean_unsigned_to_nat(4u); +x_59 = lean_nat_mul(x_55, x_58); +x_60 = lean_unsigned_to_nat(3u); +x_61 = lean_nat_div(x_59, x_60); +lean_dec(x_59); +x_62 = lean_array_get_size(x_57); +x_63 = lean_nat_dec_le(x_61, x_62); +lean_dec(x_61); +if (x_63 == 0) +{ +lean_object* x_64; lean_object* x_65; +x_64 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__3___redArg(x_57); +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_55); +lean_ctor_set(x_65, 1, x_64); +return x_65; +} +else +{ +lean_object* x_66; +x_66 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_66, 0, x_55); +lean_ctor_set(x_66, 1, x_57); +return x_66; +} +} +else +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_67 = lean_box(0); +x_68 = lean_array_uset(x_38, x_51, x_67); +x_69 = l_Std_DHashMap_Internal_AssocList_replace___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__4___redArg(x_2, x_3, x_52); +x_70 = lean_array_uset(x_68, x_51, x_69); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_37); +lean_ctor_set(x_71, 1, x_70); +return x_71; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1___redArg(x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Std_Data_DHashMap_Internal_Defs_0__Std_DHashMap_Internal_Raw_u2080_expand_go___at___00Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__3_spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l___private_Std_Data_DHashMap_Internal_Defs_0__Std_DHashMap_Internal_Raw_u2080_expand_go___at___00Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__3_spec__4___redArg(x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_IR_CollectStructTypes_registerType___closed__0() { +_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_IR_CollectStructTypes_registerType(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_30; lean_object* x_33; +switch (lean_obj_tag(x_1)) { +case 10: +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; +x_36 = lean_ctor_get(x_1, 1); +x_37 = lean_unsigned_to_nat(0u); +x_38 = lean_array_get_size(x_36); +x_39 = lean_nat_dec_lt(x_37, x_38); +if (x_39 == 0) +{ +x_3 = x_2; +goto block_29; +} +else +{ +lean_object* x_40; uint8_t x_41; +x_40 = lean_box(0); +x_41 = lean_nat_dec_le(x_38, x_38); +if (x_41 == 0) +{ +if (x_39 == 0) +{ +x_3 = x_2; +goto block_29; +} +else +{ +size_t x_42; size_t x_43; lean_object* x_44; +x_42 = 0; +x_43 = lean_usize_of_nat(x_38); +x_44 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_CollectStructTypes_registerType_spec__2(x_36, x_42, x_43, x_40, x_2); +x_30 = x_44; +goto block_32; +} +} +else +{ +size_t x_45; size_t x_46; lean_object* x_47; +x_45 = 0; +x_46 = lean_usize_of_nat(x_38); +x_47 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_CollectStructTypes_registerType_spec__2(x_36, x_45, x_46, x_40, x_2); +x_30 = x_47; +goto block_32; +} +} +} +case 11: +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; +x_48 = lean_ctor_get(x_1, 1); +x_49 = lean_unsigned_to_nat(0u); +x_50 = lean_array_get_size(x_48); +x_51 = lean_nat_dec_lt(x_49, x_50); +if (x_51 == 0) +{ +x_3 = x_2; +goto block_29; +} +else +{ +lean_object* x_52; uint8_t x_53; +x_52 = lean_box(0); +x_53 = lean_nat_dec_le(x_50, x_50); +if (x_53 == 0) +{ +if (x_51 == 0) +{ +x_3 = x_2; +goto block_29; +} +else +{ +size_t x_54; size_t x_55; lean_object* x_56; +x_54 = 0; +x_55 = lean_usize_of_nat(x_50); +x_56 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_CollectStructTypes_registerType_spec__2(x_48, x_54, x_55, x_52, x_2); +x_33 = x_56; +goto block_35; +} +} +else +{ +size_t x_57; size_t x_58; lean_object* x_59; +x_57 = 0; +x_58 = lean_usize_of_nat(x_50); +x_59 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_CollectStructTypes_registerType_spec__2(x_48, x_57, x_58, x_52, x_2); +x_33 = x_59; +goto block_35; +} +} +} +default: +{ +lean_object* x_60; lean_object* x_61; +lean_dec(x_1); +x_60 = lean_box(0); +x_61 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_2); +return x_61; +} +} +block_29: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_4 = lean_ctor_get(x_3, 0); +x_5 = lean_ctor_get(x_3, 1); +x_6 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_CollectStructTypes_registerType_spec__0___redArg(x_5, x_1); +if (lean_obj_tag(x_6) == 0) +{ +uint8_t x_7; +lean_inc(x_5); +lean_inc(x_4); +x_7 = !lean_is_exclusive(x_3); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_8 = lean_ctor_get(x_3, 1); +lean_dec(x_8); +x_9 = lean_ctor_get(x_3, 0); +lean_dec(x_9); +x_10 = lean_array_get_size(x_4); +x_11 = l_Lean_IR_IRType_normalizeObject(x_1); +x_12 = lean_box(0); +x_13 = l_Lean_IR_CollectStructTypes_registerType___closed__0; +lean_inc(x_11); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_11); +lean_ctor_set(x_14, 1, x_13); +x_15 = lean_array_push(x_4, x_14); +x_16 = l_Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1___redArg(x_5, x_11, x_10); +lean_ctor_set(x_3, 1, x_16); +lean_ctor_set(x_3, 0, x_15); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_12); +lean_ctor_set(x_17, 1, x_3); +return x_17; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_dec(x_3); +x_18 = lean_array_get_size(x_4); +x_19 = l_Lean_IR_IRType_normalizeObject(x_1); +x_20 = lean_box(0); +x_21 = l_Lean_IR_CollectStructTypes_registerType___closed__0; +lean_inc(x_19); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_19); +lean_ctor_set(x_22, 1, x_21); +x_23 = lean_array_push(x_4, x_22); +x_24 = l_Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1___redArg(x_5, x_19, x_18); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_20); +lean_ctor_set(x_26, 1, x_25); +return x_26; +} +} +else +{ +lean_object* x_27; lean_object* x_28; +lean_dec_ref(x_6); +lean_dec(x_1); +x_27 = lean_box(0); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_3); +return x_28; +} +} +block_32: +{ +lean_object* x_31; +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec_ref(x_30); +x_3 = x_31; +goto block_29; +} +block_35: +{ +lean_object* x_34; +x_34 = lean_ctor_get(x_33, 1); +lean_inc(x_34); +lean_dec_ref(x_33); +x_3 = x_34; +goto block_29; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_CollectStructTypes_registerType_spec__2(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_2, x_3); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; size_t x_11; size_t x_12; +x_7 = lean_array_uget(x_1, x_2); +x_8 = l_Lean_IR_CollectStructTypes_registerType(x_7, x_5); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +lean_dec_ref(x_8); +x_11 = 1; +x_12 = lean_usize_add(x_2, x_11); +x_2 = x_12; +x_4 = x_9; +x_5 = x_10; +goto _start; +} +else +{ +lean_object* x_14; +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_4); +lean_ctor_set(x_14, 1, x_5); +return x_14; +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = l_Std_DHashMap_Internal_AssocList_contains___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__2(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +x_5 = lean_box(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_CollectStructTypes_registerType_spec__0_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_CollectStructTypes_registerType_spec__0_spec__0(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_CollectStructTypes_registerType_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_CollectStructTypes_registerType_spec__0(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec_ref(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__2___redArg___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___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_CollectStructTypes_registerType_spec__1_spec__2___redArg(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_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_CollectStructTypes_registerType_spec__0_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_CollectStructTypes_registerType_spec__0_spec__0___redArg(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_CollectStructTypes_registerType_spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_7 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_8 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_CollectStructTypes_registerType_spec__2(x_1, x_6, x_7, x_4, x_5); +lean_dec_ref(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_CollectStructTypes_registerType_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_CollectStructTypes_registerType_spec__0___redArg(x_1, x_2); +lean_dec(x_2); +lean_dec_ref(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___redArg___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("key is not present in hash table", 32, 32); +return x_1; +} +} +static lean_object* _init_l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___redArg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.DHashMap.Internal.AssocList.get!", 36, 36); +return x_1; +} +} +static lean_object* _init_l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___redArg___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Data.DHashMap.Internal.AssocList.Basic", 42, 42); +return x_1; +} +} +static lean_object* _init_l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___redArg___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_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___redArg___closed__2; +x_2 = lean_unsigned_to_nat(11u); +x_3 = lean_unsigned_to_nat(163u); +x_4 = l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___redArg___closed__1; +x_5 = l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___redArg___closed__0; +x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___redArg(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_object* x_5; +x_4 = l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___redArg___closed__3; +x_5 = lean_panic_fn(x_1, x_4); +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; uint8_t x_10; +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 = 1; +x_10 = l_Lean_IR_IRType_compatibleWith(x_6, x_2, x_9); +if (x_10 == 0) +{ +x_3 = x_8; +goto _start; +} +else +{ +lean_dec(x_1); +lean_inc(x_7); +return x_7; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___redArg(x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; uint64_t 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; size_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_array_get_size(x_4); +x_6 = l_Lean_IR_IRType_hashApprox(x_3); +x_7 = 32; +x_8 = lean_uint64_shift_right(x_6, x_7); +x_9 = lean_uint64_xor(x_6, x_8); +x_10 = 16; +x_11 = lean_uint64_shift_right(x_9, x_10); +x_12 = lean_uint64_xor(x_9, x_11); +x_13 = lean_uint64_to_usize(x_12); +x_14 = lean_usize_of_nat(x_5); +x_15 = 1; +x_16 = lean_usize_sub(x_14, x_15); +x_17 = lean_usize_land(x_13, x_16); +x_18 = lean_array_uget(x_4, x_17); +x_19 = l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___redArg(x_1, x_3, x_18); +lean_dec(x_18); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0___redArg(x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_panic___at___00Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0_spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_panic_fn(x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_panic___at___00Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0_spec__1___redArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_panic_fn(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT uint8_t l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Array_contains___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__1_spec__2(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_3, x_4); +if (x_5 == 0) +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_array_uget(x_2, x_3); +x_7 = lean_nat_dec_eq(x_1, x_6); +lean_dec(x_6); +if (x_7 == 0) +{ +size_t x_8; size_t x_9; +x_8 = 1; +x_9 = lean_usize_add(x_3, x_8); +x_3 = x_9; +goto _start; +} +else +{ +return x_7; +} +} +else +{ +uint8_t x_11; +x_11 = 0; +return x_11; +} +} +} +LEAN_EXPORT uint8_t l_Array_contains___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = lean_unsigned_to_nat(0u); +x_4 = lean_array_get_size(x_1); +x_5 = lean_nat_dec_lt(x_3, x_4); +if (x_5 == 0) +{ +return x_5; +} +else +{ +if (x_5 == 0) +{ +return x_5; +} +else +{ +size_t x_6; size_t x_7; uint8_t x_8; +x_6 = 0; +x_7 = lean_usize_of_nat(x_4); +x_8 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Array_contains___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__1_spec__2(x_2, x_1, x_6, x_7); +return x_8; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_CollectStructTypes_addReshapeEntry(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_14; uint8_t x_15; +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_3, 1); +lean_inc(x_5); +if (lean_is_exclusive(x_3)) { + lean_ctor_release(x_3, 0); + lean_ctor_release(x_3, 1); + x_6 = x_3; +} else { + lean_dec_ref(x_3); + x_6 = lean_box(0); +} +x_7 = lean_unsigned_to_nat(0u); +x_8 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0___redArg(x_7, x_5, x_1); +x_9 = lean_box(0); +x_14 = lean_array_get_size(x_4); +x_15 = lean_nat_dec_lt(x_8, x_14); +if (x_15 == 0) +{ +lean_dec(x_8); +x_10 = x_4; +goto block_13; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_24; +x_16 = lean_array_fget(x_4, x_8); +x_17 = lean_ctor_get(x_16, 0); +x_18 = lean_ctor_get(x_16, 1); +x_19 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0___redArg(x_7, x_5, x_2); +x_20 = lean_array_fset(x_4, x_8, x_9); +x_24 = l_Array_contains___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__1(x_18, x_19); +if (x_24 == 0) +{ +uint8_t x_25; +lean_inc_ref(x_18); +lean_inc(x_17); +x_25 = !lean_is_exclusive(x_16); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_16, 1); +lean_dec(x_26); +x_27 = lean_ctor_get(x_16, 0); +lean_dec(x_27); +x_28 = lean_array_push(x_18, x_19); +lean_ctor_set(x_16, 1, x_28); +x_21 = x_16; +goto block_23; +} +else +{ +lean_object* x_29; lean_object* x_30; +lean_dec(x_16); +x_29 = lean_array_push(x_18, x_19); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_17); +lean_ctor_set(x_30, 1, x_29); +x_21 = x_30; +goto block_23; +} +} +else +{ +lean_dec(x_19); +x_21 = x_16; +goto block_23; +} +block_23: +{ +lean_object* x_22; +x_22 = lean_array_fset(x_20, x_8, x_21); +lean_dec(x_8); +x_10 = x_22; +goto block_13; +} +} +block_13: +{ +lean_object* x_11; lean_object* x_12; +if (lean_is_scalar(x_6)) { + x_11 = lean_alloc_ctor(0, 2, 0); +} else { + x_11 = x_6; +} +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_5); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_9); +lean_ctor_set(x_12, 1, x_11); +return x_12; +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0(x_1, x_2, x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0(x_1, x_2, x_3, x_4); +lean_dec(x_4); +lean_dec_ref(x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Array_contains___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Array_contains___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__1(x_1, x_2); +lean_dec(x_2); +lean_dec_ref(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Array_contains___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__1_spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; +x_5 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_6 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_7 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Array_contains___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__1_spec__2(x_1, x_2, x_5, x_6); +lean_dec_ref(x_2); +lean_dec(x_1); +x_8 = lean_box(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___redArg(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0___redArg(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec_ref(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_CollectStructTypes_addReshapeEntry___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_IR_CollectStructTypes_addReshapeEntry(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_CollectStructTypes_addReshape(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +switch (lean_obj_tag(x_1)) { +case 10: +{ +if (lean_obj_tag(x_2) == 10) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; size_t x_13; size_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_8 = lean_ctor_get(x_1, 1); +x_9 = lean_ctor_get(x_2, 1); +x_10 = lean_unsigned_to_nat(0u); +x_11 = lean_array_get_size(x_9); +lean_inc_ref(x_9); +x_12 = l_Array_toSubarray___redArg(x_9, x_10, x_11); +x_13 = lean_array_size(x_8); +x_14 = 0; +x_15 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_CollectStructTypes_addReshape_spec__0(x_8, x_13, x_14, x_12, x_3); +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec_ref(x_15); +x_17 = l_Lean_IR_CollectStructTypes_addReshapeEntry(x_1, x_2, x_16); +lean_dec_ref(x_2); +return x_17; +} +else +{ +lean_dec(x_2); +x_4 = x_3; +goto block_7; +} +} +case 11: +{ +if (lean_obj_tag(x_2) == 11) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; size_t x_23; size_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_18 = lean_ctor_get(x_1, 1); +x_19 = lean_ctor_get(x_2, 1); +x_20 = lean_unsigned_to_nat(0u); +x_21 = lean_array_get_size(x_19); +lean_inc_ref(x_19); +x_22 = l_Array_toSubarray___redArg(x_19, x_20, x_21); +x_23 = lean_array_size(x_18); +x_24 = 0; +x_25 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_CollectStructTypes_addReshape_spec__0(x_18, x_23, x_24, x_22, x_3); +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec_ref(x_25); +x_27 = l_Lean_IR_CollectStructTypes_addReshapeEntry(x_1, x_2, x_26); +lean_dec_ref(x_2); +return x_27; +} +else +{ +lean_dec(x_2); +x_4 = x_3; +goto block_7; +} +} +default: +{ +lean_dec(x_2); +x_4 = x_3; +goto block_7; +} +} +block_7: +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_box(0); +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_4); +return x_6; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_CollectStructTypes_addReshape_spec__0(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_usize_dec_lt(x_3, x_2); +if (x_6 == 0) +{ +lean_object* x_7; +x_7 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_7, 0, x_4); +lean_ctor_set(x_7, 1, x_5); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_ctor_get(x_4, 0); +x_9 = lean_ctor_get(x_4, 1); +x_10 = lean_ctor_get(x_4, 2); +x_11 = lean_nat_dec_lt(x_9, x_10); +if (x_11 == 0) +{ +lean_object* x_12; +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_4); +lean_ctor_set(x_12, 1, x_5); +return x_12; +} +else +{ +uint8_t x_13; +lean_inc(x_10); +lean_inc(x_9); +lean_inc_ref(x_8); +x_13 = !lean_is_exclusive(x_4); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; size_t x_23; size_t x_24; +x_14 = lean_ctor_get(x_4, 2); +lean_dec(x_14); +x_15 = lean_ctor_get(x_4, 1); +lean_dec(x_15); +x_16 = lean_ctor_get(x_4, 0); +lean_dec(x_16); +x_17 = lean_array_uget(x_1, x_3); +x_18 = lean_array_fget_borrowed(x_8, x_9); +lean_inc(x_18); +x_19 = l_Lean_IR_CollectStructTypes_addReshape(x_17, x_18, x_5); +lean_dec(x_17); +x_20 = lean_ctor_get(x_19, 1); +lean_inc(x_20); +lean_dec_ref(x_19); +x_21 = lean_unsigned_to_nat(1u); +x_22 = lean_nat_add(x_9, x_21); +lean_dec(x_9); +lean_ctor_set(x_4, 1, x_22); +x_23 = 1; +x_24 = lean_usize_add(x_3, x_23); +x_3 = x_24; +x_5 = x_20; +goto _start; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; size_t x_33; size_t x_34; +lean_dec(x_4); +x_26 = lean_array_uget(x_1, x_3); +x_27 = lean_array_fget_borrowed(x_8, x_9); +lean_inc(x_27); +x_28 = l_Lean_IR_CollectStructTypes_addReshape(x_26, x_27, x_5); +lean_dec(x_26); +x_29 = lean_ctor_get(x_28, 1); +lean_inc(x_29); +lean_dec_ref(x_28); +x_30 = lean_unsigned_to_nat(1u); +x_31 = lean_nat_add(x_9, x_30); +lean_dec(x_9); +x_32 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_32, 0, x_8); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set(x_32, 2, x_10); +x_33 = 1; +x_34 = lean_usize_add(x_3, x_33); +x_3 = x_34; +x_4 = x_32; +x_5 = x_29; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_CollectStructTypes_addReshape_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_7 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_8 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_CollectStructTypes_addReshape_spec__0(x_1, x_6, x_7, x_4, x_5); +lean_dec_ref(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_CollectStructTypes_addReshape___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_IR_CollectStructTypes_addReshape(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_EmitUtil_0__Lean_IR_CollectStructTypes_addReshape_match__3_splitter(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +switch (lean_obj_tag(x_2)) { +case 10: +{ +lean_dec(x_5); +if (lean_obj_tag(x_3) == 10) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_dec(x_6); +x_7 = lean_ctor_get(x_2, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_2, 1); +lean_inc_ref(x_8); +x_9 = lean_ctor_get(x_2, 2); +lean_inc(x_9); +x_10 = lean_ctor_get(x_2, 3); +lean_inc(x_10); +lean_dec_ref(x_2); +x_11 = lean_ctor_get(x_3, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_3, 1); +lean_inc_ref(x_12); +x_13 = lean_ctor_get(x_3, 2); +lean_inc(x_13); +x_14 = lean_ctor_get(x_3, 3); +lean_inc(x_14); +lean_dec_ref(x_3); +x_15 = lean_apply_8(x_4, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_15; +} +else +{ +lean_object* x_16; +lean_dec(x_4); +x_16 = lean_apply_4(x_6, x_2, x_3, lean_box(0), lean_box(0)); +return x_16; +} +} +case 11: +{ +lean_dec(x_4); +if (lean_obj_tag(x_3) == 11) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +lean_dec(x_6); +x_17 = lean_ctor_get(x_2, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_2, 1); +lean_inc_ref(x_18); +lean_dec_ref(x_2); +x_19 = lean_ctor_get(x_3, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_3, 1); +lean_inc_ref(x_20); +lean_dec_ref(x_3); +x_21 = lean_apply_4(x_5, x_17, x_18, x_19, x_20); +return x_21; +} +else +{ +lean_object* x_22; +lean_dec(x_5); +x_22 = lean_apply_4(x_6, x_2, x_3, lean_box(0), lean_box(0)); +return x_22; +} +} +default: +{ +lean_object* x_23; +lean_dec(x_5); +lean_dec(x_4); +x_23 = lean_apply_4(x_6, x_2, x_3, lean_box(0), lean_box(0)); +return x_23; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_EmitUtil_0__Lean_IR_CollectStructTypes_addReshape_match__3_splitter___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 10: +{ +lean_dec(x_4); +if (lean_obj_tag(x_2) == 10) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +lean_dec(x_5); +x_6 = lean_ctor_get(x_1, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_7); +x_8 = lean_ctor_get(x_1, 2); +lean_inc(x_8); +x_9 = lean_ctor_get(x_1, 3); +lean_inc(x_9); +lean_dec_ref(x_1); +x_10 = lean_ctor_get(x_2, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_2, 1); +lean_inc_ref(x_11); +x_12 = lean_ctor_get(x_2, 2); +lean_inc(x_12); +x_13 = lean_ctor_get(x_2, 3); +lean_inc(x_13); +lean_dec_ref(x_2); +x_14 = lean_apply_8(x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_14; +} +else +{ +lean_object* x_15; +lean_dec(x_3); +x_15 = lean_apply_4(x_5, x_1, x_2, lean_box(0), lean_box(0)); +return x_15; +} +} +case 11: +{ +lean_dec(x_3); +if (lean_obj_tag(x_2) == 11) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +lean_dec(x_5); +x_16 = lean_ctor_get(x_1, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_17); +lean_dec_ref(x_1); +x_18 = lean_ctor_get(x_2, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_2, 1); +lean_inc_ref(x_19); +lean_dec_ref(x_2); +x_20 = lean_apply_4(x_4, x_16, x_17, x_18, x_19); +return x_20; +} +else +{ +lean_object* x_21; +lean_dec(x_4); +x_21 = lean_apply_4(x_5, x_1, x_2, lean_box(0), lean_box(0)); +return x_21; +} +} +default: +{ +lean_object* x_22; +lean_dec(x_4); +lean_dec(x_3); +x_22 = lean_apply_4(x_5, x_1, x_2, lean_box(0), lean_box(0)); +return x_22; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_EmitUtil_0__Lean_IR_CollectStructTypes_addReshape_match__1_splitter___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_4; lean_object* x_5; +lean_dec(x_3); +x_4 = lean_box(0); +x_5 = lean_apply_1(x_2, x_4); +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +lean_dec(x_2); +x_6 = lean_ctor_get(x_1, 0); +lean_inc(x_6); +lean_dec_ref(x_1); +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +lean_dec(x_6); +x_9 = lean_apply_2(x_3, x_7, x_8); +return x_9; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_EmitUtil_0__Lean_IR_CollectStructTypes_addReshape_match__1_splitter(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l___private_Lean_Compiler_IR_EmitUtil_0__Lean_IR_CollectStructTypes_addReshape_match__1_splitter___redArg(x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_CollectStructTypes_collectParams_spec__0(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; uint8_t x_13; +x_13 = lean_usize_dec_lt(x_4, x_3); +if (x_13 == 0) +{ +lean_object* x_14; +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_5); +lean_ctor_set(x_14, 1, x_6); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = lean_array_uget(x_2, x_4); +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +x_17 = l_Lean_IR_IRType_isStruct(x_16); +if (x_17 == 0) +{ +lean_dec(x_16); +x_7 = x_1; +x_8 = x_6; +goto block_12; +} +else +{ +lean_object* x_18; lean_object* x_19; +x_18 = l_Lean_IR_CollectStructTypes_registerType(x_16, x_6); +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +lean_dec_ref(x_18); +x_7 = x_1; +x_8 = x_19; +goto block_12; +} +} +block_12: +{ +size_t x_9; size_t x_10; +x_9 = 1; +x_10 = lean_usize_add(x_4, x_9); +x_4 = x_10; +x_5 = x_7; +x_6 = x_8; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_CollectStructTypes_collectParams(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; size_t x_4; size_t x_5; lean_object* x_6; uint8_t x_7; +x_3 = lean_box(0); +x_4 = lean_array_size(x_1); +x_5 = 0; +x_6 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_CollectStructTypes_collectParams_spec__0(x_3, x_1, x_4, x_5, x_3, x_2); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; +x_8 = lean_ctor_get(x_6, 0); +lean_dec(x_8); +lean_ctor_set(x_6, 0, x_3); +return x_6; +} +else +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_ctor_get(x_6, 1); +lean_inc(x_9); +lean_dec(x_6); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_3); +lean_ctor_set(x_10, 1, x_9); +return x_10; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_CollectStructTypes_collectParams___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_IR_CollectStructTypes_collectParams(x_1, x_2); +lean_dec_ref(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_CollectStructTypes_collectParams_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_8 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_9 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_CollectStructTypes_collectParams_spec__0(x_1, x_2, x_7, x_8, x_5, x_6); +lean_dec_ref(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_CollectStructTypes_collectFnBody(lean_object* x_1, lean_object* x_2) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 0: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc_ref(x_4); +x_5 = lean_ctor_get(x_1, 3); +lean_inc(x_5); +lean_dec_ref(x_1); +x_6 = l_Lean_IR_IRType_isStruct(x_3); +if (x_6 == 0) +{ +lean_dec_ref(x_4); +lean_dec(x_3); +x_1 = x_5; +goto _start; +} +else +{ +lean_object* x_8; +lean_inc(x_3); +x_8 = l_Lean_IR_CollectStructTypes_registerType(x_3, x_2); +if (lean_obj_tag(x_4) == 9) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_8, 1); +lean_inc(x_9); +lean_dec_ref(x_8); +x_10 = lean_ctor_get(x_4, 0); +lean_inc(x_10); +lean_dec_ref(x_4); +x_11 = l_Lean_IR_CollectStructTypes_addReshape(x_10, x_3, x_9); +lean_dec(x_10); +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +lean_dec_ref(x_11); +x_1 = x_5; +x_2 = x_12; +goto _start; +} +else +{ +lean_object* x_14; +lean_dec_ref(x_4); +lean_dec(x_3); +x_14 = lean_ctor_get(x_8, 1); +lean_inc(x_14); +lean_dec_ref(x_8); +x_1 = x_5; +x_2 = x_14; +goto _start; +} +} +} +case 1: +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_16 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_16); +x_17 = lean_ctor_get(x_1, 2); +lean_inc(x_17); +x_18 = lean_ctor_get(x_1, 3); +lean_inc(x_18); +lean_dec_ref(x_1); +x_19 = l_Lean_IR_CollectStructTypes_collectParams(x_16, x_2); +lean_dec_ref(x_16); +x_20 = lean_ctor_get(x_19, 1); +lean_inc(x_20); +lean_dec_ref(x_19); +x_21 = l_Lean_IR_CollectStructTypes_collectFnBody(x_17, x_20); +x_22 = lean_ctor_get(x_21, 1); +lean_inc(x_22); +lean_dec_ref(x_21); +x_1 = x_18; +x_2 = x_22; +goto _start; +} +case 9: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_24 = lean_ctor_get(x_1, 3); +lean_inc_ref(x_24); +lean_dec_ref(x_1); +x_25 = lean_unsigned_to_nat(0u); +x_26 = lean_array_get_size(x_24); +x_27 = lean_box(0); +x_28 = lean_nat_dec_lt(x_25, x_26); +if (x_28 == 0) +{ +lean_object* x_29; +lean_dec_ref(x_24); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_2); +return x_29; +} +else +{ +uint8_t x_30; +x_30 = lean_nat_dec_le(x_26, x_26); +if (x_30 == 0) +{ +if (x_28 == 0) +{ +lean_object* x_31; +lean_dec_ref(x_24); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_27); +lean_ctor_set(x_31, 1, x_2); +return x_31; +} +else +{ +size_t x_32; size_t x_33; lean_object* x_34; +x_32 = 0; +x_33 = lean_usize_of_nat(x_26); +x_34 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_CollectStructTypes_collectFnBody_spec__0(x_24, x_32, x_33, x_27, x_2); +lean_dec_ref(x_24); +return x_34; +} +} +else +{ +size_t x_35; size_t x_36; lean_object* x_37; +x_35 = 0; +x_36 = lean_usize_of_nat(x_26); +x_37 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_CollectStructTypes_collectFnBody_spec__0(x_24, x_35, x_36, x_27, x_2); +lean_dec_ref(x_24); +return x_37; +} +} +} +default: +{ +uint8_t x_38; +x_38 = l_Lean_IR_FnBody_isTerminal(x_1); +if (x_38 == 0) +{ +lean_object* x_39; +x_39 = l_Lean_IR_FnBody_body(x_1); +lean_dec(x_1); +x_1 = x_39; +goto _start; +} +else +{ +lean_object* x_41; lean_object* x_42; +lean_dec(x_1); +x_41 = lean_box(0); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_2); +return x_42; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_CollectStructTypes_collectFnBody_spec__0(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_2, x_3); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; +x_7 = lean_array_uget(x_1, x_2); +x_8 = l_Lean_IR_Alt_body(x_7); +lean_dec(x_7); +x_9 = l_Lean_IR_CollectStructTypes_collectFnBody(x_8, x_5); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec_ref(x_9); +x_12 = 1; +x_13 = lean_usize_add(x_2, x_12); +x_2 = x_13; +x_4 = x_10; +x_5 = x_11; +goto _start; +} +else +{ +lean_object* x_15; +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_4); +lean_ctor_set(x_15, 1, x_5); +return x_15; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_CollectStructTypes_collectFnBody_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_7 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_8 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_CollectStructTypes_collectFnBody_spec__0(x_1, x_6, x_7, x_4, x_5); +lean_dec_ref(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_CollectStructTypes_collectDecl(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_3 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +x_5 = lean_ctor_get(x_1, 3); +lean_inc(x_5); +lean_dec_ref(x_1); +x_6 = l_Lean_IR_CollectStructTypes_collectParams(x_3, x_2); +lean_dec_ref(x_3); +x_7 = lean_ctor_get(x_6, 1); +lean_inc(x_7); +lean_dec_ref(x_6); +x_8 = l_Lean_IR_CollectStructTypes_registerType(x_4, x_7); +x_9 = lean_ctor_get(x_8, 1); +lean_inc(x_9); +lean_dec_ref(x_8); +x_10 = l_Lean_IR_CollectStructTypes_collectFnBody(x_5, x_9); +return x_10; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_11 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_11); +x_12 = lean_ctor_get(x_1, 2); +lean_inc(x_12); +lean_dec_ref(x_1); +x_13 = l_Lean_IR_CollectStructTypes_collectParams(x_11, x_2); +lean_dec_ref(x_11); +x_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec_ref(x_13); +x_15 = l_Lean_IR_CollectStructTypes_registerType(x_12, x_14); +return x_15; +} +} +} +LEAN_EXPORT lean_object* l_List_forM___at___00Lean_IR_collectStructTypes_spec__0(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; lean_object* x_4; +x_3 = lean_box(0); +x_4 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +lean_dec_ref(x_1); +x_7 = l_Lean_IR_CollectStructTypes_collectDecl(x_5, x_2); +x_8 = lean_ctor_get(x_7, 1); +lean_inc(x_8); +lean_dec_ref(x_7); +x_1 = x_6; +x_2 = x_8; +goto _start; +} +} +} +static lean_object* _init_l_Lean_IR_collectStructTypes___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = l_Array_empty(lean_box(0)); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_collectStructTypes___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = lean_unsigned_to_nat(16u); +x_3 = lean_mk_array(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_IR_collectStructTypes___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_collectStructTypes___closed__1; +x_2 = lean_unsigned_to_nat(0u); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_IR_collectStructTypes___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_collectStructTypes___closed__2; +x_2 = l_Lean_IR_collectStructTypes___closed__0; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_collectStructTypes(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Lean_IR_collectStructTypes___closed__3; +x_3 = l_List_forM___at___00Lean_IR_collectStructTypes_spec__0(x_1, x_2); +x_4 = lean_ctor_get(x_3, 1); +lean_inc(x_4); +lean_dec_ref(x_3); +return x_4; +} +} LEAN_EXPORT lean_object* l_Lean_IR_CollectUsedDecls_collect(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { @@ -1907,6 +4288,7 @@ return x_3; } lean_object* initialize_Lean_Compiler_InitAttr(uint8_t builtin); lean_object* initialize_Lean_Compiler_IR_CompilerM(uint8_t builtin); +lean_object* initialize_Lean_Compiler_NameMangling(uint8_t builtin); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Compiler_IR_EmitUtil(uint8_t builtin) { lean_object * res; @@ -1918,6 +4300,42 @@ lean_dec_ref(res); res = initialize_Lean_Compiler_IR_CompilerM(builtin); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Compiler_NameMangling(builtin); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_IR_IRType_hashApprox___closed__0 = _init_l_Lean_IR_IRType_hashApprox___closed__0(); +l_Lean_IR_instBEqIRTypeApprox = _init_l_Lean_IR_instBEqIRTypeApprox(); +lean_mark_persistent(l_Lean_IR_instBEqIRTypeApprox); +l_Lean_IR_instHashableIRTypeApprox___closed__0 = _init_l_Lean_IR_instHashableIRTypeApprox___closed__0(); +lean_mark_persistent(l_Lean_IR_instHashableIRTypeApprox___closed__0); +l_Lean_IR_instHashableIRTypeApprox = _init_l_Lean_IR_instHashableIRTypeApprox(); +lean_mark_persistent(l_Lean_IR_instHashableIRTypeApprox); +l_Lean_IR_instInhabitedStructTypeInfo_default___closed__0 = _init_l_Lean_IR_instInhabitedStructTypeInfo_default___closed__0(); +lean_mark_persistent(l_Lean_IR_instInhabitedStructTypeInfo_default___closed__0); +l_Lean_IR_instInhabitedStructTypeInfo_default___closed__1 = _init_l_Lean_IR_instInhabitedStructTypeInfo_default___closed__1(); +lean_mark_persistent(l_Lean_IR_instInhabitedStructTypeInfo_default___closed__1); +l_Lean_IR_instInhabitedStructTypeInfo_default = _init_l_Lean_IR_instInhabitedStructTypeInfo_default(); +lean_mark_persistent(l_Lean_IR_instInhabitedStructTypeInfo_default); +l_Lean_IR_instInhabitedStructTypeInfo = _init_l_Lean_IR_instInhabitedStructTypeInfo(); +lean_mark_persistent(l_Lean_IR_instInhabitedStructTypeInfo); +l_Lean_IR_CollectStructTypes_registerType___closed__0 = _init_l_Lean_IR_CollectStructTypes_registerType___closed__0(); +lean_mark_persistent(l_Lean_IR_CollectStructTypes_registerType___closed__0); +l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___redArg___closed__2 = _init_l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___redArg___closed__2(); +lean_mark_persistent(l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___redArg___closed__2); +l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___redArg___closed__1 = _init_l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___redArg___closed__1(); +lean_mark_persistent(l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___redArg___closed__1); +l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___redArg___closed__0 = _init_l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___redArg___closed__0(); +lean_mark_persistent(l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___redArg___closed__0); +l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___redArg___closed__3 = _init_l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___redArg___closed__3(); +lean_mark_persistent(l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_CollectStructTypes_addReshapeEntry_spec__0_spec__0___redArg___closed__3); +l_Lean_IR_collectStructTypes___closed__0 = _init_l_Lean_IR_collectStructTypes___closed__0(); +lean_mark_persistent(l_Lean_IR_collectStructTypes___closed__0); +l_Lean_IR_collectStructTypes___closed__1 = _init_l_Lean_IR_collectStructTypes___closed__1(); +lean_mark_persistent(l_Lean_IR_collectStructTypes___closed__1); +l_Lean_IR_collectStructTypes___closed__2 = _init_l_Lean_IR_collectStructTypes___closed__2(); +lean_mark_persistent(l_Lean_IR_collectStructTypes___closed__2); +l_Lean_IR_collectStructTypes___closed__3 = _init_l_Lean_IR_collectStructTypes___closed__3(); +lean_mark_persistent(l_Lean_IR_collectStructTypes___closed__3); l_Lean_IR_CollectMaps_collectVar___closed__0 = _init_l_Lean_IR_CollectMaps_collectVar___closed__0(); lean_mark_persistent(l_Lean_IR_CollectMaps_collectVar___closed__0); l_Lean_IR_CollectMaps_collectVar___closed__1 = _init_l_Lean_IR_CollectMaps_collectVar___closed__1(); diff --git a/stage0/stdlib/Lean/Compiler/IR/ExpandResetReuse.c b/stage0/stdlib/Lean/Compiler/IR/ExpandResetReuse.c index 5622d9835b98..3f99e9dd6aee 100644 --- a/stage0/stdlib/Lean/Compiler/IR/ExpandResetReuse.c +++ b/stage0/stdlib/Lean/Compiler/IR/ExpandResetReuse.c @@ -14,8 +14,10 @@ extern "C" { #endif LEAN_EXPORT lean_object* l_Lean_IR_ExpandResetReuse_eraseProjIncFor(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_searchAndExpand_spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00Lean_IR_ExpandResetReuse_setFields_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExpandResetReuse_CollectProjMap_collectVDecl(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_searchAndExpand_spec__1___lam__0___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExpandResetReuse_eraseProjIncForAux(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExpandResetReuse_mkSlowPath___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -32,8 +34,9 @@ static lean_object* l___private_Lean_Compiler_IR_ExpandResetReuse_0__Lean_IR_ini lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_ExpandResetReuse_CollectProjMap_collectFnBody_spec__0_spec__0___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___at___00Lean_IR_ExpandResetReuse_releaseUnreadFields_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_searchAndExpand_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExpandResetReuse_removeSelfSet___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_ExpandResetReuse_expand___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_ExpandResetReuse_expand___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); static lean_object* l___private_Lean_Compiler_IR_ExpandResetReuse_0__Lean_IR_initFn___closed__9_00___x40_Lean_Compiler_IR_ExpandResetReuse_1583541931____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l_Lean_IR_ExpandResetReuse_removeSelfSet(lean_object*, lean_object*); @@ -72,15 +75,14 @@ LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___00Std LEAN_EXPORT lean_object* l_Lean_IR_ExpandResetReuse_mkFresh(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_ExpandResetReuse_0__Lean_IR_initFn___closed__23_00___x40_Lean_Compiler_IR_ExpandResetReuse_1583541931____hygCtx___hyg_2_; lean_object* l_Std_DHashMap_Internal_Raw_u2080_insert___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_searchAndExpand_spec__0(size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExpandResetReuse_reuseToSet(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_IR_instBEqVarId_beq___boxed(lean_object*, lean_object*); uint64_t lean_uint64_shift_right(uint64_t, uint64_t); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_searchAndExpand_spec__1___lam__0(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_ExpandResetReuse_CollectProjMap_collectVDecl___closed__0; lean_object* lean_nat_div(lean_object*, lean_object*); lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___at___00Lean_IR_ExpandResetReuse_releaseUnreadFields_spec__0___boxed(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_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_searchAndExpand_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_ExpandResetReuse_isSelfSet_spec__0___boxed(lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_IR_instInhabitedFnBody_default__1; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_reuseToCtor_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -116,11 +118,11 @@ lean_object* l_Lean_IR_mkIf(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___at___00Lean_IR_ExpandResetReuse_releaseUnreadFields_spec__0___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExpandResetReuse_consumed___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_ExpandResetReuse_isSelfSet_spec__0___redArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_ExpandResetReuse_searchAndExpand___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_ExpandResetReuse_0__Lean_IR_initFn___closed__20_00___x40_Lean_Compiler_IR_ExpandResetReuse_1583541931____hygCtx___hyg_2_; lean_object* l_Lean_IR_FnBody_setBody(lean_object*, lean_object*); lean_object* l_Lean_IR_Decl_updateBody_x21(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_searchAndExpand_spec__0___lam__0(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_ExpandResetReuse_consumed_spec__0(lean_object*, lean_object*, size_t, size_t); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExpandResetReuse_mkFastPath___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -177,7 +179,7 @@ uint64_t l_Lean_IR_instHashableVarId_hash(lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_ExpandResetReuse_CollectProjMap_collectFnBody_spec__0_spec__1___redArg(lean_object*); lean_object* l_Lean_IR_Decl_normalizeIds(lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_ExpandResetReuse_expand(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_ExpandResetReuse_expand(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ExpandResetReuse_isSelfSet___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_removeSelfSet_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -1143,8 +1145,8 @@ if (lean_obj_tag(x_29) == 3) { lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; uint8_t x_48; x_30 = lean_ctor_get(x_28, 0); -x_31 = lean_ctor_get(x_29, 0); -x_32 = lean_ctor_get(x_29, 1); +x_31 = lean_ctor_get(x_29, 1); +x_32 = lean_ctor_get(x_29, 2); x_48 = l_Lean_IR_instBEqVarId_beq(x_30, x_20); if (x_48 == 0) { @@ -1561,7 +1563,7 @@ goto block_8; case 4: { lean_object* x_70; -x_70 = lean_ctor_get(x_2, 3); +x_70 = lean_ctor_get(x_2, 4); lean_inc(x_70); x_3 = x_70; goto block_8; @@ -1569,7 +1571,7 @@ goto block_8; case 5: { lean_object* x_71; -x_71 = lean_ctor_get(x_2, 5); +x_71 = lean_ctor_get(x_2, 6); lean_inc(x_71); x_3 = x_71; goto block_8; @@ -1916,71 +1918,39 @@ lean_dec(x_12); x_14 = lean_array_fget_borrowed(x_1, x_13); if (lean_obj_tag(x_14) == 0) { -lean_object* x_15; uint8_t x_16; +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; x_15 = l_Lean_IR_ExpandResetReuse_mkFresh___redArg(x_6); -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; -x_17 = lean_ctor_get(x_15, 0); -x_18 = lean_ctor_get(x_15, 1); -x_19 = lean_box(8); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec_ref(x_15); +x_18 = lean_box(8); lean_inc(x_2); -lean_ctor_set_tag(x_15, 3); -lean_ctor_set(x_15, 1, x_2); -lean_ctor_set(x_15, 0, x_13); +x_19 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_19, 0, x_7); +lean_ctor_set(x_19, 1, x_13); +lean_ctor_set(x_19, 2, x_2); x_20 = 1; -lean_inc(x_17); +lean_inc(x_16); x_21 = lean_alloc_ctor(7, 3, 2); -lean_ctor_set(x_21, 0, x_17); +lean_ctor_set(x_21, 0, x_16); lean_ctor_set(x_21, 1, x_10); lean_ctor_set(x_21, 2, x_5); lean_ctor_set_uint8(x_21, sizeof(void*)*3, x_20); lean_ctor_set_uint8(x_21, sizeof(void*)*3 + 1, x_8); x_22 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_22, 0, x_17); -lean_ctor_set(x_22, 1, x_19); -lean_ctor_set(x_22, 2, x_15); +lean_ctor_set(x_22, 0, x_16); +lean_ctor_set(x_22, 1, x_18); +lean_ctor_set(x_22, 2, x_19); lean_ctor_set(x_22, 3, x_21); x_4 = x_11; x_5 = x_22; -x_6 = x_18; +x_6 = x_17; goto _start; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; lean_object* x_30; -x_24 = lean_ctor_get(x_15, 0); -x_25 = lean_ctor_get(x_15, 1); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_15); -x_26 = lean_box(8); -lean_inc(x_2); -x_27 = lean_alloc_ctor(3, 2, 0); -lean_ctor_set(x_27, 0, x_13); -lean_ctor_set(x_27, 1, x_2); -x_28 = 1; -lean_inc(x_24); -x_29 = lean_alloc_ctor(7, 3, 2); -lean_ctor_set(x_29, 0, x_24); -lean_ctor_set(x_29, 1, x_10); -lean_ctor_set(x_29, 2, x_5); -lean_ctor_set_uint8(x_29, sizeof(void*)*3, x_28); -lean_ctor_set_uint8(x_29, sizeof(void*)*3 + 1, x_8); -x_30 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_30, 0, x_24); -lean_ctor_set(x_30, 1, x_26); -lean_ctor_set(x_30, 2, x_27); -lean_ctor_set(x_30, 3, x_29); -x_4 = x_11; -x_5 = x_30; -x_6 = x_25; -goto _start; -} -} -else -{ lean_dec(x_13); x_4 = x_11; goto _start; @@ -2204,9 +2174,9 @@ lean_dec_ref(x_6); if (lean_obj_tag(x_7) == 3) { lean_object* x_8; lean_object* x_9; uint8_t x_10; -x_8 = lean_ctor_get(x_7, 0); +x_8 = lean_ctor_get(x_7, 1); lean_inc(x_8); -x_9 = lean_ctor_get(x_7, 1); +x_9 = lean_ctor_get(x_7, 2); lean_inc(x_9); lean_dec_ref(x_7); x_10 = lean_nat_dec_eq(x_8, x_3); @@ -2315,9 +2285,9 @@ lean_dec_ref(x_5); if (lean_obj_tag(x_6) == 4) { lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_7 = lean_ctor_get(x_6, 0); +x_7 = lean_ctor_get(x_6, 1); lean_inc(x_7); -x_8 = lean_ctor_get(x_6, 1); +x_8 = lean_ctor_get(x_6, 2); lean_inc(x_8); lean_dec_ref(x_6); x_9 = lean_nat_dec_eq(x_7, x_3); @@ -2379,11 +2349,11 @@ lean_dec_ref(x_6); if (lean_obj_tag(x_7) == 5) { lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; uint8_t x_14; -x_8 = lean_ctor_get(x_7, 0); +x_8 = lean_ctor_get(x_7, 1); lean_inc(x_8); -x_9 = lean_ctor_get(x_7, 1); +x_9 = lean_ctor_get(x_7, 2); lean_inc(x_9); -x_10 = lean_ctor_get(x_7, 2); +x_10 = lean_ctor_get(x_7, 3); lean_inc(x_10); lean_dec_ref(x_7); x_14 = lean_nat_dec_eq(x_3, x_8); @@ -2523,252 +2493,264 @@ uint8_t x_25; x_25 = !lean_is_exclusive(x_2); if (x_25 == 0) { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; x_26 = lean_ctor_get(x_2, 0); x_27 = lean_ctor_get(x_2, 1); x_28 = lean_ctor_get(x_2, 2); x_29 = lean_ctor_get(x_2, 3); -x_30 = l_Lean_IR_ExpandResetReuse_isSelfUSet(x_1, x_26, x_27, x_28); -if (x_30 == 0) +x_30 = lean_ctor_get(x_2, 4); +x_31 = l_Lean_IR_ExpandResetReuse_isSelfUSet(x_1, x_26, x_28, x_29); +if (x_31 == 0) { -lean_object* x_31; -x_31 = l_Lean_IR_ExpandResetReuse_removeSelfSet(x_1, x_29); -lean_ctor_set(x_2, 3, x_31); +lean_object* x_32; +x_32 = l_Lean_IR_ExpandResetReuse_removeSelfSet(x_1, x_30); +lean_ctor_set(x_2, 4, x_32); return x_2; } else { lean_free_object(x_2); +lean_dec(x_29); lean_dec(x_28); lean_dec(x_27); lean_dec(x_26); -x_2 = x_29; +x_2 = x_30; goto _start; } } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_33 = lean_ctor_get(x_2, 0); -x_34 = lean_ctor_get(x_2, 1); -x_35 = lean_ctor_get(x_2, 2); -x_36 = lean_ctor_get(x_2, 3); +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; +x_34 = lean_ctor_get(x_2, 0); +x_35 = lean_ctor_get(x_2, 1); +x_36 = lean_ctor_get(x_2, 2); +x_37 = lean_ctor_get(x_2, 3); +x_38 = lean_ctor_get(x_2, 4); +lean_inc(x_38); +lean_inc(x_37); lean_inc(x_36); lean_inc(x_35); lean_inc(x_34); -lean_inc(x_33); lean_dec(x_2); -x_37 = l_Lean_IR_ExpandResetReuse_isSelfUSet(x_1, x_33, x_34, x_35); -if (x_37 == 0) +x_39 = l_Lean_IR_ExpandResetReuse_isSelfUSet(x_1, x_34, x_36, x_37); +if (x_39 == 0) { -lean_object* x_38; lean_object* x_39; -x_38 = l_Lean_IR_ExpandResetReuse_removeSelfSet(x_1, x_36); -x_39 = lean_alloc_ctor(4, 4, 0); -lean_ctor_set(x_39, 0, x_33); -lean_ctor_set(x_39, 1, x_34); -lean_ctor_set(x_39, 2, x_35); -lean_ctor_set(x_39, 3, x_38); -return x_39; +lean_object* x_40; lean_object* x_41; +x_40 = l_Lean_IR_ExpandResetReuse_removeSelfSet(x_1, x_38); +x_41 = lean_alloc_ctor(4, 5, 0); +lean_ctor_set(x_41, 0, x_34); +lean_ctor_set(x_41, 1, x_35); +lean_ctor_set(x_41, 2, x_36); +lean_ctor_set(x_41, 3, x_37); +lean_ctor_set(x_41, 4, x_40); +return x_41; } else { +lean_dec(x_37); +lean_dec(x_36); lean_dec(x_35); lean_dec(x_34); -lean_dec(x_33); -x_2 = x_36; +x_2 = x_38; goto _start; } } } case 5: { -uint8_t x_41; -x_41 = !lean_is_exclusive(x_2); -if (x_41 == 0) -{ -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; -x_42 = lean_ctor_get(x_2, 0); -x_43 = lean_ctor_get(x_2, 1); -x_44 = lean_ctor_get(x_2, 2); -x_45 = lean_ctor_get(x_2, 3); -x_46 = lean_ctor_get(x_2, 4); -x_47 = lean_ctor_get(x_2, 5); -x_48 = l_Lean_IR_ExpandResetReuse_isSelfSSet(x_1, x_42, x_43, x_44, x_45); -if (x_48 == 0) +uint8_t x_43; +x_43 = !lean_is_exclusive(x_2); +if (x_43 == 0) { -lean_object* x_49; -x_49 = l_Lean_IR_ExpandResetReuse_removeSelfSet(x_1, x_47); -lean_ctor_set(x_2, 5, x_49); +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; +x_44 = lean_ctor_get(x_2, 0); +x_45 = lean_ctor_get(x_2, 1); +x_46 = lean_ctor_get(x_2, 2); +x_47 = lean_ctor_get(x_2, 3); +x_48 = lean_ctor_get(x_2, 4); +x_49 = lean_ctor_get(x_2, 5); +x_50 = lean_ctor_get(x_2, 6); +x_51 = l_Lean_IR_ExpandResetReuse_isSelfSSet(x_1, x_44, x_46, x_47, x_48); +if (x_51 == 0) +{ +lean_object* x_52; +x_52 = l_Lean_IR_ExpandResetReuse_removeSelfSet(x_1, x_50); +lean_ctor_set(x_2, 6, x_52); return x_2; } else { lean_free_object(x_2); +lean_dec(x_49); +lean_dec(x_48); +lean_dec(x_47); lean_dec(x_46); lean_dec(x_45); lean_dec(x_44); -lean_dec(x_43); -lean_dec(x_42); -x_2 = x_47; +x_2 = x_50; goto _start; } } else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; -x_51 = lean_ctor_get(x_2, 0); -x_52 = lean_ctor_get(x_2, 1); -x_53 = lean_ctor_get(x_2, 2); -x_54 = lean_ctor_get(x_2, 3); -x_55 = lean_ctor_get(x_2, 4); -x_56 = lean_ctor_get(x_2, 5); +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; +x_54 = lean_ctor_get(x_2, 0); +x_55 = lean_ctor_get(x_2, 1); +x_56 = lean_ctor_get(x_2, 2); +x_57 = lean_ctor_get(x_2, 3); +x_58 = lean_ctor_get(x_2, 4); +x_59 = lean_ctor_get(x_2, 5); +x_60 = lean_ctor_get(x_2, 6); +lean_inc(x_60); +lean_inc(x_59); +lean_inc(x_58); +lean_inc(x_57); lean_inc(x_56); lean_inc(x_55); lean_inc(x_54); -lean_inc(x_53); -lean_inc(x_52); -lean_inc(x_51); lean_dec(x_2); -x_57 = l_Lean_IR_ExpandResetReuse_isSelfSSet(x_1, x_51, x_52, x_53, x_54); -if (x_57 == 0) -{ -lean_object* x_58; lean_object* x_59; -x_58 = l_Lean_IR_ExpandResetReuse_removeSelfSet(x_1, x_56); -x_59 = lean_alloc_ctor(5, 6, 0); -lean_ctor_set(x_59, 0, x_51); -lean_ctor_set(x_59, 1, x_52); -lean_ctor_set(x_59, 2, x_53); -lean_ctor_set(x_59, 3, x_54); -lean_ctor_set(x_59, 4, x_55); -lean_ctor_set(x_59, 5, x_58); -return x_59; +x_61 = l_Lean_IR_ExpandResetReuse_isSelfSSet(x_1, x_54, x_56, x_57, x_58); +if (x_61 == 0) +{ +lean_object* x_62; lean_object* x_63; +x_62 = l_Lean_IR_ExpandResetReuse_removeSelfSet(x_1, x_60); +x_63 = lean_alloc_ctor(5, 7, 0); +lean_ctor_set(x_63, 0, x_54); +lean_ctor_set(x_63, 1, x_55); +lean_ctor_set(x_63, 2, x_56); +lean_ctor_set(x_63, 3, x_57); +lean_ctor_set(x_63, 4, x_58); +lean_ctor_set(x_63, 5, x_59); +lean_ctor_set(x_63, 6, x_62); +return x_63; } else { +lean_dec(x_59); +lean_dec(x_58); +lean_dec(x_57); +lean_dec(x_56); lean_dec(x_55); lean_dec(x_54); -lean_dec(x_53); -lean_dec(x_52); -lean_dec(x_51); -x_2 = x_56; +x_2 = x_60; goto _start; } } } case 9: { -uint8_t x_61; -x_61 = !lean_is_exclusive(x_2); -if (x_61 == 0) +uint8_t x_65; +x_65 = !lean_is_exclusive(x_2); +if (x_65 == 0) { -lean_object* x_62; size_t x_63; size_t x_64; lean_object* x_65; -x_62 = lean_ctor_get(x_2, 3); -x_63 = lean_array_size(x_62); -x_64 = 0; -x_65 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_removeSelfSet_spec__0(x_1, x_63, x_64, x_62); -lean_ctor_set(x_2, 3, x_65); +lean_object* x_66; size_t x_67; size_t x_68; lean_object* x_69; +x_66 = lean_ctor_get(x_2, 3); +x_67 = lean_array_size(x_66); +x_68 = 0; +x_69 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_removeSelfSet_spec__0(x_1, x_67, x_68, x_66); +lean_ctor_set(x_2, 3, x_69); return x_2; } else { -lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; size_t x_70; size_t x_71; lean_object* x_72; lean_object* x_73; -x_66 = lean_ctor_get(x_2, 0); -x_67 = lean_ctor_get(x_2, 1); -x_68 = lean_ctor_get(x_2, 2); -x_69 = lean_ctor_get(x_2, 3); -lean_inc(x_69); -lean_inc(x_68); -lean_inc(x_67); -lean_inc(x_66); +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; size_t x_74; size_t x_75; lean_object* x_76; lean_object* x_77; +x_70 = lean_ctor_get(x_2, 0); +x_71 = lean_ctor_get(x_2, 1); +x_72 = lean_ctor_get(x_2, 2); +x_73 = lean_ctor_get(x_2, 3); +lean_inc(x_73); +lean_inc(x_72); +lean_inc(x_71); +lean_inc(x_70); lean_dec(x_2); -x_70 = lean_array_size(x_69); -x_71 = 0; -x_72 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_removeSelfSet_spec__0(x_1, x_70, x_71, x_69); -x_73 = lean_alloc_ctor(9, 4, 0); -lean_ctor_set(x_73, 0, x_66); -lean_ctor_set(x_73, 1, x_67); -lean_ctor_set(x_73, 2, x_68); -lean_ctor_set(x_73, 3, x_72); -return x_73; +x_74 = lean_array_size(x_73); +x_75 = 0; +x_76 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_removeSelfSet_spec__0(x_1, x_74, x_75, x_73); +x_77 = lean_alloc_ctor(9, 4, 0); +lean_ctor_set(x_77, 0, x_70); +lean_ctor_set(x_77, 1, x_71); +lean_ctor_set(x_77, 2, x_72); +lean_ctor_set(x_77, 3, x_76); +return x_77; } } default: { -uint8_t x_74; -x_74 = l_Lean_IR_FnBody_isTerminal(x_2); -if (x_74 == 0) +uint8_t x_78; +x_78 = l_Lean_IR_FnBody_isTerminal(x_2); +if (x_78 == 0) { switch (lean_obj_tag(x_2)) { case 0: { -lean_object* x_75; -x_75 = lean_ctor_get(x_2, 3); -lean_inc(x_75); -x_3 = x_75; +lean_object* x_79; +x_79 = lean_ctor_get(x_2, 3); +lean_inc(x_79); +x_3 = x_79; goto block_8; } case 1: { -lean_object* x_76; -x_76 = lean_ctor_get(x_2, 3); -lean_inc(x_76); -x_3 = x_76; +lean_object* x_80; +x_80 = lean_ctor_get(x_2, 3); +lean_inc(x_80); +x_3 = x_80; goto block_8; } case 2: { -lean_object* x_77; -x_77 = lean_ctor_get(x_2, 3); -lean_inc(x_77); -x_3 = x_77; +lean_object* x_81; +x_81 = lean_ctor_get(x_2, 3); +lean_inc(x_81); +x_3 = x_81; goto block_8; } case 4: { -lean_object* x_78; -x_78 = lean_ctor_get(x_2, 3); -lean_inc(x_78); -x_3 = x_78; +lean_object* x_82; +x_82 = lean_ctor_get(x_2, 4); +lean_inc(x_82); +x_3 = x_82; goto block_8; } case 5: { -lean_object* x_79; -x_79 = lean_ctor_get(x_2, 5); -lean_inc(x_79); -x_3 = x_79; +lean_object* x_83; +x_83 = lean_ctor_get(x_2, 6); +lean_inc(x_83); +x_3 = x_83; goto block_8; } case 3: { -lean_object* x_80; -x_80 = lean_ctor_get(x_2, 2); -lean_inc(x_80); -x_3 = x_80; +lean_object* x_84; +x_84 = lean_ctor_get(x_2, 2); +lean_inc(x_84); +x_3 = x_84; goto block_8; } case 6: { -lean_object* x_81; -x_81 = lean_ctor_get(x_2, 2); -lean_inc(x_81); -x_3 = x_81; +lean_object* x_85; +x_85 = lean_ctor_get(x_2, 2); +lean_inc(x_85); +x_3 = x_85; goto block_8; } case 7: { -lean_object* x_82; -x_82 = lean_ctor_get(x_2, 2); -lean_inc(x_82); -x_3 = x_82; +lean_object* x_86; +x_86 = lean_ctor_get(x_2, 2); +lean_inc(x_86); +x_3 = x_86; goto block_8; } case 8: { -lean_object* x_83; -x_83 = lean_ctor_get(x_2, 1); -lean_inc(x_83); -x_3 = x_83; +lean_object* x_87; +x_87 = lean_ctor_get(x_2, 1); +lean_inc(x_87); +x_3 = x_87; goto block_8; } default: @@ -3205,7 +3187,7 @@ goto block_10; case 4: { lean_object* x_87; -x_87 = lean_ctor_get(x_4, 3); +x_87 = lean_ctor_get(x_4, 4); lean_inc(x_87); x_5 = x_87; goto block_10; @@ -3213,7 +3195,7 @@ goto block_10; case 5: { lean_object* x_88; -x_88 = lean_ctor_get(x_4, 5); +x_88 = lean_ctor_get(x_4, 6); lean_inc(x_88); x_5 = x_88; goto block_10; @@ -3409,93 +3391,84 @@ lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_IR_ExpandResetReuse_expand(lean_object* x_1, lean_object* x_2, lean_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_IR_ExpandResetReuse_expand(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_9 = l_Lean_IR_ExpandResetReuse_eraseProjIncFor(x_4, x_5, x_2); -x_10 = lean_ctor_get(x_9, 0); +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_8 = l_Lean_IR_ExpandResetReuse_eraseProjIncFor(x_3, x_4, x_1); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); lean_inc(x_10); -x_11 = lean_ctor_get(x_9, 1); -lean_inc(x_11); -lean_dec_ref(x_9); -lean_inc(x_6); +lean_dec_ref(x_8); lean_inc(x_5); -x_12 = l_Lean_IR_ExpandResetReuse_mkFastPath(x_3, x_5, x_11, x_6, x_7, x_8); -x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_4); +x_11 = l_Lean_IR_ExpandResetReuse_mkFastPath(x_2, x_4, x_10, x_5, x_6, x_7); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec_ref(x_12); -x_15 = l_Lean_IR_ExpandResetReuse_eraseProjIncFor___closed__0; -x_16 = lean_apply_4(x_1, x_13, x_15, x_7, x_14); -x_17 = lean_ctor_get(x_16, 0); +lean_dec_ref(x_11); +x_14 = l_Lean_IR_ExpandResetReuse_eraseProjIncFor___closed__0; +x_15 = l_Lean_IR_ExpandResetReuse_searchAndExpand(x_12, x_14, x_6, x_13); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec_ref(x_16); -x_19 = l_Lean_IR_ExpandResetReuse_mkFresh___redArg(x_18); -x_20 = !lean_is_exclusive(x_19); -if (x_20 == 0) +lean_dec_ref(x_15); +x_18 = l_Lean_IR_ExpandResetReuse_mkFresh___redArg(x_17); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_21 = lean_ctor_get(x_19, 0); -lean_inc(x_5); -x_22 = l_Lean_IR_ExpandResetReuse_mkSlowPath(x_3, x_5, x_11, x_6); -lean_dec(x_11); -x_23 = lean_box(1); -x_24 = lean_alloc_ctor(12, 1, 0); -lean_ctor_set(x_24, 0, x_5); -lean_inc(x_21); -x_25 = l_Lean_IR_mkIf(x_21, x_22, x_17); -x_26 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_26, 0, x_21); -lean_ctor_set(x_26, 1, x_23); -lean_ctor_set(x_26, 2, x_24); -lean_ctor_set(x_26, 3, x_25); -x_27 = l_Lean_IR_reshape(x_10, x_26); -lean_ctor_set(x_19, 0, x_27); -return x_19; +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_20 = lean_ctor_get(x_18, 0); +lean_inc(x_4); +x_21 = l_Lean_IR_ExpandResetReuse_mkSlowPath(x_2, x_4, x_10, x_5); +lean_dec(x_10); +x_22 = lean_box(1); +x_23 = lean_alloc_ctor(12, 1, 0); +lean_ctor_set(x_23, 0, x_4); +lean_inc(x_20); +x_24 = l_Lean_IR_mkIf(x_20, x_21, x_16); +x_25 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_25, 0, x_20); +lean_ctor_set(x_25, 1, x_22); +lean_ctor_set(x_25, 2, x_23); +lean_ctor_set(x_25, 3, x_24); +x_26 = l_Lean_IR_reshape(x_9, x_25); +lean_ctor_set(x_18, 0, x_26); +return x_18; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_28 = lean_ctor_get(x_19, 0); -x_29 = lean_ctor_get(x_19, 1); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_19); -lean_inc(x_5); -x_30 = l_Lean_IR_ExpandResetReuse_mkSlowPath(x_3, x_5, x_11, x_6); -lean_dec(x_11); -x_31 = lean_box(1); -x_32 = lean_alloc_ctor(12, 1, 0); -lean_ctor_set(x_32, 0, x_5); +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_27 = lean_ctor_get(x_18, 0); +x_28 = lean_ctor_get(x_18, 1); lean_inc(x_28); -x_33 = l_Lean_IR_mkIf(x_28, x_30, x_17); -x_34 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_34, 0, x_28); -lean_ctor_set(x_34, 1, x_31); -lean_ctor_set(x_34, 2, x_32); -lean_ctor_set(x_34, 3, x_33); -x_35 = l_Lean_IR_reshape(x_10, x_34); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_29); -return x_36; -} -} +lean_inc(x_27); +lean_dec(x_18); +lean_inc(x_4); +x_29 = l_Lean_IR_ExpandResetReuse_mkSlowPath(x_2, x_4, x_10, x_5); +lean_dec(x_10); +x_30 = lean_box(1); +x_31 = lean_alloc_ctor(12, 1, 0); +lean_ctor_set(x_31, 0, x_4); +lean_inc(x_27); +x_32 = l_Lean_IR_mkIf(x_27, x_29, x_16); +x_33 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_33, 0, x_27); +lean_ctor_set(x_33, 1, x_30); +lean_ctor_set(x_33, 2, x_31); +lean_ctor_set(x_33, 3, x_32); +x_34 = l_Lean_IR_reshape(x_9, x_33); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_28); +return x_35; } -LEAN_EXPORT lean_object* l_Lean_IR_ExpandResetReuse_expand___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_IR_ExpandResetReuse_expand(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_3); -return x_9; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_searchAndExpand_spec__0(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_searchAndExpand_spec__1(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { uint8_t x_6; @@ -3503,7 +3476,6 @@ x_6 = lean_usize_dec_lt(x_2, x_1); if (x_6 == 0) { lean_object* x_7; -lean_dec_ref(x_4); x_7 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_7, 0, x_3); lean_ctor_set(x_7, 1, x_5); @@ -3523,8 +3495,7 @@ if (x_18 == 0) { lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_8, 1); -lean_inc_ref(x_4); -x_20 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_searchAndExpand_spec__0___lam__0(x_19, x_4, x_5); +x_20 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_searchAndExpand_spec__1___lam__0(x_19, x_4, x_5); x_21 = lean_ctor_get(x_20, 0); lean_inc(x_21); x_22 = lean_ctor_get(x_20, 1); @@ -3543,8 +3514,7 @@ x_24 = lean_ctor_get(x_8, 1); lean_inc(x_24); lean_inc(x_23); lean_dec(x_8); -lean_inc_ref(x_4); -x_25 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_searchAndExpand_spec__0___lam__0(x_24, x_4, x_5); +x_25 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_searchAndExpand_spec__1___lam__0(x_24, x_4, x_5); x_26 = lean_ctor_get(x_25, 0); lean_inc(x_26); x_27 = lean_ctor_get(x_25, 1); @@ -3566,8 +3536,7 @@ if (x_29 == 0) { lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; x_30 = lean_ctor_get(x_8, 0); -lean_inc_ref(x_4); -x_31 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_searchAndExpand_spec__0___lam__0(x_30, x_4, x_5); +x_31 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_searchAndExpand_spec__1___lam__0(x_30, x_4, x_5); x_32 = lean_ctor_get(x_31, 0); lean_inc(x_32); x_33 = lean_ctor_get(x_31, 1); @@ -3584,8 +3553,7 @@ lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean x_34 = lean_ctor_get(x_8, 0); lean_inc(x_34); lean_dec(x_8); -lean_inc_ref(x_4); -x_35 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_searchAndExpand_spec__0___lam__0(x_34, x_4, x_5); +x_35 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_searchAndExpand_spec__1___lam__0(x_34, x_4, x_5); x_36 = lean_ctor_get(x_35, 0); lean_inc(x_36); x_37 = lean_ctor_get(x_35, 1); @@ -3641,15 +3609,14 @@ goto _start; } else { -lean_object* x_24; lean_object* x_25; +lean_object* x_24; lean_inc(x_20); lean_inc(x_19); lean_inc(x_17); lean_dec_ref(x_1); -x_24 = lean_alloc_closure((void*)(l_Lean_IR_ExpandResetReuse_searchAndExpand), 4, 0); -x_25 = l_Lean_IR_ExpandResetReuse_expand(x_24, x_2, x_17, x_19, x_20, x_18, x_3, x_4); +x_24 = l_Lean_IR_ExpandResetReuse_expand(x_2, x_17, x_19, x_20, x_18, x_3, x_4); lean_dec(x_17); -return x_25; +return x_24; } } else @@ -3663,141 +3630,139 @@ goto block_15; } case 1: { -uint8_t x_26; -x_26 = !lean_is_exclusive(x_1); -if (x_26 == 0) +uint8_t x_25; +x_25 = !lean_is_exclusive(x_1); +if (x_25 == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_27 = lean_ctor_get(x_1, 2); -x_28 = lean_ctor_get(x_1, 3); -x_29 = l_Lean_IR_ExpandResetReuse_eraseProjIncFor___closed__0; -lean_inc_ref(x_3); -x_30 = l_Lean_IR_ExpandResetReuse_searchAndExpand(x_27, x_29, x_3, x_4); -x_31 = lean_ctor_get(x_30, 0); +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_26 = lean_ctor_get(x_1, 2); +x_27 = lean_ctor_get(x_1, 3); +x_28 = l_Lean_IR_ExpandResetReuse_eraseProjIncFor___closed__0; +x_29 = l_Lean_IR_ExpandResetReuse_searchAndExpand(x_26, x_28, x_3, x_4); +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_29, 1); lean_inc(x_31); -x_32 = lean_ctor_get(x_30, 1); -lean_inc(x_32); -lean_dec_ref(x_30); -x_33 = lean_box(12); -lean_ctor_set(x_1, 3, x_33); -lean_ctor_set(x_1, 2, x_31); -x_34 = l_Lean_IR_push(x_2, x_1); -x_1 = x_28; -x_2 = x_34; -x_4 = x_32; +lean_dec_ref(x_29); +x_32 = lean_box(12); +lean_ctor_set(x_1, 3, x_32); +lean_ctor_set(x_1, 2, x_30); +x_33 = l_Lean_IR_push(x_2, x_1); +x_1 = x_27; +x_2 = x_33; +x_4 = x_31; goto _start; } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_36 = lean_ctor_get(x_1, 0); -x_37 = lean_ctor_get(x_1, 1); -x_38 = lean_ctor_get(x_1, 2); -x_39 = lean_ctor_get(x_1, 3); -lean_inc(x_39); +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_35 = lean_ctor_get(x_1, 0); +x_36 = lean_ctor_get(x_1, 1); +x_37 = lean_ctor_get(x_1, 2); +x_38 = lean_ctor_get(x_1, 3); lean_inc(x_38); lean_inc(x_37); lean_inc(x_36); +lean_inc(x_35); lean_dec(x_1); -x_40 = l_Lean_IR_ExpandResetReuse_eraseProjIncFor___closed__0; -lean_inc_ref(x_3); -x_41 = l_Lean_IR_ExpandResetReuse_searchAndExpand(x_38, x_40, x_3, x_4); -x_42 = lean_ctor_get(x_41, 0); +x_39 = l_Lean_IR_ExpandResetReuse_eraseProjIncFor___closed__0; +x_40 = l_Lean_IR_ExpandResetReuse_searchAndExpand(x_37, x_39, x_3, x_4); +x_41 = lean_ctor_get(x_40, 0); +lean_inc(x_41); +x_42 = lean_ctor_get(x_40, 1); lean_inc(x_42); -x_43 = lean_ctor_get(x_41, 1); -lean_inc(x_43); -lean_dec_ref(x_41); -x_44 = lean_box(12); -x_45 = lean_alloc_ctor(1, 4, 0); -lean_ctor_set(x_45, 0, x_36); -lean_ctor_set(x_45, 1, x_37); -lean_ctor_set(x_45, 2, x_42); -lean_ctor_set(x_45, 3, x_44); -x_46 = l_Lean_IR_push(x_2, x_45); -x_1 = x_39; -x_2 = x_46; -x_4 = x_43; +lean_dec_ref(x_40); +x_43 = lean_box(12); +x_44 = lean_alloc_ctor(1, 4, 0); +lean_ctor_set(x_44, 0, x_35); +lean_ctor_set(x_44, 1, x_36); +lean_ctor_set(x_44, 2, x_41); +lean_ctor_set(x_44, 3, x_43); +x_45 = l_Lean_IR_push(x_2, x_44); +x_1 = x_38; +x_2 = x_45; +x_4 = x_42; goto _start; } } case 9: { -uint8_t x_48; -x_48 = !lean_is_exclusive(x_1); -if (x_48 == 0) +uint8_t x_47; +x_47 = !lean_is_exclusive(x_1); +if (x_47 == 0) { -lean_object* x_49; size_t x_50; size_t x_51; lean_object* x_52; uint8_t x_53; -x_49 = lean_ctor_get(x_1, 3); -x_50 = lean_array_size(x_49); -x_51 = 0; -x_52 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_searchAndExpand_spec__0(x_50, x_51, x_49, x_3, x_4); -x_53 = !lean_is_exclusive(x_52); -if (x_53 == 0) +lean_object* x_48; size_t x_49; size_t x_50; lean_object* x_51; uint8_t x_52; +x_48 = lean_ctor_get(x_1, 3); +x_49 = lean_array_size(x_48); +x_50 = 0; +x_51 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_searchAndExpand_spec__1(x_49, x_50, x_48, x_3, x_4); +x_52 = !lean_is_exclusive(x_51); +if (x_52 == 0) { -lean_object* x_54; lean_object* x_55; -x_54 = lean_ctor_get(x_52, 0); -lean_ctor_set(x_1, 3, x_54); -x_55 = l_Lean_IR_reshape(x_2, x_1); -lean_ctor_set(x_52, 0, x_55); -return x_52; +lean_object* x_53; lean_object* x_54; +x_53 = lean_ctor_get(x_51, 0); +lean_ctor_set(x_1, 3, x_53); +x_54 = l_Lean_IR_reshape(x_2, x_1); +lean_ctor_set(x_51, 0, x_54); +return x_51; } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_56 = lean_ctor_get(x_52, 0); -x_57 = lean_ctor_get(x_52, 1); -lean_inc(x_57); +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_55 = lean_ctor_get(x_51, 0); +x_56 = lean_ctor_get(x_51, 1); lean_inc(x_56); -lean_dec(x_52); -lean_ctor_set(x_1, 3, x_56); -x_58 = l_Lean_IR_reshape(x_2, x_1); -x_59 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_59, 0, x_58); -lean_ctor_set(x_59, 1, x_57); -return x_59; +lean_inc(x_55); +lean_dec(x_51); +lean_ctor_set(x_1, 3, x_55); +x_57 = l_Lean_IR_reshape(x_2, x_1); +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 { -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; size_t x_64; size_t x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_60 = lean_ctor_get(x_1, 0); -x_61 = lean_ctor_get(x_1, 1); -x_62 = lean_ctor_get(x_1, 2); -x_63 = lean_ctor_get(x_1, 3); -lean_inc(x_63); +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; size_t x_63; size_t x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_59 = lean_ctor_get(x_1, 0); +x_60 = lean_ctor_get(x_1, 1); +x_61 = lean_ctor_get(x_1, 2); +x_62 = lean_ctor_get(x_1, 3); lean_inc(x_62); lean_inc(x_61); lean_inc(x_60); +lean_inc(x_59); lean_dec(x_1); -x_64 = lean_array_size(x_63); -x_65 = 0; -x_66 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_searchAndExpand_spec__0(x_64, x_65, x_63, x_3, x_4); -x_67 = lean_ctor_get(x_66, 0); +x_63 = lean_array_size(x_62); +x_64 = 0; +x_65 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_searchAndExpand_spec__1(x_63, x_64, x_62, x_3, x_4); +x_66 = lean_ctor_get(x_65, 0); +lean_inc(x_66); +x_67 = lean_ctor_get(x_65, 1); lean_inc(x_67); -x_68 = lean_ctor_get(x_66, 1); -lean_inc(x_68); -if (lean_is_exclusive(x_66)) { - lean_ctor_release(x_66, 0); - lean_ctor_release(x_66, 1); - x_69 = x_66; +if (lean_is_exclusive(x_65)) { + lean_ctor_release(x_65, 0); + lean_ctor_release(x_65, 1); + x_68 = x_65; } else { - lean_dec_ref(x_66); - x_69 = lean_box(0); -} -x_70 = lean_alloc_ctor(9, 4, 0); -lean_ctor_set(x_70, 0, x_60); -lean_ctor_set(x_70, 1, x_61); -lean_ctor_set(x_70, 2, x_62); -lean_ctor_set(x_70, 3, x_67); -x_71 = l_Lean_IR_reshape(x_2, x_70); -if (lean_is_scalar(x_69)) { - x_72 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_65); + x_68 = lean_box(0); +} +x_69 = lean_alloc_ctor(9, 4, 0); +lean_ctor_set(x_69, 0, x_59); +lean_ctor_set(x_69, 1, x_60); +lean_ctor_set(x_69, 2, x_61); +lean_ctor_set(x_69, 3, x_66); +x_70 = l_Lean_IR_reshape(x_2, x_69); +if (lean_is_scalar(x_68)) { + x_71 = lean_alloc_ctor(0, 2, 0); } else { - x_72 = x_69; + x_71 = x_68; } -lean_ctor_set(x_72, 0, x_71); -lean_ctor_set(x_72, 1, x_68); -return x_72; +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_67); +return x_71; } } default: @@ -3827,7 +3792,6 @@ goto _start; else { lean_object* x_13; lean_object* x_14; -lean_dec_ref(x_7); x_13 = l_Lean_IR_reshape(x_6, x_5); x_14 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_14, 0, x_13); @@ -3837,7 +3801,7 @@ return x_14; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_searchAndExpand_spec__0___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_searchAndExpand_spec__1___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; @@ -3846,7 +3810,26 @@ x_5 = l_Lean_IR_ExpandResetReuse_searchAndExpand(x_1, x_4, x_2, x_3); return x_5; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_searchAndExpand_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_searchAndExpand_spec__1___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_searchAndExpand_spec__1___lam__0(x_1, x_2, x_3); +lean_dec_ref(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_ExpandResetReuse_expand___boxed(lean_object* x_1, lean_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_IR_ExpandResetReuse_expand(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec_ref(x_6); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_searchAndExpand_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { size_t x_6; size_t x_7; lean_object* x_8; @@ -3854,10 +3837,20 @@ x_6 = lean_unbox_usize(x_1); lean_dec(x_1); x_7 = lean_unbox_usize(x_2); lean_dec(x_2); -x_8 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_searchAndExpand_spec__0(x_6, x_7, x_3, x_4, x_5); +x_8 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_searchAndExpand_spec__1(x_6, x_7, x_3, x_4, x_5); +lean_dec_ref(x_4); return x_8; } } +LEAN_EXPORT lean_object* l_Lean_IR_ExpandResetReuse_searchAndExpand___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_IR_ExpandResetReuse_searchAndExpand(x_1, x_2, x_3, x_4); +lean_dec_ref(x_3); +return x_5; +} +} LEAN_EXPORT lean_object* l_Lean_IR_ExpandResetReuse_main(lean_object* x_1) { _start: { @@ -3875,6 +3868,7 @@ lean_dec(x_4); x_7 = l_Lean_IR_ExpandResetReuse_eraseProjIncFor___closed__0; lean_inc(x_2); x_8 = l_Lean_IR_ExpandResetReuse_searchAndExpand(x_2, x_7, x_3, x_6); +lean_dec_ref(x_3); x_9 = lean_ctor_get(x_8, 0); lean_inc(x_9); lean_dec_ref(x_8); diff --git a/stage0/stdlib/Lean/Compiler/IR/Format.c b/stage0/stdlib/Lean/Compiler/IR/Format.c index 2c138543dd24..665dedee7b09 100644 --- a/stage0/stdlib/Lean/Compiler/IR/Format.c +++ b/stage0/stdlib/Lean/Compiler/IR/Format.c @@ -53,7 +53,6 @@ LEAN_EXPORT lean_object* l_Lean_IR_instToFormatDecl___lam__0(lean_object*); static lean_object* l_Lean_IR_formatFnBodyHead___closed__2; LEAN_EXPORT lean_object* l_Lean_IR_instToFormatIRType; static lean_object* l_Lean_IR_formatArray___redArg___closed__4; -static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__6; static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__10; static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__15; static lean_object* l_Lean_IR_formatFnBodyHead___closed__3; @@ -96,6 +95,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_instToFormatDecl; lean_object* lean_nat_to_int(lean_object*); static lean_object* l_Lean_IR_instToFormatIRType___closed__0; static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__27; +static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__32; static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__15; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_formatArray___at___00Lean_IR_formatParams_spec__0_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0___boxed(lean_object*); @@ -111,6 +111,7 @@ static lean_object* l_Lean_IR_instToFormatLitVal___closed__0; static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__6; static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__12; LEAN_EXPORT lean_object* l_Lean_IR_formatParams(lean_object*); +static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__35; static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__2; static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__29; LEAN_EXPORT lean_object* l_Lean_IR_instToFormatCtorInfo___private__1(lean_object*); @@ -140,6 +141,7 @@ static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType_ static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__0; LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam(lean_object*); static lean_object* l_Lean_IR_formatDecl___closed__0; +static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__33; static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop___closed__1; LEAN_EXPORT lean_object* l_Lean_IR_instToStringDecl; static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__14; @@ -148,7 +150,6 @@ LEAN_EXPORT lean_object* l_Lean_IR_instToFormatLitVal; static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__8; extern lean_object* l_Std_Format_defWidth; static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop___closed__0; -static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__5; static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__1; LEAN_EXPORT lean_object* l_Lean_IR_instToFormatLitVal___private__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instToFormatCtorInfo; @@ -171,7 +172,6 @@ lean_object* l_Id_instMonad___lam__1___boxed(lean_object*, lean_object*, lean_ob LEAN_EXPORT lean_object* l_Lean_IR_instToFormatFnBody; static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__22; static lean_object* l_Lean_IR_formatArray___redArg___lam__0___closed__1; -static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__3; static lean_object* l_Lean_IR_formatFnBodyHead___closed__4; static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__7; static lean_object* l_Lean_IR_formatAlt___closed__1; @@ -188,8 +188,10 @@ static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___ LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instToFormatExpr; LEAN_EXPORT lean_object* l_Lean_IR_formatArray___at___00Lean_IR_formatParams_spec__0(lean_object*); +static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__36; static lean_object* l_Lean_IR_formatFnBodyHead___closed__28; static lean_object* l_Lean_IR_formatFnBodyHead___closed__15; +static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__34; static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__7; LEAN_EXPORT lean_object* l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(lean_object*); static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__12; @@ -212,7 +214,6 @@ static lean_object* l_Lean_IR_formatFnBodyHead___closed__30; size_t lean_usize_add(size_t, size_t); static lean_object* l_Lean_IR_formatArray___redArg___closed__0; static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__3; -static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__4; lean_object* lean_array_uget(lean_object*, size_t); static lean_object* l_Lean_IR_instToFormatParam___closed__0; static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop___closed__5; @@ -982,7 +983,7 @@ static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatE _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("uproj[", 6, 6); +x_1 = lean_mk_string_unchecked(", ", 2, 2); return x_1; } } @@ -1000,7 +1001,7 @@ static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatE _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("sproj[", 6, 6); +x_1 = lean_mk_string_unchecked("uproj[", 6, 6); return x_1; } } @@ -1018,7 +1019,7 @@ static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatE _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked(", ", 2, 2); +x_1 = lean_mk_string_unchecked("sproj[", 6, 6); return x_1; } } @@ -1287,267 +1288,263 @@ return x_58; } case 3: { -uint8_t x_62; -x_62 = !lean_is_exclusive(x_1); -if (x_62 == 0) -{ -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_63 = lean_ctor_get(x_1, 0); -x_64 = lean_ctor_get(x_1, 1); +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_62 = lean_ctor_get(x_1, 0); +lean_inc(x_62); +x_63 = lean_ctor_get(x_1, 1); +lean_inc(x_63); +x_64 = lean_ctor_get(x_1, 2); +lean_inc(x_64); +lean_dec_ref(x_1); x_65 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__11; -x_66 = l_Nat_reprFast(x_63); +x_66 = l_Nat_reprFast(x_62); x_67 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_67, 0, x_66); -lean_ctor_set_tag(x_1, 5); -lean_ctor_set(x_1, 1, x_67); -lean_ctor_set(x_1, 0, x_65); -x_68 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__3; -x_69 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_69, 0, x_1); -lean_ctor_set(x_69, 1, x_68); -x_70 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; -x_71 = l_Nat_reprFast(x_64); -x_72 = lean_string_append(x_70, x_71); -lean_dec_ref(x_71); -x_73 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_73, 0, x_72); -x_74 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_74, 0, x_69); -lean_ctor_set(x_74, 1, x_73); -return x_74; -} -else -{ -lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_75 = lean_ctor_get(x_1, 0); -x_76 = lean_ctor_get(x_1, 1); -lean_inc(x_76); -lean_inc(x_75); -lean_dec(x_1); -x_77 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__11; -x_78 = l_Nat_reprFast(x_75); +x_68 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_68, 0, x_65); +lean_ctor_set(x_68, 1, x_67); +x_69 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__13; +x_70 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set(x_70, 1, x_69); +x_71 = l_Nat_reprFast(x_63); +x_72 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_72, 0, x_71); +x_73 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_73, 0, x_70); +lean_ctor_set(x_73, 1, x_72); +x_74 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__3; +x_75 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_74); +x_76 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_77 = l_Nat_reprFast(x_64); +x_78 = lean_string_append(x_76, x_77); +lean_dec_ref(x_77); x_79 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_79, 0, x_78); x_80 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_80, 0, x_77); +lean_ctor_set(x_80, 0, x_75); lean_ctor_set(x_80, 1, x_79); -x_81 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__3; -x_82 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_82, 0, x_80); -lean_ctor_set(x_82, 1, x_81); -x_83 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; -x_84 = l_Nat_reprFast(x_76); -x_85 = lean_string_append(x_83, x_84); -lean_dec_ref(x_84); +return x_80; +} +case 4: +{ +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_81 = lean_ctor_get(x_1, 0); +lean_inc(x_81); +x_82 = lean_ctor_get(x_1, 1); +lean_inc(x_82); +x_83 = lean_ctor_get(x_1, 2); +lean_inc(x_83); +lean_dec_ref(x_1); +x_84 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__15; +x_85 = l_Nat_reprFast(x_81); x_86 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_86, 0, x_85); x_87 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_87, 0, x_82); +lean_ctor_set(x_87, 0, x_84); lean_ctor_set(x_87, 1, x_86); -return x_87; -} -} -case 4: -{ -uint8_t x_88; -x_88 = !lean_is_exclusive(x_1); -if (x_88 == 0) -{ -lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_89 = lean_ctor_get(x_1, 0); -x_90 = lean_ctor_get(x_1, 1); -x_91 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__13; -x_92 = l_Nat_reprFast(x_89); -x_93 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_93, 0, x_92); -lean_ctor_set_tag(x_1, 5); -lean_ctor_set(x_1, 1, x_93); -lean_ctor_set(x_1, 0, x_91); -x_94 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__3; -x_95 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_95, 0, x_1); -lean_ctor_set(x_95, 1, x_94); -x_96 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; -x_97 = l_Nat_reprFast(x_90); -x_98 = lean_string_append(x_96, x_97); -lean_dec_ref(x_97); -x_99 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_99, 0, x_98); -x_100 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_100, 0, x_95); -lean_ctor_set(x_100, 1, x_99); -return x_100; -} -else -{ -lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; -x_101 = lean_ctor_get(x_1, 0); -x_102 = lean_ctor_get(x_1, 1); -lean_inc(x_102); -lean_inc(x_101); -lean_dec(x_1); -x_103 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__13; -x_104 = l_Nat_reprFast(x_101); -x_105 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_105, 0, x_104); -x_106 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_106, 0, x_103); -lean_ctor_set(x_106, 1, x_105); -x_107 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__3; -x_108 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_108, 0, x_106); -lean_ctor_set(x_108, 1, x_107); -x_109 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; -x_110 = l_Nat_reprFast(x_102); -x_111 = lean_string_append(x_109, x_110); -lean_dec_ref(x_110); -x_112 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_112, 0, x_111); -x_113 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_113, 0, x_108); -lean_ctor_set(x_113, 1, x_112); -return x_113; -} +x_88 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__13; +x_89 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_89, 0, x_87); +lean_ctor_set(x_89, 1, x_88); +x_90 = l_Nat_reprFast(x_82); +x_91 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_91, 0, x_90); +x_92 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_92, 0, x_89); +lean_ctor_set(x_92, 1, x_91); +x_93 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__3; +x_94 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_94, 0, x_92); +lean_ctor_set(x_94, 1, x_93); +x_95 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_96 = l_Nat_reprFast(x_83); +x_97 = lean_string_append(x_95, x_96); +lean_dec_ref(x_96); +x_98 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_98, 0, x_97); +x_99 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_99, 0, x_94); +lean_ctor_set(x_99, 1, x_98); +return x_99; } case 5: { -lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; -x_114 = lean_ctor_get(x_1, 0); -lean_inc(x_114); -x_115 = lean_ctor_get(x_1, 1); -lean_inc(x_115); -x_116 = lean_ctor_get(x_1, 2); -lean_inc(x_116); +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_100 = lean_ctor_get(x_1, 0); +lean_inc(x_100); +x_101 = lean_ctor_get(x_1, 1); +lean_inc(x_101); +x_102 = lean_ctor_get(x_1, 2); +lean_inc(x_102); +x_103 = lean_ctor_get(x_1, 3); +lean_inc(x_103); lean_dec_ref(x_1); -x_117 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__15; -x_118 = l_Nat_reprFast(x_114); -x_119 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_119, 0, x_118); -x_120 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_120, 0, x_117); -lean_ctor_set(x_120, 1, x_119); -x_121 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__17; -x_122 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_122, 0, x_120); -lean_ctor_set(x_122, 1, x_121); -x_123 = l_Nat_reprFast(x_115); -x_124 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_124, 0, x_123); -x_125 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_125, 0, x_122); -lean_ctor_set(x_125, 1, x_124); -x_126 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__3; -x_127 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_127, 0, x_125); -lean_ctor_set(x_127, 1, x_126); -x_128 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; -x_129 = l_Nat_reprFast(x_116); -x_130 = lean_string_append(x_128, x_129); -lean_dec_ref(x_129); -x_131 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_131, 0, x_130); -x_132 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_132, 0, x_127); -lean_ctor_set(x_132, 1, x_131); -return x_132; +x_104 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__17; +x_105 = l_Nat_reprFast(x_100); +x_106 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_106, 0, x_105); +x_107 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_107, 0, x_104); +lean_ctor_set(x_107, 1, x_106); +x_108 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__13; +x_109 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_109, 0, x_107); +lean_ctor_set(x_109, 1, x_108); +x_110 = l_Nat_reprFast(x_101); +x_111 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_111, 0, x_110); +x_112 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_112, 0, x_109); +lean_ctor_set(x_112, 1, x_111); +x_113 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_113, 0, x_112); +lean_ctor_set(x_113, 1, x_108); +x_114 = l_Nat_reprFast(x_102); +x_115 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_115, 0, x_114); +x_116 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_116, 0, x_113); +lean_ctor_set(x_116, 1, x_115); +x_117 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__3; +x_118 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_118, 0, x_116); +lean_ctor_set(x_118, 1, x_117); +x_119 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_120 = l_Nat_reprFast(x_103); +x_121 = lean_string_append(x_119, x_120); +lean_dec_ref(x_120); +x_122 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_122, 0, x_121); +x_123 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_123, 0, x_118); +lean_ctor_set(x_123, 1, x_122); +return x_123; } case 6: { -uint8_t x_133; -x_133 = !lean_is_exclusive(x_1); -if (x_133 == 0) -{ -lean_object* x_134; lean_object* x_135; uint8_t x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; -x_134 = lean_ctor_get(x_1, 0); -x_135 = lean_ctor_get(x_1, 1); -x_136 = 1; -x_137 = l_Lean_Name_toString(x_134, x_136); -x_138 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_138, 0, x_137); -x_139 = l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(x_135); -lean_dec_ref(x_135); +uint8_t x_124; +x_124 = !lean_is_exclusive(x_1); +if (x_124 == 0) +{ +lean_object* x_125; lean_object* x_126; uint8_t x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; +x_125 = lean_ctor_get(x_1, 0); +x_126 = lean_ctor_get(x_1, 1); +x_127 = 1; +x_128 = l_Lean_Name_toString(x_125, x_127); +x_129 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_129, 0, x_128); +x_130 = l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(x_126); +lean_dec_ref(x_126); lean_ctor_set_tag(x_1, 5); -lean_ctor_set(x_1, 1, x_139); -lean_ctor_set(x_1, 0, x_138); +lean_ctor_set(x_1, 1, x_130); +lean_ctor_set(x_1, 0, x_129); return x_1; } else { -lean_object* x_140; lean_object* x_141; uint8_t x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; -x_140 = lean_ctor_get(x_1, 0); -x_141 = lean_ctor_get(x_1, 1); -lean_inc(x_141); -lean_inc(x_140); +lean_object* x_131; lean_object* x_132; uint8_t x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; +x_131 = lean_ctor_get(x_1, 0); +x_132 = lean_ctor_get(x_1, 1); +lean_inc(x_132); +lean_inc(x_131); lean_dec(x_1); +x_133 = 1; +x_134 = l_Lean_Name_toString(x_131, x_133); +x_135 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_135, 0, x_134); +x_136 = l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(x_132); +lean_dec_ref(x_132); +x_137 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_137, 0, x_135); +lean_ctor_set(x_137, 1, x_136); +return x_137; +} +} +case 7: +{ +uint8_t x_138; +x_138 = !lean_is_exclusive(x_1); +if (x_138 == 0) +{ +lean_object* x_139; lean_object* x_140; lean_object* x_141; uint8_t x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; +x_139 = lean_ctor_get(x_1, 0); +x_140 = lean_ctor_get(x_1, 1); +x_141 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__19; x_142 = 1; -x_143 = l_Lean_Name_toString(x_140, x_142); +x_143 = l_Lean_Name_toString(x_139, x_142); x_144 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_144, 0, x_143); -x_145 = l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(x_141); -lean_dec_ref(x_141); +lean_ctor_set_tag(x_1, 5); +lean_ctor_set(x_1, 1, x_144); +lean_ctor_set(x_1, 0, x_141); +x_145 = l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(x_140); +lean_dec_ref(x_140); x_146 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_146, 0, x_144); +lean_ctor_set(x_146, 0, x_1); lean_ctor_set(x_146, 1, x_145); return x_146; } -} -case 7: -{ -uint8_t x_147; -x_147 = !lean_is_exclusive(x_1); -if (x_147 == 0) +else { -lean_object* x_148; lean_object* x_149; lean_object* x_150; uint8_t x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; -x_148 = lean_ctor_get(x_1, 0); -x_149 = lean_ctor_get(x_1, 1); -x_150 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__19; -x_151 = 1; -x_152 = l_Lean_Name_toString(x_148, x_151); -x_153 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_153, 0, x_152); -lean_ctor_set_tag(x_1, 5); -lean_ctor_set(x_1, 1, x_153); -lean_ctor_set(x_1, 0, x_150); -x_154 = l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(x_149); -lean_dec_ref(x_149); +lean_object* x_147; lean_object* x_148; lean_object* x_149; uint8_t x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; +x_147 = lean_ctor_get(x_1, 0); +x_148 = lean_ctor_get(x_1, 1); +lean_inc(x_148); +lean_inc(x_147); +lean_dec(x_1); +x_149 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__19; +x_150 = 1; +x_151 = l_Lean_Name_toString(x_147, x_150); +x_152 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_152, 0, x_151); +x_153 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_153, 0, x_149); +lean_ctor_set(x_153, 1, x_152); +x_154 = l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(x_148); +lean_dec_ref(x_148); x_155 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_155, 0, x_1); +lean_ctor_set(x_155, 0, x_153); lean_ctor_set(x_155, 1, x_154); return x_155; } -else -{ -lean_object* x_156; lean_object* x_157; lean_object* x_158; uint8_t x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; -x_156 = lean_ctor_get(x_1, 0); -x_157 = lean_ctor_get(x_1, 1); -lean_inc(x_157); -lean_inc(x_156); -lean_dec(x_1); -x_158 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__19; -x_159 = 1; -x_160 = l_Lean_Name_toString(x_156, x_159); -x_161 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_161, 0, x_160); -x_162 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_162, 0, x_158); -lean_ctor_set(x_162, 1, x_161); -x_163 = l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(x_157); -lean_dec_ref(x_157); -x_164 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_164, 0, x_162); -lean_ctor_set(x_164, 1, x_163); -return x_164; -} } case 8: { -uint8_t x_165; -x_165 = !lean_is_exclusive(x_1); -if (x_165 == 0) +uint8_t x_156; +x_156 = !lean_is_exclusive(x_1); +if (x_156 == 0) +{ +lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; +x_157 = lean_ctor_get(x_1, 0); +x_158 = lean_ctor_get(x_1, 1); +x_159 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__21; +x_160 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_161 = l_Nat_reprFast(x_157); +x_162 = lean_string_append(x_160, x_161); +lean_dec_ref(x_161); +x_163 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_163, 0, x_162); +lean_ctor_set_tag(x_1, 5); +lean_ctor_set(x_1, 1, x_163); +lean_ctor_set(x_1, 0, x_159); +x_164 = l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(x_158); +lean_dec_ref(x_158); +x_165 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_165, 0, x_1); +lean_ctor_set(x_165, 1, x_164); +return x_165; +} +else { -lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; +lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; x_166 = lean_ctor_get(x_1, 0); x_167 = lean_ctor_get(x_1, 1); +lean_inc(x_167); +lean_inc(x_166); +lean_dec(x_1); x_168 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__21; x_169 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; x_170 = l_Nat_reprFast(x_166); @@ -1555,168 +1552,143 @@ x_171 = lean_string_append(x_169, x_170); lean_dec_ref(x_170); x_172 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_172, 0, x_171); -lean_ctor_set_tag(x_1, 5); -lean_ctor_set(x_1, 1, x_172); -lean_ctor_set(x_1, 0, x_168); -x_173 = l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(x_167); +x_173 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_173, 0, x_168); +lean_ctor_set(x_173, 1, x_172); +x_174 = l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(x_167); lean_dec_ref(x_167); -x_174 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_174, 0, x_1); -lean_ctor_set(x_174, 1, x_173); -return x_174; -} -else -{ -lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; -x_175 = lean_ctor_get(x_1, 0); -x_176 = lean_ctor_get(x_1, 1); -lean_inc(x_176); -lean_inc(x_175); -lean_dec(x_1); -x_177 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__21; -x_178 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; -x_179 = l_Nat_reprFast(x_175); -x_180 = lean_string_append(x_178, x_179); -lean_dec_ref(x_179); -x_181 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_181, 0, x_180); -x_182 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_182, 0, x_177); -lean_ctor_set(x_182, 1, x_181); -x_183 = l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(x_176); -lean_dec_ref(x_176); -x_184 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_184, 0, x_182); -lean_ctor_set(x_184, 1, x_183); -return x_184; +x_175 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_175, 0, x_173); +lean_ctor_set(x_175, 1, x_174); +return x_175; } } case 9: { -uint8_t x_185; -x_185 = !lean_is_exclusive(x_1); -if (x_185 == 0) -{ -lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; -x_186 = lean_ctor_get(x_1, 1); -x_187 = lean_ctor_get(x_1, 0); -lean_dec(x_187); -x_188 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__23; -x_189 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; -x_190 = l_Nat_reprFast(x_186); -x_191 = lean_string_append(x_189, x_190); -lean_dec_ref(x_190); -x_192 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_192, 0, x_191); +uint8_t x_176; +x_176 = !lean_is_exclusive(x_1); +if (x_176 == 0) +{ +lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; +x_177 = lean_ctor_get(x_1, 1); +x_178 = lean_ctor_get(x_1, 0); +lean_dec(x_178); +x_179 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__23; +x_180 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_181 = l_Nat_reprFast(x_177); +x_182 = lean_string_append(x_180, x_181); +lean_dec_ref(x_181); +x_183 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_183, 0, x_182); lean_ctor_set_tag(x_1, 5); -lean_ctor_set(x_1, 1, x_192); -lean_ctor_set(x_1, 0, x_188); +lean_ctor_set(x_1, 1, x_183); +lean_ctor_set(x_1, 0, x_179); return x_1; } else { -lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; -x_193 = lean_ctor_get(x_1, 1); -lean_inc(x_193); +lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; +x_184 = lean_ctor_get(x_1, 1); +lean_inc(x_184); lean_dec(x_1); -x_194 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__23; -x_195 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; -x_196 = l_Nat_reprFast(x_193); -x_197 = lean_string_append(x_195, x_196); -lean_dec_ref(x_196); -x_198 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_198, 0, x_197); -x_199 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_199, 0, x_194); -lean_ctor_set(x_199, 1, x_198); -return x_199; +x_185 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__23; +x_186 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_187 = l_Nat_reprFast(x_184); +x_188 = lean_string_append(x_186, x_187); +lean_dec_ref(x_187); +x_189 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_189, 0, x_188); +x_190 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_190, 0, x_185); +lean_ctor_set(x_190, 1, x_189); +return x_190; } } case 10: { -uint8_t x_200; -x_200 = !lean_is_exclusive(x_1); -if (x_200 == 0) +uint8_t x_191; +x_191 = !lean_is_exclusive(x_1); +if (x_191 == 0) { -lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; -x_201 = lean_ctor_get(x_1, 0); -x_202 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__25; -x_203 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; -x_204 = l_Nat_reprFast(x_201); -x_205 = lean_string_append(x_203, x_204); -lean_dec_ref(x_204); +lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; +x_192 = lean_ctor_get(x_1, 0); +x_193 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__25; +x_194 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_195 = l_Nat_reprFast(x_192); +x_196 = lean_string_append(x_194, x_195); +lean_dec_ref(x_195); lean_ctor_set_tag(x_1, 3); -lean_ctor_set(x_1, 0, x_205); -x_206 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_206, 0, x_202); -lean_ctor_set(x_206, 1, x_1); -return x_206; +lean_ctor_set(x_1, 0, x_196); +x_197 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_197, 0, x_193); +lean_ctor_set(x_197, 1, x_1); +return x_197; } else { -lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; -x_207 = lean_ctor_get(x_1, 0); -lean_inc(x_207); +lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; +x_198 = lean_ctor_get(x_1, 0); +lean_inc(x_198); lean_dec(x_1); -x_208 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__25; -x_209 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; -x_210 = l_Nat_reprFast(x_207); -x_211 = lean_string_append(x_209, x_210); -lean_dec_ref(x_210); -x_212 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_212, 0, x_211); -x_213 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_213, 0, x_208); -lean_ctor_set(x_213, 1, x_212); -return x_213; +x_199 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__25; +x_200 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_201 = l_Nat_reprFast(x_198); +x_202 = lean_string_append(x_200, x_201); +lean_dec_ref(x_201); +x_203 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_203, 0, x_202); +x_204 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_204, 0, x_199); +lean_ctor_set(x_204, 1, x_203); +return x_204; } } case 11: { -lean_object* x_214; lean_object* x_215; -x_214 = lean_ctor_get(x_1, 0); -lean_inc_ref(x_214); +lean_object* x_205; lean_object* x_206; +x_205 = lean_ctor_get(x_1, 0); +lean_inc_ref(x_205); lean_dec_ref(x_1); -x_215 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatLitVal(x_214); -return x_215; +x_206 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatLitVal(x_205); +return x_206; } default: { -uint8_t x_216; -x_216 = !lean_is_exclusive(x_1); -if (x_216 == 0) -{ -lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; -x_217 = lean_ctor_get(x_1, 0); -x_218 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__27; -x_219 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; -x_220 = l_Nat_reprFast(x_217); -x_221 = lean_string_append(x_219, x_220); -lean_dec_ref(x_220); +uint8_t x_207; +x_207 = !lean_is_exclusive(x_1); +if (x_207 == 0) +{ +lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; +x_208 = lean_ctor_get(x_1, 0); +x_209 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__27; +x_210 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_211 = l_Nat_reprFast(x_208); +x_212 = lean_string_append(x_210, x_211); +lean_dec_ref(x_211); lean_ctor_set_tag(x_1, 3); -lean_ctor_set(x_1, 0, x_221); -x_222 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_222, 0, x_218); -lean_ctor_set(x_222, 1, x_1); -return x_222; +lean_ctor_set(x_1, 0, x_212); +x_213 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_213, 0, x_209); +lean_ctor_set(x_213, 1, x_1); +return x_213; } else { -lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; -x_223 = lean_ctor_get(x_1, 0); -lean_inc(x_223); +lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; +x_214 = lean_ctor_get(x_1, 0); +lean_inc(x_214); lean_dec(x_1); -x_224 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__27; -x_225 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; -x_226 = l_Nat_reprFast(x_223); -x_227 = lean_string_append(x_225, x_226); -lean_dec_ref(x_226); -x_228 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_228, 0, x_227); -x_229 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_229, 0, x_224); -lean_ctor_set(x_229, 1, x_228); -return x_229; +x_215 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__27; +x_216 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_217 = l_Nat_reprFast(x_214); +x_218 = lean_string_append(x_216, x_217); +lean_dec_ref(x_217); +x_219 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_219, 0, x_218); +x_220 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_220, 0, x_215); +lean_ctor_set(x_220, 1, x_219); +return x_220; } } } @@ -1965,24 +1937,6 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__18() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("struct ", 7, 7); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__19() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__18; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType_spec__0(lean_object* x_1, lean_object* x_2) { _start: { @@ -2021,7 +1975,7 @@ return x_9; } } } -static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__20() { +static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__18() { _start: { lean_object* x_1; @@ -2029,35 +1983,35 @@ x_1 = lean_mk_string_unchecked("{", 1, 1); return x_1; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__22() { +static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__20() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__20; +x_1 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__18; x_2 = lean_string_length(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__23() { +static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__21() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__22; +x_1 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__20; x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__24() { +static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__22() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__20; +x_1 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__18; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__21() { +static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__19() { _start: { lean_object* x_1; @@ -2065,21 +2019,29 @@ x_1 = lean_mk_string_unchecked("}", 1, 1); return x_1; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__25() { +static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__23() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__21; +x_1 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__19; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } +static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__24() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(":", 1, 1); +return x_1; +} +} static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__26() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("union ", 6, 6); +x_1 = lean_mk_string_unchecked(" | ", 3, 3); return x_1; } } @@ -2093,11 +2055,11 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__28() { +static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__25() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("tagged", 6, 6); +x_1 = lean_mk_string_unchecked("(", 1, 1); return x_1; } } @@ -2105,13 +2067,67 @@ static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatI _start: { lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__25; +x_2 = lean_string_length(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__30() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__29; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__31() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__25; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__28() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(")", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__32() { +_start: +{ +lean_object* x_1; lean_object* x_2; x_1 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__28; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__30() { +static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__33() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("tagged", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__34() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__33; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__35() { _start: { lean_object* x_1; @@ -2119,11 +2135,11 @@ x_1 = lean_mk_string_unchecked("void", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__31() { +static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__36() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__30; +x_1 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__35; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; @@ -2195,149 +2211,141 @@ return x_11; } case 10: { -uint8_t x_12; -x_12 = !lean_is_exclusive(x_1); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; -x_13 = lean_ctor_get(x_1, 1); -x_14 = lean_ctor_get(x_1, 0); -lean_dec(x_14); -x_15 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__19; -x_16 = lean_array_to_list(x_13); -x_17 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__17; -x_18 = l_Std_Format_joinSep___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType_spec__0(x_16, x_17); -x_19 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__23; -x_20 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__24; -lean_ctor_set_tag(x_1, 5); -lean_ctor_set(x_1, 1, x_18); -lean_ctor_set(x_1, 0, x_20); -x_21 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__25; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_34; uint8_t x_35; +x_12 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_12); +x_13 = lean_ctor_get(x_1, 2); +lean_inc(x_13); +x_14 = lean_ctor_get(x_1, 3); +lean_inc(x_14); +lean_dec_ref(x_1); +x_15 = lean_array_to_list(x_12); +x_16 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__13; +x_17 = l_Std_Format_joinSep___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType_spec__0(x_15, x_16); +x_18 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__21; +x_19 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__22; +x_20 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_17); +x_21 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__23; x_22 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_22, 0, x_1); +lean_ctor_set(x_22, 0, x_20); lean_ctor_set(x_22, 1, x_21); x_23 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_23, 0, x_19); +lean_ctor_set(x_23, 0, x_18); lean_ctor_set(x_23, 1, x_22); x_24 = 0; x_25 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_25, 0, x_23); lean_ctor_set_uint8(x_25, sizeof(void*)*1, x_24); -x_26 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_26, 0, x_15); -lean_ctor_set(x_26, 1, x_25); -return x_26; +x_34 = lean_unsigned_to_nat(0u); +x_35 = lean_nat_dec_eq(x_13, x_34); +if (x_35 == 0) +{ +goto block_33; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; lean_object* x_39; lean_object* x_40; -x_27 = lean_ctor_get(x_1, 1); -lean_inc(x_27); -lean_dec(x_1); -x_28 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__19; -x_29 = lean_array_to_list(x_27); -x_30 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__17; -x_31 = l_Std_Format_joinSep___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType_spec__0(x_29, x_30); -x_32 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__23; -x_33 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__24; -x_34 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_31); -x_35 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__25; -x_36 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -x_37 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_37, 0, x_32); -lean_ctor_set(x_37, 1, x_36); -x_38 = 0; -x_39 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_39, 0, x_37); -lean_ctor_set_uint8(x_39, sizeof(void*)*1, x_38); -x_40 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_40, 0, x_28); -lean_ctor_set(x_40, 1, x_39); -return x_40; +uint8_t x_36; +x_36 = lean_nat_dec_eq(x_14, x_34); +if (x_36 == 0) +{ +goto block_33; +} +else +{ +lean_dec(x_14); +lean_dec(x_13); +return x_25; +} +} +block_33: +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_26 = l_Nat_reprFast(x_13); +x_27 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__24; +x_28 = lean_string_append(x_26, x_27); +x_29 = l_Nat_reprFast(x_14); +x_30 = lean_string_append(x_28, x_29); +lean_dec_ref(x_29); +x_31 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_31, 0, x_30); +x_32 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_25); +return x_32; } } case 11: { -uint8_t x_41; -x_41 = !lean_is_exclusive(x_1); -if (x_41 == 0) -{ -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; lean_object* x_54; lean_object* x_55; -x_42 = lean_ctor_get(x_1, 1); -x_43 = lean_ctor_get(x_1, 0); -lean_dec(x_43); -x_44 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__27; -x_45 = lean_array_to_list(x_42); -x_46 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__17; -x_47 = l_Std_Format_joinSep___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType_spec__0(x_45, x_46); -x_48 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__23; -x_49 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__24; +uint8_t x_37; +x_37 = !lean_is_exclusive(x_1); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; lean_object* x_49; +x_38 = lean_ctor_get(x_1, 1); +x_39 = lean_ctor_get(x_1, 0); +lean_dec(x_39); +x_40 = lean_array_to_list(x_38); +x_41 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__27; +x_42 = l_Std_Format_joinSep___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType_spec__0(x_40, x_41); +x_43 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__30; +x_44 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__31; lean_ctor_set_tag(x_1, 5); -lean_ctor_set(x_1, 1, x_47); -lean_ctor_set(x_1, 0, x_49); -x_50 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__25; -x_51 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_51, 0, x_1); -lean_ctor_set(x_51, 1, x_50); -x_52 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_52, 0, x_48); -lean_ctor_set(x_52, 1, x_51); -x_53 = 0; -x_54 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_54, 0, x_52); -lean_ctor_set_uint8(x_54, sizeof(void*)*1, x_53); -x_55 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_55, 0, x_44); -lean_ctor_set(x_55, 1, x_54); -return x_55; +lean_ctor_set(x_1, 1, x_42); +lean_ctor_set(x_1, 0, x_44); +x_45 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__32; +x_46 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_46, 0, x_1); +lean_ctor_set(x_46, 1, x_45); +x_47 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_47, 0, x_43); +lean_ctor_set(x_47, 1, x_46); +x_48 = 0; +x_49 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set_uint8(x_49, sizeof(void*)*1, x_48); +return x_49; } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; lean_object* x_68; lean_object* x_69; -x_56 = lean_ctor_get(x_1, 1); -lean_inc(x_56); +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; lean_object* x_61; +x_50 = lean_ctor_get(x_1, 1); +lean_inc(x_50); lean_dec(x_1); -x_57 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__27; -x_58 = lean_array_to_list(x_56); -x_59 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__17; -x_60 = l_Std_Format_joinSep___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType_spec__0(x_58, x_59); -x_61 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__23; -x_62 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__24; -x_63 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_63, 0, x_62); -lean_ctor_set(x_63, 1, x_60); -x_64 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__25; -x_65 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_65, 0, x_63); -lean_ctor_set(x_65, 1, x_64); -x_66 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_66, 0, x_61); -lean_ctor_set(x_66, 1, x_65); -x_67 = 0; -x_68 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_68, 0, x_66); -lean_ctor_set_uint8(x_68, sizeof(void*)*1, x_67); -x_69 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_69, 0, x_57); -lean_ctor_set(x_69, 1, x_68); -return x_69; +x_51 = lean_array_to_list(x_50); +x_52 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__27; +x_53 = l_Std_Format_joinSep___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType_spec__0(x_51, x_52); +x_54 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__30; +x_55 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__31; +x_56 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_53); +x_57 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__32; +x_58 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_58, 0, x_56); +lean_ctor_set(x_58, 1, x_57); +x_59 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_59, 0, x_54); +lean_ctor_set(x_59, 1, x_58); +x_60 = 0; +x_61 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_61, 0, x_59); +lean_ctor_set_uint8(x_61, sizeof(void*)*1, x_60); +return x_61; } } case 12: { -lean_object* x_70; -x_70 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__29; -return x_70; +lean_object* x_62; +x_62 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__34; +return x_62; } default: { -lean_object* x_71; -x_71 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__31; -return x_71; +lean_object* x_63; +x_63 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__36; +return x_63; } } } @@ -2447,57 +2455,21 @@ static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatP _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("(", 1, 1); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__0; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__2() { -_start: -{ -lean_object* x_1; x_1 = lean_mk_string_unchecked(" : ", 3, 3); return x_1; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__2; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked(")", 1, 1); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__5() { +static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__1() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__4; +x_1 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__0; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__6() { +static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__2() { _start: { lean_object* x_1; @@ -2515,7 +2487,7 @@ x_3 = lean_ctor_get_uint8(x_1, sizeof(void*)*2); x_4 = lean_ctor_get(x_1, 1); lean_inc(x_4); lean_dec_ref(x_1); -x_5 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__1; +x_5 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__31; x_6 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; x_7 = l_Nat_reprFast(x_2); x_8 = lean_string_append(x_6, x_7); @@ -2525,7 +2497,7 @@ lean_ctor_set(x_9, 0, x_8); x_10 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_10, 0, x_5); lean_ctor_set(x_10, 1, x_9); -x_11 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__3; +x_11 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__1; x_12 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_12, 0, x_10); lean_ctor_set(x_12, 1, x_11); @@ -2539,7 +2511,7 @@ goto block_20; else { lean_object* x_22; -x_22 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__6; +x_22 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__2; x_13 = x_22; goto block_20; } @@ -2555,7 +2527,7 @@ x_16 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType(x_4); x_17 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_17, 0, x_15); lean_ctor_set(x_17, 1, x_16); -x_18 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__5; +x_18 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__32; x_19 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_19, 0, x_17); lean_ctor_set(x_19, 1, x_18); @@ -3193,7 +3165,7 @@ lean_ctor_set(x_9, 0, x_8); x_10 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_10, 0, x_5); lean_ctor_set(x_10, 1, x_9); -x_11 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__3; +x_11 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__1; x_12 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_12, 0, x_10); lean_ctor_set(x_12, 1, x_11); @@ -3308,379 +3280,402 @@ return x_59; } case 4: { -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; x_60 = lean_ctor_get(x_1, 0); lean_inc(x_60); x_61 = lean_ctor_get(x_1, 1); lean_inc(x_61); x_62 = lean_ctor_get(x_1, 2); lean_inc(x_62); +x_63 = lean_ctor_get(x_1, 3); +lean_inc(x_63); lean_dec_ref(x_1); -x_63 = l_Lean_IR_formatFnBodyHead___closed__14; -x_64 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; -x_65 = l_Nat_reprFast(x_60); -x_66 = lean_string_append(x_64, x_65); -lean_dec_ref(x_65); -x_67 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_67, 0, x_66); -x_68 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_68, 0, x_63); -lean_ctor_set(x_68, 1, x_67); -x_69 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__1; -x_70 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_69); -x_71 = l_Nat_reprFast(x_61); -x_72 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_72, 0, x_71); -x_73 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_73, 0, x_70); -lean_ctor_set(x_73, 1, x_72); -x_74 = l_Lean_IR_formatFnBodyHead___closed__10; -x_75 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_75, 0, x_73); -lean_ctor_set(x_75, 1, x_74); -x_76 = l_Nat_reprFast(x_62); -x_77 = lean_string_append(x_64, x_76); -lean_dec_ref(x_76); +x_64 = l_Lean_IR_formatFnBodyHead___closed__14; +x_65 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_66 = l_Nat_reprFast(x_60); +x_67 = lean_string_append(x_65, x_66); +lean_dec_ref(x_66); +x_68 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_68, 0, x_67); +x_69 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_69, 0, x_64); +lean_ctor_set(x_69, 1, x_68); +x_70 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__1; +x_71 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_71, 0, x_69); +lean_ctor_set(x_71, 1, x_70); +x_72 = l_Nat_reprFast(x_61); +x_73 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_73, 0, x_72); +x_74 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_74, 0, x_71); +lean_ctor_set(x_74, 1, x_73); +x_75 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__13; +x_76 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_76, 0, x_74); +lean_ctor_set(x_76, 1, x_75); +x_77 = l_Nat_reprFast(x_62); x_78 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_78, 0, x_77); x_79 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_79, 0, x_75); +lean_ctor_set(x_79, 0, x_76); lean_ctor_set(x_79, 1, x_78); -return x_79; +x_80 = l_Lean_IR_formatFnBodyHead___closed__10; +x_81 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_81, 0, x_79); +lean_ctor_set(x_81, 1, x_80); +x_82 = l_Nat_reprFast(x_63); +x_83 = lean_string_append(x_65, x_82); +lean_dec_ref(x_82); +x_84 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_84, 0, x_83); +x_85 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_85, 0, x_81); +lean_ctor_set(x_85, 1, x_84); +return x_85; } case 5: { -lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_80 = lean_ctor_get(x_1, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_1, 1); -lean_inc(x_81); -x_82 = lean_ctor_get(x_1, 2); -lean_inc(x_82); -x_83 = lean_ctor_get(x_1, 3); -lean_inc(x_83); -x_84 = lean_ctor_get(x_1, 4); -lean_inc(x_84); +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; +x_86 = lean_ctor_get(x_1, 0); +lean_inc(x_86); +x_87 = lean_ctor_get(x_1, 1); +lean_inc(x_87); +x_88 = lean_ctor_get(x_1, 2); +lean_inc(x_88); +x_89 = lean_ctor_get(x_1, 3); +lean_inc(x_89); +x_90 = lean_ctor_get(x_1, 4); +lean_inc(x_90); +x_91 = lean_ctor_get(x_1, 5); +lean_inc(x_91); lean_dec_ref(x_1); -x_85 = l_Lean_IR_formatFnBodyHead___closed__16; -x_86 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; -x_87 = l_Nat_reprFast(x_80); -x_88 = lean_string_append(x_86, x_87); -lean_dec_ref(x_87); -x_89 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_89, 0, x_88); -x_90 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_90, 0, x_85); -lean_ctor_set(x_90, 1, x_89); -x_91 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__1; -x_92 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_92, 0, x_90); -lean_ctor_set(x_92, 1, x_91); -x_93 = l_Nat_reprFast(x_81); +x_92 = l_Lean_IR_formatFnBodyHead___closed__16; +x_93 = l_Nat_reprFast(x_87); x_94 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_94, 0, x_93); x_95 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_95, 0, x_92); lean_ctor_set(x_95, 1, x_94); -x_96 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__17; +x_96 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__1; x_97 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_97, 0, x_95); lean_ctor_set(x_97, 1, x_96); -x_98 = l_Nat_reprFast(x_82); -x_99 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_99, 0, x_98); -x_100 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_100, 0, x_97); -lean_ctor_set(x_100, 1, x_99); -x_101 = l_Lean_IR_formatFnBodyHead___closed__18; +x_98 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_99 = l_Nat_reprFast(x_86); +x_100 = lean_string_append(x_98, x_99); +lean_dec_ref(x_99); +x_101 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_101, 0, x_100); x_102 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_102, 0, x_100); +lean_ctor_set(x_102, 0, x_97); lean_ctor_set(x_102, 1, x_101); -x_103 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType(x_84); +x_103 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__13; x_104 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_104, 0, x_102); lean_ctor_set(x_104, 1, x_103); -x_105 = l_Lean_IR_formatFnBodyHead___closed__3; -x_106 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_106, 0, x_104); -lean_ctor_set(x_106, 1, x_105); -x_107 = l_Nat_reprFast(x_83); -x_108 = lean_string_append(x_86, x_107); -lean_dec_ref(x_107); -x_109 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_109, 0, x_108); -x_110 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_110, 0, x_106); -lean_ctor_set(x_110, 1, x_109); -return x_110; -} -case 6: -{ -lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_124; uint8_t x_125; -x_111 = lean_ctor_get(x_1, 0); -lean_inc(x_111); -x_112 = lean_ctor_get(x_1, 1); -lean_inc(x_112); -lean_dec_ref(x_1); -x_113 = l_Lean_IR_formatFnBodyHead___closed__20; -x_124 = lean_unsigned_to_nat(1u); -x_125 = lean_nat_dec_eq(x_112, x_124); -if (x_125 == 0) -{ -lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; uint8_t x_134; lean_object* x_135; -x_126 = l_Nat_reprFast(x_112); -x_127 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_127, 0, x_126); -x_128 = l_Lean_IR_formatFnBodyHead___closed__22; -x_129 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__1; -x_130 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_130, 0, x_129); -lean_ctor_set(x_130, 1, x_127); -x_131 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__3; -x_132 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_132, 0, x_130); -lean_ctor_set(x_132, 1, x_131); -x_133 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_133, 0, x_128); -lean_ctor_set(x_133, 1, x_132); -x_134 = 0; -x_135 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_135, 0, x_133); -lean_ctor_set_uint8(x_135, sizeof(void*)*1, x_134); -x_114 = x_135; -goto block_123; -} -else -{ -lean_object* x_136; -lean_dec(x_112); -x_136 = l_Lean_IR_formatFnBodyHead___closed__23; -x_114 = x_136; -goto block_123; -} -block_123: -{ -lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; +x_105 = l_Nat_reprFast(x_88); +x_106 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_106, 0, x_105); +x_107 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_107, 0, x_104); +lean_ctor_set(x_107, 1, x_106); +x_108 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_108, 0, x_107); +lean_ctor_set(x_108, 1, x_103); +x_109 = l_Nat_reprFast(x_89); +x_110 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_110, 0, x_109); +x_111 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_110); +x_112 = l_Lean_IR_formatFnBodyHead___closed__18; +x_113 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_113, 0, x_111); +lean_ctor_set(x_113, 1, x_112); +x_114 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType(x_91); x_115 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_115, 0, x_113); lean_ctor_set(x_115, 1, x_114); -x_116 = l_Lean_IR_formatArray___redArg___lam__0___closed__1; +x_116 = l_Lean_IR_formatFnBodyHead___closed__3; x_117 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_117, 0, x_115); lean_ctor_set(x_117, 1, x_116); -x_118 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; -x_119 = l_Nat_reprFast(x_111); -x_120 = lean_string_append(x_118, x_119); -lean_dec_ref(x_119); -x_121 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_121, 0, x_120); -x_122 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_122, 0, x_117); -lean_ctor_set(x_122, 1, x_121); -return x_122; -} +x_118 = l_Nat_reprFast(x_90); +x_119 = lean_string_append(x_98, x_118); +lean_dec_ref(x_118); +x_120 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_120, 0, x_119); +x_121 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_121, 0, x_117); +lean_ctor_set(x_121, 1, x_120); +return x_121; } -case 7: +case 6: { -lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_150; uint8_t x_151; -x_137 = lean_ctor_get(x_1, 0); -lean_inc(x_137); -x_138 = lean_ctor_get(x_1, 1); -lean_inc(x_138); +lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_135; uint8_t x_136; +x_122 = lean_ctor_get(x_1, 0); +lean_inc(x_122); +x_123 = lean_ctor_get(x_1, 1); +lean_inc(x_123); lean_dec_ref(x_1); -x_139 = l_Lean_IR_formatFnBodyHead___closed__25; -x_150 = lean_unsigned_to_nat(1u); -x_151 = lean_nat_dec_eq(x_138, x_150); -if (x_151 == 0) +x_124 = l_Lean_IR_formatFnBodyHead___closed__20; +x_135 = lean_unsigned_to_nat(1u); +x_136 = lean_nat_dec_eq(x_123, x_135); +if (x_136 == 0) { -lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; uint8_t x_160; lean_object* x_161; -x_152 = l_Nat_reprFast(x_138); -x_153 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_153, 0, x_152); -x_154 = l_Lean_IR_formatFnBodyHead___closed__22; -x_155 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__1; -x_156 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_156, 0, x_155); -lean_ctor_set(x_156, 1, x_153); -x_157 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__3; -x_158 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_158, 0, x_156); -lean_ctor_set(x_158, 1, x_157); -x_159 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_159, 0, x_154); -lean_ctor_set(x_159, 1, x_158); -x_160 = 0; -x_161 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_161, 0, x_159); -lean_ctor_set_uint8(x_161, sizeof(void*)*1, x_160); -x_140 = x_161; -goto block_149; +lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; uint8_t x_145; lean_object* x_146; +x_137 = l_Nat_reprFast(x_123); +x_138 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_138, 0, x_137); +x_139 = l_Lean_IR_formatFnBodyHead___closed__22; +x_140 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__1; +x_141 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_141, 0, x_140); +lean_ctor_set(x_141, 1, x_138); +x_142 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__3; +x_143 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_143, 0, x_141); +lean_ctor_set(x_143, 1, x_142); +x_144 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_144, 0, x_139); +lean_ctor_set(x_144, 1, x_143); +x_145 = 0; +x_146 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_146, 0, x_144); +lean_ctor_set_uint8(x_146, sizeof(void*)*1, x_145); +x_125 = x_146; +goto block_134; } else { -lean_object* x_162; -lean_dec(x_138); -x_162 = l_Lean_IR_formatFnBodyHead___closed__23; -x_140 = x_162; -goto block_149; +lean_object* x_147; +lean_dec(x_123); +x_147 = l_Lean_IR_formatFnBodyHead___closed__23; +x_125 = x_147; +goto block_134; +} +block_134: +{ +lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; +x_126 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_126, 0, x_124); +lean_ctor_set(x_126, 1, x_125); +x_127 = l_Lean_IR_formatArray___redArg___lam__0___closed__1; +x_128 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_128, 0, x_126); +lean_ctor_set(x_128, 1, x_127); +x_129 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_130 = l_Nat_reprFast(x_122); +x_131 = lean_string_append(x_129, x_130); +lean_dec_ref(x_130); +x_132 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_132, 0, x_131); +x_133 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_133, 0, x_128); +lean_ctor_set(x_133, 1, x_132); +return x_133; } -block_149: +} +case 7: { -lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; -x_141 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_141, 0, x_139); -lean_ctor_set(x_141, 1, x_140); -x_142 = l_Lean_IR_formatArray___redArg___lam__0___closed__1; -x_143 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_143, 0, x_141); -lean_ctor_set(x_143, 1, x_142); -x_144 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; -x_145 = l_Nat_reprFast(x_137); -x_146 = lean_string_append(x_144, x_145); -lean_dec_ref(x_145); -x_147 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_147, 0, x_146); -x_148 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_148, 0, x_143); -lean_ctor_set(x_148, 1, x_147); -return x_148; +lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_161; uint8_t x_162; +x_148 = lean_ctor_get(x_1, 0); +lean_inc(x_148); +x_149 = lean_ctor_get(x_1, 1); +lean_inc(x_149); +lean_dec_ref(x_1); +x_150 = l_Lean_IR_formatFnBodyHead___closed__25; +x_161 = lean_unsigned_to_nat(1u); +x_162 = lean_nat_dec_eq(x_149, x_161); +if (x_162 == 0) +{ +lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; uint8_t x_171; lean_object* x_172; +x_163 = l_Nat_reprFast(x_149); +x_164 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_164, 0, x_163); +x_165 = l_Lean_IR_formatFnBodyHead___closed__22; +x_166 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__1; +x_167 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_167, 0, x_166); +lean_ctor_set(x_167, 1, x_164); +x_168 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__3; +x_169 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_169, 0, x_167); +lean_ctor_set(x_169, 1, x_168); +x_170 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_170, 0, x_165); +lean_ctor_set(x_170, 1, x_169); +x_171 = 0; +x_172 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_172, 0, x_170); +lean_ctor_set_uint8(x_172, sizeof(void*)*1, x_171); +x_151 = x_172; +goto block_160; +} +else +{ +lean_object* x_173; +lean_dec(x_149); +x_173 = l_Lean_IR_formatFnBodyHead___closed__23; +x_151 = x_173; +goto block_160; +} +block_160: +{ +lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; +x_152 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_152, 0, x_150); +lean_ctor_set(x_152, 1, x_151); +x_153 = l_Lean_IR_formatArray___redArg___lam__0___closed__1; +x_154 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_154, 0, x_152); +lean_ctor_set(x_154, 1, x_153); +x_155 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_156 = l_Nat_reprFast(x_148); +x_157 = lean_string_append(x_155, x_156); +lean_dec_ref(x_156); +x_158 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_158, 0, x_157); +x_159 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_159, 0, x_154); +lean_ctor_set(x_159, 1, x_158); +return x_159; } } case 8: { -uint8_t x_163; -x_163 = !lean_is_exclusive(x_1); -if (x_163 == 0) -{ -lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; -x_164 = lean_ctor_get(x_1, 0); -x_165 = lean_ctor_get(x_1, 1); -lean_dec(x_165); -x_166 = l_Lean_IR_formatFnBodyHead___closed__27; -x_167 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; -x_168 = l_Nat_reprFast(x_164); -x_169 = lean_string_append(x_167, x_168); -lean_dec_ref(x_168); -x_170 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_170, 0, x_169); +uint8_t x_174; +x_174 = !lean_is_exclusive(x_1); +if (x_174 == 0) +{ +lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; +x_175 = lean_ctor_get(x_1, 0); +x_176 = lean_ctor_get(x_1, 1); +lean_dec(x_176); +x_177 = l_Lean_IR_formatFnBodyHead___closed__27; +x_178 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_179 = l_Nat_reprFast(x_175); +x_180 = lean_string_append(x_178, x_179); +lean_dec_ref(x_179); +x_181 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_181, 0, x_180); lean_ctor_set_tag(x_1, 5); -lean_ctor_set(x_1, 1, x_170); -lean_ctor_set(x_1, 0, x_166); +lean_ctor_set(x_1, 1, x_181); +lean_ctor_set(x_1, 0, x_177); return x_1; } else { -lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; -x_171 = lean_ctor_get(x_1, 0); -lean_inc(x_171); +lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; +x_182 = lean_ctor_get(x_1, 0); +lean_inc(x_182); lean_dec(x_1); -x_172 = l_Lean_IR_formatFnBodyHead___closed__27; -x_173 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; -x_174 = l_Nat_reprFast(x_171); -x_175 = lean_string_append(x_173, x_174); -lean_dec_ref(x_174); -x_176 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_176, 0, x_175); -x_177 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_177, 0, x_172); -lean_ctor_set(x_177, 1, x_176); -return x_177; +x_183 = l_Lean_IR_formatFnBodyHead___closed__27; +x_184 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_185 = l_Nat_reprFast(x_182); +x_186 = lean_string_append(x_184, x_185); +lean_dec_ref(x_185); +x_187 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_187, 0, x_186); +x_188 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_188, 0, x_183); +lean_ctor_set(x_188, 1, x_187); +return x_188; } } case 9: { -lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; -x_178 = lean_ctor_get(x_1, 1); -lean_inc(x_178); +lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; +x_189 = lean_ctor_get(x_1, 1); +lean_inc(x_189); lean_dec_ref(x_1); -x_179 = l_Lean_IR_formatFnBodyHead___closed__29; -x_180 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; -x_181 = l_Nat_reprFast(x_178); -x_182 = lean_string_append(x_180, x_181); -lean_dec_ref(x_181); -x_183 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_183, 0, x_182); -x_184 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_184, 0, x_179); -lean_ctor_set(x_184, 1, x_183); -x_185 = l_Lean_IR_formatFnBodyHead___closed__31; -x_186 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_186, 0, x_184); -lean_ctor_set(x_186, 1, x_185); -return x_186; +x_190 = l_Lean_IR_formatFnBodyHead___closed__29; +x_191 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_192 = l_Nat_reprFast(x_189); +x_193 = lean_string_append(x_191, x_192); +lean_dec_ref(x_192); +x_194 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_194, 0, x_193); +x_195 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_195, 0, x_190); +lean_ctor_set(x_195, 1, x_194); +x_196 = l_Lean_IR_formatFnBodyHead___closed__31; +x_197 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_197, 0, x_195); +lean_ctor_set(x_197, 1, x_196); +return x_197; } case 10: { -lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; -x_187 = lean_ctor_get(x_1, 0); -lean_inc(x_187); +lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; +x_198 = lean_ctor_get(x_1, 0); +lean_inc(x_198); lean_dec_ref(x_1); -x_188 = l_Lean_IR_formatFnBodyHead___closed__33; -x_189 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg(x_187); -x_190 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_190, 0, x_188); -lean_ctor_set(x_190, 1, x_189); -return x_190; +x_199 = l_Lean_IR_formatFnBodyHead___closed__33; +x_200 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg(x_198); +x_201 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_201, 0, x_199); +lean_ctor_set(x_201, 1, x_200); +return x_201; } case 11: { -uint8_t x_191; -x_191 = !lean_is_exclusive(x_1); -if (x_191 == 0) -{ -lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; -x_192 = lean_ctor_get(x_1, 0); -x_193 = lean_ctor_get(x_1, 1); -x_194 = l_Lean_IR_formatFnBodyHead___closed__35; -x_195 = l_Lean_IR_formatFnBodyHead___closed__4; -x_196 = l_Nat_reprFast(x_192); -x_197 = lean_string_append(x_195, x_196); -lean_dec_ref(x_196); -x_198 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_198, 0, x_197); +uint8_t x_202; +x_202 = !lean_is_exclusive(x_1); +if (x_202 == 0) +{ +lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; +x_203 = lean_ctor_get(x_1, 0); +x_204 = lean_ctor_get(x_1, 1); +x_205 = l_Lean_IR_formatFnBodyHead___closed__35; +x_206 = l_Lean_IR_formatFnBodyHead___closed__4; +x_207 = l_Nat_reprFast(x_203); +x_208 = lean_string_append(x_206, x_207); +lean_dec_ref(x_207); +x_209 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_209, 0, x_208); lean_ctor_set_tag(x_1, 5); -lean_ctor_set(x_1, 1, x_198); -lean_ctor_set(x_1, 0, x_194); -x_199 = l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(x_193); -lean_dec_ref(x_193); -x_200 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_200, 0, x_1); -lean_ctor_set(x_200, 1, x_199); -return x_200; +lean_ctor_set(x_1, 1, x_209); +lean_ctor_set(x_1, 0, x_205); +x_210 = l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(x_204); +lean_dec_ref(x_204); +x_211 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_211, 0, x_1); +lean_ctor_set(x_211, 1, x_210); +return x_211; } else { -lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; -x_201 = lean_ctor_get(x_1, 0); -x_202 = lean_ctor_get(x_1, 1); -lean_inc(x_202); -lean_inc(x_201); +lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; +x_212 = lean_ctor_get(x_1, 0); +x_213 = lean_ctor_get(x_1, 1); +lean_inc(x_213); +lean_inc(x_212); lean_dec(x_1); -x_203 = l_Lean_IR_formatFnBodyHead___closed__35; -x_204 = l_Lean_IR_formatFnBodyHead___closed__4; -x_205 = l_Nat_reprFast(x_201); -x_206 = lean_string_append(x_204, x_205); -lean_dec_ref(x_205); -x_207 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_207, 0, x_206); -x_208 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_208, 0, x_203); -lean_ctor_set(x_208, 1, x_207); -x_209 = l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(x_202); -lean_dec_ref(x_202); -x_210 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_210, 0, x_208); -lean_ctor_set(x_210, 1, x_209); -return x_210; +x_214 = l_Lean_IR_formatFnBodyHead___closed__35; +x_215 = l_Lean_IR_formatFnBodyHead___closed__4; +x_216 = l_Nat_reprFast(x_212); +x_217 = lean_string_append(x_215, x_216); +lean_dec_ref(x_216); +x_218 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_218, 0, x_217); +x_219 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_219, 0, x_214); +lean_ctor_set(x_219, 1, x_218); +x_220 = l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(x_213); +lean_dec_ref(x_213); +x_221 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_221, 0, x_219); +lean_ctor_set(x_221, 1, x_220); +return x_221; } } default: { -lean_object* x_211; -x_211 = l_Lean_IR_formatFnBodyHead___closed__37; -return x_211; +lean_object* x_222; +x_222 = l_Lean_IR_formatFnBodyHead___closed__37; +return x_222; } } } @@ -3776,7 +3771,7 @@ lean_ctor_set(x_11, 0, x_10); x_12 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_12, 0, x_7); lean_ctor_set(x_12, 1, x_11); -x_13 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__3; +x_13 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__1; x_14 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_14, 0, x_12); lean_ctor_set(x_14, 1, x_13); @@ -3960,7 +3955,7 @@ return x_94; } case 4: { -lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; +lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; x_95 = lean_ctor_get(x_2, 0); lean_inc(x_95); x_96 = lean_ctor_get(x_2, 1); @@ -3969,521 +3964,544 @@ x_97 = lean_ctor_get(x_2, 2); lean_inc(x_97); x_98 = lean_ctor_get(x_2, 3); lean_inc(x_98); +x_99 = lean_ctor_get(x_2, 4); +lean_inc(x_99); lean_dec_ref(x_2); -x_99 = l_Lean_IR_formatFnBodyHead___closed__14; -x_100 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; -x_101 = l_Nat_reprFast(x_95); -x_102 = lean_string_append(x_100, x_101); -lean_dec_ref(x_101); -x_103 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_103, 0, x_102); -x_104 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_104, 0, x_99); -lean_ctor_set(x_104, 1, x_103); -x_105 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__1; -x_106 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_106, 0, x_104); -lean_ctor_set(x_106, 1, x_105); -x_107 = l_Nat_reprFast(x_96); -x_108 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_108, 0, x_107); -x_109 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_109, 0, x_106); -lean_ctor_set(x_109, 1, x_108); -x_110 = l_Lean_IR_formatFnBodyHead___closed__10; -x_111 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_111, 0, x_109); -lean_ctor_set(x_111, 1, x_110); -x_112 = l_Nat_reprFast(x_97); -x_113 = lean_string_append(x_100, x_112); -lean_dec_ref(x_112); +x_100 = l_Lean_IR_formatFnBodyHead___closed__14; +x_101 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_102 = l_Nat_reprFast(x_95); +x_103 = lean_string_append(x_101, x_102); +lean_dec_ref(x_102); +x_104 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_104, 0, x_103); +x_105 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_105, 0, x_100); +lean_ctor_set(x_105, 1, x_104); +x_106 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__1; +x_107 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_107, 0, x_105); +lean_ctor_set(x_107, 1, x_106); +x_108 = l_Nat_reprFast(x_96); +x_109 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_109, 0, x_108); +x_110 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_109); +x_111 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__13; +x_112 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_112, 0, x_110); +lean_ctor_set(x_112, 1, x_111); +x_113 = l_Nat_reprFast(x_97); x_114 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_114, 0, x_113); x_115 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_115, 0, x_111); +lean_ctor_set(x_115, 0, x_112); lean_ctor_set(x_115, 1, x_114); -x_116 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop___closed__1; +x_116 = l_Lean_IR_formatFnBodyHead___closed__10; x_117 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_117, 0, x_115); lean_ctor_set(x_117, 1, x_116); -x_118 = lean_box(1); -x_119 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_119, 0, x_117); -lean_ctor_set(x_119, 1, x_118); -x_120 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop(x_1, x_98); +x_118 = l_Nat_reprFast(x_98); +x_119 = lean_string_append(x_101, x_118); +lean_dec_ref(x_118); +x_120 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_120, 0, x_119); x_121 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_121, 0, x_119); +lean_ctor_set(x_121, 0, x_117); lean_ctor_set(x_121, 1, x_120); -return x_121; +x_122 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop___closed__1; +x_123 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_123, 0, x_121); +lean_ctor_set(x_123, 1, x_122); +x_124 = lean_box(1); +x_125 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_125, 0, x_123); +lean_ctor_set(x_125, 1, x_124); +x_126 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop(x_1, x_99); +x_127 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; } case 5: { -lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; -x_122 = lean_ctor_get(x_2, 0); -lean_inc(x_122); -x_123 = lean_ctor_get(x_2, 1); -lean_inc(x_123); -x_124 = lean_ctor_get(x_2, 2); -lean_inc(x_124); -x_125 = lean_ctor_get(x_2, 3); -lean_inc(x_125); -x_126 = lean_ctor_get(x_2, 4); -lean_inc(x_126); -x_127 = lean_ctor_get(x_2, 5); -lean_inc(x_127); +lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; +x_128 = lean_ctor_get(x_2, 0); +lean_inc(x_128); +x_129 = lean_ctor_get(x_2, 1); +lean_inc(x_129); +x_130 = lean_ctor_get(x_2, 2); +lean_inc(x_130); +x_131 = lean_ctor_get(x_2, 3); +lean_inc(x_131); +x_132 = lean_ctor_get(x_2, 4); +lean_inc(x_132); +x_133 = lean_ctor_get(x_2, 5); +lean_inc(x_133); +x_134 = lean_ctor_get(x_2, 6); +lean_inc(x_134); lean_dec_ref(x_2); -x_128 = l_Lean_IR_formatFnBodyHead___closed__16; -x_129 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; -x_130 = l_Nat_reprFast(x_122); -x_131 = lean_string_append(x_129, x_130); -lean_dec_ref(x_130); -x_132 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_132, 0, x_131); -x_133 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_133, 0, x_128); -lean_ctor_set(x_133, 1, x_132); -x_134 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__1; -x_135 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_135, 0, x_133); -lean_ctor_set(x_135, 1, x_134); -x_136 = l_Nat_reprFast(x_123); -x_137 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_137, 0, x_136); -x_138 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_138, 0, x_135); -lean_ctor_set(x_138, 1, x_137); -x_139 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__17; +x_135 = l_Lean_IR_formatFnBodyHead___closed__16; +x_136 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_137 = l_Nat_reprFast(x_128); +x_138 = lean_string_append(x_136, x_137); +lean_dec_ref(x_137); +x_139 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_139, 0, x_138); x_140 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_140, 0, x_138); +lean_ctor_set(x_140, 0, x_135); lean_ctor_set(x_140, 1, x_139); -x_141 = l_Nat_reprFast(x_124); -x_142 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_142, 0, x_141); -x_143 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_143, 0, x_140); -lean_ctor_set(x_143, 1, x_142); -x_144 = l_Lean_IR_formatFnBodyHead___closed__18; +x_141 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__1; +x_142 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_142, 0, x_140); +lean_ctor_set(x_142, 1, x_141); +x_143 = l_Nat_reprFast(x_129); +x_144 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_144, 0, x_143); x_145 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_145, 0, x_143); +lean_ctor_set(x_145, 0, x_142); lean_ctor_set(x_145, 1, x_144); -x_146 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType(x_126); +x_146 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__13; x_147 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_147, 0, x_145); lean_ctor_set(x_147, 1, x_146); -x_148 = l_Lean_IR_formatFnBodyHead___closed__3; -x_149 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_149, 0, x_147); -lean_ctor_set(x_149, 1, x_148); -x_150 = l_Nat_reprFast(x_125); -x_151 = lean_string_append(x_129, x_150); -lean_dec_ref(x_150); -x_152 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_152, 0, x_151); -x_153 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_153, 0, x_149); -lean_ctor_set(x_153, 1, x_152); -x_154 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop___closed__1; -x_155 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_155, 0, x_153); -lean_ctor_set(x_155, 1, x_154); -x_156 = lean_box(1); -x_157 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_157, 0, x_155); -lean_ctor_set(x_157, 1, x_156); -x_158 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop(x_1, x_127); -x_159 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_159, 0, x_157); -lean_ctor_set(x_159, 1, x_158); -return x_159; +x_148 = l_Nat_reprFast(x_130); +x_149 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_149, 0, x_148); +x_150 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_150, 0, x_147); +lean_ctor_set(x_150, 1, x_149); +x_151 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_151, 0, x_150); +lean_ctor_set(x_151, 1, x_146); +x_152 = l_Nat_reprFast(x_131); +x_153 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_153, 0, x_152); +x_154 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_154, 0, x_151); +lean_ctor_set(x_154, 1, x_153); +x_155 = l_Lean_IR_formatFnBodyHead___closed__18; +x_156 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_156, 0, x_154); +lean_ctor_set(x_156, 1, x_155); +x_157 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType(x_133); +x_158 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_158, 0, x_156); +lean_ctor_set(x_158, 1, x_157); +x_159 = l_Lean_IR_formatFnBodyHead___closed__3; +x_160 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_160, 0, x_158); +lean_ctor_set(x_160, 1, x_159); +x_161 = l_Nat_reprFast(x_132); +x_162 = lean_string_append(x_136, x_161); +lean_dec_ref(x_161); +x_163 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_163, 0, x_162); +x_164 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_164, 0, x_160); +lean_ctor_set(x_164, 1, x_163); +x_165 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop___closed__1; +x_166 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_166, 0, x_164); +lean_ctor_set(x_166, 1, x_165); +x_167 = lean_box(1); +x_168 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_168, 0, x_166); +lean_ctor_set(x_168, 1, x_167); +x_169 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop(x_1, x_134); +x_170 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_170, 0, x_168); +lean_ctor_set(x_170, 1, x_169); +return x_170; } case 6: { -lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_180; uint8_t x_181; -x_160 = lean_ctor_get(x_2, 0); -lean_inc(x_160); -x_161 = lean_ctor_get(x_2, 1); -lean_inc(x_161); -x_162 = lean_ctor_get(x_2, 2); -lean_inc(x_162); +lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_191; uint8_t x_192; +x_171 = lean_ctor_get(x_2, 0); +lean_inc(x_171); +x_172 = lean_ctor_get(x_2, 1); +lean_inc(x_172); +x_173 = lean_ctor_get(x_2, 2); +lean_inc(x_173); lean_dec_ref(x_2); -x_163 = l_Lean_IR_formatFnBodyHead___closed__20; -x_180 = lean_unsigned_to_nat(1u); -x_181 = lean_nat_dec_eq(x_161, x_180); -if (x_181 == 0) -{ -lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; uint8_t x_190; lean_object* x_191; -x_182 = l_Nat_reprFast(x_161); -x_183 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_183, 0, x_182); -x_184 = l_Lean_IR_formatFnBodyHead___closed__22; -x_185 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__1; -x_186 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_186, 0, x_185); -lean_ctor_set(x_186, 1, x_183); -x_187 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__3; -x_188 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_188, 0, x_186); -lean_ctor_set(x_188, 1, x_187); -x_189 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_189, 0, x_184); -lean_ctor_set(x_189, 1, x_188); -x_190 = 0; -x_191 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_191, 0, x_189); -lean_ctor_set_uint8(x_191, sizeof(void*)*1, x_190); -x_164 = x_191; -goto block_179; +x_174 = l_Lean_IR_formatFnBodyHead___closed__20; +x_191 = lean_unsigned_to_nat(1u); +x_192 = lean_nat_dec_eq(x_172, x_191); +if (x_192 == 0) +{ +lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; uint8_t x_201; lean_object* x_202; +x_193 = l_Nat_reprFast(x_172); +x_194 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_194, 0, x_193); +x_195 = l_Lean_IR_formatFnBodyHead___closed__22; +x_196 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__1; +x_197 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_197, 0, x_196); +lean_ctor_set(x_197, 1, x_194); +x_198 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__3; +x_199 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_199, 0, x_197); +lean_ctor_set(x_199, 1, x_198); +x_200 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_200, 0, x_195); +lean_ctor_set(x_200, 1, x_199); +x_201 = 0; +x_202 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_202, 0, x_200); +lean_ctor_set_uint8(x_202, sizeof(void*)*1, x_201); +x_175 = x_202; +goto block_190; } else { -lean_object* x_192; -lean_dec(x_161); -x_192 = l_Lean_IR_formatFnBodyHead___closed__23; -x_164 = x_192; -goto block_179; +lean_object* x_203; +lean_dec(x_172); +x_203 = l_Lean_IR_formatFnBodyHead___closed__23; +x_175 = x_203; +goto block_190; } -block_179: +block_190: { -lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; -x_165 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_165, 0, x_163); -lean_ctor_set(x_165, 1, x_164); -x_166 = l_Lean_IR_formatArray___redArg___lam__0___closed__1; -x_167 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_167, 0, x_165); -lean_ctor_set(x_167, 1, x_166); -x_168 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; -x_169 = l_Nat_reprFast(x_160); -x_170 = lean_string_append(x_168, x_169); -lean_dec_ref(x_169); -x_171 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_171, 0, x_170); -x_172 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_172, 0, x_167); -lean_ctor_set(x_172, 1, x_171); -x_173 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop___closed__1; -x_174 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_174, 0, x_172); -lean_ctor_set(x_174, 1, x_173); -x_175 = lean_box(1); +lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; x_176 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_176, 0, x_174); lean_ctor_set(x_176, 1, x_175); -x_177 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop(x_1, x_162); +x_177 = l_Lean_IR_formatArray___redArg___lam__0___closed__1; x_178 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_178, 0, x_176); lean_ctor_set(x_178, 1, x_177); -return x_178; +x_179 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_180 = l_Nat_reprFast(x_171); +x_181 = lean_string_append(x_179, x_180); +lean_dec_ref(x_180); +x_182 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_182, 0, x_181); +x_183 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_183, 0, x_178); +lean_ctor_set(x_183, 1, x_182); +x_184 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop___closed__1; +x_185 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_185, 0, x_183); +lean_ctor_set(x_185, 1, x_184); +x_186 = lean_box(1); +x_187 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_187, 0, x_185); +lean_ctor_set(x_187, 1, x_186); +x_188 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop(x_1, x_173); +x_189 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_189, 0, x_187); +lean_ctor_set(x_189, 1, x_188); +return x_189; } } case 7: { -lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_213; uint8_t x_214; -x_193 = lean_ctor_get(x_2, 0); -lean_inc(x_193); -x_194 = lean_ctor_get(x_2, 1); -lean_inc(x_194); -x_195 = lean_ctor_get(x_2, 2); -lean_inc(x_195); +lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_224; uint8_t x_225; +x_204 = lean_ctor_get(x_2, 0); +lean_inc(x_204); +x_205 = lean_ctor_get(x_2, 1); +lean_inc(x_205); +x_206 = lean_ctor_get(x_2, 2); +lean_inc(x_206); lean_dec_ref(x_2); -x_196 = l_Lean_IR_formatFnBodyHead___closed__25; -x_213 = lean_unsigned_to_nat(1u); -x_214 = lean_nat_dec_eq(x_194, x_213); -if (x_214 == 0) -{ -lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; uint8_t x_223; lean_object* x_224; -x_215 = l_Nat_reprFast(x_194); -x_216 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_216, 0, x_215); -x_217 = l_Lean_IR_formatFnBodyHead___closed__22; -x_218 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__1; -x_219 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_219, 0, x_218); -lean_ctor_set(x_219, 1, x_216); -x_220 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__3; -x_221 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_221, 0, x_219); -lean_ctor_set(x_221, 1, x_220); -x_222 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_222, 0, x_217); -lean_ctor_set(x_222, 1, x_221); -x_223 = 0; -x_224 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_224, 0, x_222); -lean_ctor_set_uint8(x_224, sizeof(void*)*1, x_223); -x_197 = x_224; -goto block_212; +x_207 = l_Lean_IR_formatFnBodyHead___closed__25; +x_224 = lean_unsigned_to_nat(1u); +x_225 = lean_nat_dec_eq(x_205, x_224); +if (x_225 == 0) +{ +lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; uint8_t x_234; lean_object* x_235; +x_226 = l_Nat_reprFast(x_205); +x_227 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_227, 0, x_226); +x_228 = l_Lean_IR_formatFnBodyHead___closed__22; +x_229 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__1; +x_230 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_230, 0, x_229); +lean_ctor_set(x_230, 1, x_227); +x_231 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__3; +x_232 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_232, 0, x_230); +lean_ctor_set(x_232, 1, x_231); +x_233 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_233, 0, x_228); +lean_ctor_set(x_233, 1, x_232); +x_234 = 0; +x_235 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_235, 0, x_233); +lean_ctor_set_uint8(x_235, sizeof(void*)*1, x_234); +x_208 = x_235; +goto block_223; } else { -lean_object* x_225; -lean_dec(x_194); -x_225 = l_Lean_IR_formatFnBodyHead___closed__23; -x_197 = x_225; -goto block_212; -} -block_212: -{ -lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; -x_198 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_198, 0, x_196); -lean_ctor_set(x_198, 1, x_197); -x_199 = l_Lean_IR_formatArray___redArg___lam__0___closed__1; -x_200 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_200, 0, x_198); -lean_ctor_set(x_200, 1, x_199); -x_201 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; -x_202 = l_Nat_reprFast(x_193); -x_203 = lean_string_append(x_201, x_202); -lean_dec_ref(x_202); -x_204 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_204, 0, x_203); -x_205 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_205, 0, x_200); -lean_ctor_set(x_205, 1, x_204); -x_206 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop___closed__1; -x_207 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_207, 0, x_205); -lean_ctor_set(x_207, 1, x_206); -x_208 = lean_box(1); +lean_object* x_236; +lean_dec(x_205); +x_236 = l_Lean_IR_formatFnBodyHead___closed__23; +x_208 = x_236; +goto block_223; +} +block_223: +{ +lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; x_209 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_209, 0, x_207); lean_ctor_set(x_209, 1, x_208); -x_210 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop(x_1, x_195); +x_210 = l_Lean_IR_formatArray___redArg___lam__0___closed__1; x_211 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_211, 0, x_209); lean_ctor_set(x_211, 1, x_210); -return x_211; +x_212 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_213 = l_Nat_reprFast(x_204); +x_214 = lean_string_append(x_212, x_213); +lean_dec_ref(x_213); +x_215 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_215, 0, x_214); +x_216 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_216, 0, x_211); +lean_ctor_set(x_216, 1, x_215); +x_217 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop___closed__1; +x_218 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_218, 0, x_216); +lean_ctor_set(x_218, 1, x_217); +x_219 = lean_box(1); +x_220 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_220, 0, x_218); +lean_ctor_set(x_220, 1, x_219); +x_221 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop(x_1, x_206); +x_222 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_222, 0, x_220); +lean_ctor_set(x_222, 1, x_221); +return x_222; } } case 8: { -uint8_t x_226; -x_226 = !lean_is_exclusive(x_2); -if (x_226 == 0) -{ -lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; -x_227 = lean_ctor_get(x_2, 0); -x_228 = lean_ctor_get(x_2, 1); -x_229 = l_Lean_IR_formatFnBodyHead___closed__27; -x_230 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; -x_231 = l_Nat_reprFast(x_227); -x_232 = lean_string_append(x_230, x_231); -lean_dec_ref(x_231); -x_233 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_233, 0, x_232); +uint8_t x_237; +x_237 = !lean_is_exclusive(x_2); +if (x_237 == 0) +{ +lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; +x_238 = lean_ctor_get(x_2, 0); +x_239 = lean_ctor_get(x_2, 1); +x_240 = l_Lean_IR_formatFnBodyHead___closed__27; +x_241 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_242 = l_Nat_reprFast(x_238); +x_243 = lean_string_append(x_241, x_242); +lean_dec_ref(x_242); +x_244 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_244, 0, x_243); lean_ctor_set_tag(x_2, 5); -lean_ctor_set(x_2, 1, x_233); -lean_ctor_set(x_2, 0, x_229); -x_234 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop___closed__1; -x_235 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_235, 0, x_2); -lean_ctor_set(x_235, 1, x_234); -x_236 = lean_box(1); -x_237 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_237, 0, x_235); -lean_ctor_set(x_237, 1, x_236); -x_238 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop(x_1, x_228); -x_239 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_239, 0, x_237); -lean_ctor_set(x_239, 1, x_238); -return x_239; +lean_ctor_set(x_2, 1, x_244); +lean_ctor_set(x_2, 0, x_240); +x_245 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop___closed__1; +x_246 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_246, 0, x_2); +lean_ctor_set(x_246, 1, x_245); +x_247 = lean_box(1); +x_248 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_248, 0, x_246); +lean_ctor_set(x_248, 1, x_247); +x_249 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop(x_1, x_239); +x_250 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_250, 0, x_248); +lean_ctor_set(x_250, 1, x_249); +return x_250; } else { -lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; -x_240 = lean_ctor_get(x_2, 0); -x_241 = lean_ctor_get(x_2, 1); -lean_inc(x_241); -lean_inc(x_240); +lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; +x_251 = lean_ctor_get(x_2, 0); +x_252 = lean_ctor_get(x_2, 1); +lean_inc(x_252); +lean_inc(x_251); lean_dec(x_2); -x_242 = l_Lean_IR_formatFnBodyHead___closed__27; -x_243 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; -x_244 = l_Nat_reprFast(x_240); -x_245 = lean_string_append(x_243, x_244); -lean_dec_ref(x_244); -x_246 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_246, 0, x_245); -x_247 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_247, 0, x_242); -lean_ctor_set(x_247, 1, x_246); -x_248 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop___closed__1; -x_249 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_249, 0, x_247); -lean_ctor_set(x_249, 1, x_248); -x_250 = lean_box(1); -x_251 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_251, 0, x_249); -lean_ctor_set(x_251, 1, x_250); -x_252 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop(x_1, x_241); -x_253 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_253, 0, x_251); -lean_ctor_set(x_253, 1, x_252); -return x_253; -} -} -case 9: -{ -lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; uint8_t x_272; -x_254 = lean_ctor_get(x_2, 1); -lean_inc(x_254); -x_255 = lean_ctor_get(x_2, 2); -lean_inc(x_255); -x_256 = lean_ctor_get(x_2, 3); -lean_inc_ref(x_256); -lean_dec_ref(x_2); -x_257 = l_Lean_IR_formatFnBodyHead___closed__29; -x_258 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; -x_259 = l_Nat_reprFast(x_254); -x_260 = lean_string_append(x_258, x_259); -lean_dec_ref(x_259); -x_261 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_261, 0, x_260); +x_253 = l_Lean_IR_formatFnBodyHead___closed__27; +x_254 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_255 = l_Nat_reprFast(x_251); +x_256 = lean_string_append(x_254, x_255); +lean_dec_ref(x_255); +x_257 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_257, 0, x_256); +x_258 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_258, 0, x_253); +lean_ctor_set(x_258, 1, x_257); +x_259 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop___closed__1; +x_260 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_260, 0, x_258); +lean_ctor_set(x_260, 1, x_259); +x_261 = lean_box(1); x_262 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_262, 0, x_257); +lean_ctor_set(x_262, 0, x_260); lean_ctor_set(x_262, 1, x_261); -x_263 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__3; +x_263 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop(x_1, x_252); x_264 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_264, 0, x_262); lean_ctor_set(x_264, 1, x_263); -x_265 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType(x_255); -x_266 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_266, 0, x_264); -lean_ctor_set(x_266, 1, x_265); -x_267 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop___closed__5; -x_268 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_268, 0, x_266); -lean_ctor_set(x_268, 1, x_267); -x_269 = lean_box(0); -x_270 = lean_unsigned_to_nat(0u); -x_271 = lean_array_get_size(x_256); -x_272 = lean_nat_dec_lt(x_270, x_271); -if (x_272 == 0) -{ -lean_object* x_273; -lean_dec_ref(x_256); -lean_dec(x_1); +return x_264; +} +} +case 9: +{ +lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; uint8_t x_283; +x_265 = lean_ctor_get(x_2, 1); +lean_inc(x_265); +x_266 = lean_ctor_get(x_2, 2); +lean_inc(x_266); +x_267 = lean_ctor_get(x_2, 3); +lean_inc_ref(x_267); +lean_dec_ref(x_2); +x_268 = l_Lean_IR_formatFnBodyHead___closed__29; +x_269 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_270 = l_Nat_reprFast(x_265); +x_271 = lean_string_append(x_269, x_270); +lean_dec_ref(x_270); +x_272 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_272, 0, x_271); x_273 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_273, 0, x_268); -lean_ctor_set(x_273, 1, x_269); -return x_273; +lean_ctor_set(x_273, 1, x_272); +x_274 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__1; +x_275 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_275, 0, x_273); +lean_ctor_set(x_275, 1, x_274); +x_276 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType(x_266); +x_277 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_277, 0, x_275); +lean_ctor_set(x_277, 1, x_276); +x_278 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop___closed__5; +x_279 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_279, 0, x_277); +lean_ctor_set(x_279, 1, x_278); +x_280 = lean_box(0); +x_281 = lean_unsigned_to_nat(0u); +x_282 = lean_array_get_size(x_267); +x_283 = lean_nat_dec_lt(x_281, x_282); +if (x_283 == 0) +{ +lean_object* x_284; +lean_dec_ref(x_267); +lean_dec(x_1); +x_284 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_284, 0, x_279); +lean_ctor_set(x_284, 1, x_280); +return x_284; } else { -uint8_t x_274; -x_274 = lean_nat_dec_le(x_271, x_271); -if (x_274 == 0) +uint8_t x_285; +x_285 = lean_nat_dec_le(x_282, x_282); +if (x_285 == 0) { -if (x_272 == 0) +if (x_283 == 0) { -lean_object* x_275; -lean_dec_ref(x_256); +lean_object* x_286; +lean_dec_ref(x_267); lean_dec(x_1); -x_275 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_275, 0, x_268); -lean_ctor_set(x_275, 1, x_269); -return x_275; +x_286 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_286, 0, x_279); +lean_ctor_set(x_286, 1, x_280); +return x_286; } else { -size_t x_276; size_t x_277; lean_object* x_278; lean_object* x_279; -x_276 = 0; -x_277 = lean_usize_of_nat(x_271); -x_278 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop_spec__0(x_1, x_256, x_276, x_277, x_269); -lean_dec_ref(x_256); -x_279 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_279, 0, x_268); -lean_ctor_set(x_279, 1, x_278); -return x_279; +size_t x_287; size_t x_288; lean_object* x_289; lean_object* x_290; +x_287 = 0; +x_288 = lean_usize_of_nat(x_282); +x_289 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop_spec__0(x_1, x_267, x_287, x_288, x_280); +lean_dec_ref(x_267); +x_290 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_290, 0, x_279); +lean_ctor_set(x_290, 1, x_289); +return x_290; } } else { -size_t x_280; size_t x_281; lean_object* x_282; lean_object* x_283; -x_280 = 0; -x_281 = lean_usize_of_nat(x_271); -x_282 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop_spec__0(x_1, x_256, x_280, x_281, x_269); -lean_dec_ref(x_256); -x_283 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_283, 0, x_268); -lean_ctor_set(x_283, 1, x_282); -return x_283; +size_t x_291; size_t x_292; lean_object* x_293; lean_object* x_294; +x_291 = 0; +x_292 = lean_usize_of_nat(x_282); +x_293 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop_spec__0(x_1, x_267, x_291, x_292, x_280); +lean_dec_ref(x_267); +x_294 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_294, 0, x_279); +lean_ctor_set(x_294, 1, x_293); +return x_294; } } } case 10: { -lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; +lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_dec(x_1); -x_284 = lean_ctor_get(x_2, 0); -lean_inc(x_284); +x_295 = lean_ctor_get(x_2, 0); +lean_inc(x_295); lean_dec_ref(x_2); -x_285 = l_Lean_IR_formatFnBodyHead___closed__33; -x_286 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg(x_284); -x_287 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_287, 0, x_285); -lean_ctor_set(x_287, 1, x_286); -return x_287; +x_296 = l_Lean_IR_formatFnBodyHead___closed__33; +x_297 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg(x_295); +x_298 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_298, 0, x_296); +lean_ctor_set(x_298, 1, x_297); +return x_298; } case 11: { -uint8_t x_288; +uint8_t x_299; lean_dec(x_1); -x_288 = !lean_is_exclusive(x_2); -if (x_288 == 0) -{ -lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; -x_289 = lean_ctor_get(x_2, 0); -x_290 = lean_ctor_get(x_2, 1); -x_291 = l_Lean_IR_formatFnBodyHead___closed__35; -x_292 = l_Lean_IR_formatFnBodyHead___closed__4; -x_293 = l_Nat_reprFast(x_289); -x_294 = lean_string_append(x_292, x_293); -lean_dec_ref(x_293); -x_295 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_295, 0, x_294); +x_299 = !lean_is_exclusive(x_2); +if (x_299 == 0) +{ +lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; +x_300 = lean_ctor_get(x_2, 0); +x_301 = lean_ctor_get(x_2, 1); +x_302 = l_Lean_IR_formatFnBodyHead___closed__35; +x_303 = l_Lean_IR_formatFnBodyHead___closed__4; +x_304 = l_Nat_reprFast(x_300); +x_305 = lean_string_append(x_303, x_304); +lean_dec_ref(x_304); +x_306 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_306, 0, x_305); lean_ctor_set_tag(x_2, 5); -lean_ctor_set(x_2, 1, x_295); -lean_ctor_set(x_2, 0, x_291); -x_296 = l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(x_290); -lean_dec_ref(x_290); -x_297 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_297, 0, x_2); -lean_ctor_set(x_297, 1, x_296); -return x_297; +lean_ctor_set(x_2, 1, x_306); +lean_ctor_set(x_2, 0, x_302); +x_307 = l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(x_301); +lean_dec_ref(x_301); +x_308 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_308, 0, x_2); +lean_ctor_set(x_308, 1, x_307); +return x_308; } else { -lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; -x_298 = lean_ctor_get(x_2, 0); -x_299 = lean_ctor_get(x_2, 1); -lean_inc(x_299); -lean_inc(x_298); +lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; +x_309 = lean_ctor_get(x_2, 0); +x_310 = lean_ctor_get(x_2, 1); +lean_inc(x_310); +lean_inc(x_309); lean_dec(x_2); -x_300 = l_Lean_IR_formatFnBodyHead___closed__35; -x_301 = l_Lean_IR_formatFnBodyHead___closed__4; -x_302 = l_Nat_reprFast(x_298); -x_303 = lean_string_append(x_301, x_302); -lean_dec_ref(x_302); -x_304 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_304, 0, x_303); -x_305 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_305, 0, x_300); -lean_ctor_set(x_305, 1, x_304); -x_306 = l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(x_299); -lean_dec_ref(x_299); -x_307 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_307, 0, x_305); -lean_ctor_set(x_307, 1, x_306); -return x_307; +x_311 = l_Lean_IR_formatFnBodyHead___closed__35; +x_312 = l_Lean_IR_formatFnBodyHead___closed__4; +x_313 = l_Nat_reprFast(x_309); +x_314 = lean_string_append(x_312, x_313); +lean_dec_ref(x_313); +x_315 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_315, 0, x_314); +x_316 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_316, 0, x_311); +lean_ctor_set(x_316, 1, x_315); +x_317 = l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(x_310); +lean_dec_ref(x_310); +x_318 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_318, 0, x_316); +lean_ctor_set(x_318, 1, x_317); +return x_318; } } default: { -lean_object* x_308; +lean_object* x_319; lean_dec(x_1); -x_308 = l_Lean_IR_formatFnBodyHead___closed__37; -return x_308; +x_319 = l_Lean_IR_formatFnBodyHead___closed__37; +return x_319; } } } @@ -4644,7 +4662,7 @@ lean_dec_ref(x_4); x_13 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_13, 0, x_11); lean_ctor_set(x_13, 1, x_12); -x_14 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__3; +x_14 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__1; x_15 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_15, 0, x_13); lean_ctor_set(x_15, 1, x_14); @@ -4695,7 +4713,7 @@ lean_dec_ref(x_27); x_35 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_35, 0, x_33); lean_ctor_set(x_35, 1, x_34); -x_36 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__3; +x_36 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__1; x_37 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_37, 0, x_35); lean_ctor_set(x_37, 1, x_36); @@ -4921,32 +4939,42 @@ l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__17 = _init lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__17); l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__18 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__18(); lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__18); -l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__19 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__19(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__19); l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__20 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__20(); lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__20); +l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__21 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__21(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__21); l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__22 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__22(); lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__22); +l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__19 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__19(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__19); l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__23 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__23(); lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__23); l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__24 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__24(); lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__24); -l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__21 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__21(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__21); -l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__25 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__25(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__25); l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__26 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__26(); lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__26); l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__27 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__27(); lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__27); -l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__28 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__28(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__28); +l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__25 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__25(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__25); l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__29 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__29(); lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__29); l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__30 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__30(); lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__30); l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__31 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__31(); lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__31); +l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__28 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__28(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__28); +l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__32 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__32(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__32); +l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__33 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__33(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__33); +l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__34 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__34(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__34); +l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__35 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__35(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__35); +l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__36 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__36(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__36); l_Lean_IR_instToFormatIRType___closed__0 = _init_l_Lean_IR_instToFormatIRType___closed__0(); lean_mark_persistent(l_Lean_IR_instToFormatIRType___closed__0); l_Lean_IR_instToFormatIRType = _init_l_Lean_IR_instToFormatIRType(); @@ -4959,14 +4987,6 @@ l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__1 = _init_l lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__1); l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__2 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__2(); lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__2); -l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__3 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__3(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__3); -l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__4 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__4(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__4); -l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__5 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__5(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__5); -l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__6 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__6(); -lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__6); l_Lean_IR_instToFormatParam___closed__0 = _init_l_Lean_IR_instToFormatParam___closed__0(); lean_mark_persistent(l_Lean_IR_instToFormatParam___closed__0); l_Lean_IR_instToFormatParam = _init_l_Lean_IR_instToFormatParam(); diff --git a/stage0/stdlib/Lean/Compiler/IR/FreeVars.c b/stage0/stdlib/Lean/Compiler/IR/FreeVars.c index 67ac8096df9b..06549152fcb8 100644 --- a/stage0/stdlib/Lean/Compiler/IR/FreeVars.c +++ b/stage0/stdlib/Lean/Compiler/IR/FreeVars.c @@ -431,7 +431,7 @@ return x_77; } } } -case 5: +case 3: { lean_object* x_78; lean_object* x_79; x_78 = lean_ctor_get(x_1, 2); @@ -440,179 +440,197 @@ lean_dec_ref(x_1); x_79 = l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_MaxIndex_visitIndex(x_78, x_2); return x_79; } +case 4: +{ +lean_object* x_80; lean_object* x_81; +x_80 = lean_ctor_get(x_1, 2); +lean_inc(x_80); +lean_dec_ref(x_1); +x_81 = l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_MaxIndex_visitIndex(x_80, x_2); +return x_81; +} +case 5: +{ +lean_object* x_82; lean_object* x_83; +x_82 = lean_ctor_get(x_1, 3); +lean_inc(x_82); +lean_dec_ref(x_1); +x_83 = l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_MaxIndex_visitIndex(x_82, x_2); +return x_83; +} case 6: { -lean_object* x_80; -x_80 = lean_ctor_get(x_1, 1); -lean_inc_ref(x_80); +lean_object* x_84; +x_84 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_84); lean_dec_ref(x_1); -x_3 = x_80; +x_3 = x_84; x_4 = x_2; goto block_18; } case 7: { -lean_object* x_81; -x_81 = lean_ctor_get(x_1, 1); -lean_inc_ref(x_81); +lean_object* x_85; +x_85 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_85); lean_dec_ref(x_1); -x_3 = x_81; +x_3 = x_85; x_4 = x_2; goto block_18; } case 8: { -lean_object* x_82; lean_object* x_83; lean_object* x_84; uint8_t x_85; -x_82 = lean_ctor_get(x_1, 0); -lean_inc(x_82); -x_83 = lean_ctor_get(x_1, 1); -lean_inc_ref(x_83); +lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; +x_86 = lean_ctor_get(x_1, 0); +lean_inc(x_86); +x_87 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_87); lean_dec_ref(x_1); -x_84 = l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_MaxIndex_visitIndex(x_82, x_2); -x_85 = !lean_is_exclusive(x_84); -if (x_85 == 0) -{ -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; uint8_t x_91; -x_86 = lean_ctor_get(x_84, 1); -x_87 = lean_ctor_get(x_84, 0); -lean_dec(x_87); -x_88 = lean_unsigned_to_nat(0u); -x_89 = lean_array_get_size(x_83); -x_90 = lean_box(0); -x_91 = lean_nat_dec_lt(x_88, x_89); -if (x_91 == 0) +x_88 = l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_MaxIndex_visitIndex(x_86, x_2); +x_89 = !lean_is_exclusive(x_88); +if (x_89 == 0) +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; uint8_t x_95; +x_90 = lean_ctor_get(x_88, 1); +x_91 = lean_ctor_get(x_88, 0); +lean_dec(x_91); +x_92 = lean_unsigned_to_nat(0u); +x_93 = lean_array_get_size(x_87); +x_94 = lean_box(0); +x_95 = lean_nat_dec_lt(x_92, x_93); +if (x_95 == 0) { -lean_dec_ref(x_83); -lean_ctor_set(x_84, 0, x_90); -return x_84; +lean_dec_ref(x_87); +lean_ctor_set(x_88, 0, x_94); +return x_88; } else { -uint8_t x_92; -x_92 = lean_nat_dec_le(x_89, x_89); -if (x_92 == 0) +uint8_t x_96; +x_96 = lean_nat_dec_le(x_93, x_93); +if (x_96 == 0) { -if (x_91 == 0) +if (x_95 == 0) { -lean_dec_ref(x_83); -lean_ctor_set(x_84, 0, x_90); -return x_84; +lean_dec_ref(x_87); +lean_ctor_set(x_88, 0, x_94); +return x_88; } else { -size_t x_93; size_t x_94; lean_object* x_95; -lean_free_object(x_84); -x_93 = 0; -x_94 = lean_usize_of_nat(x_89); -x_95 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_FreeVars_0__Lean_IR_MaxIndex_visitExpr_spec__0(x_83, x_93, x_94, x_90, x_86); -lean_dec_ref(x_83); -return x_95; +size_t x_97; size_t x_98; lean_object* x_99; +lean_free_object(x_88); +x_97 = 0; +x_98 = lean_usize_of_nat(x_93); +x_99 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_FreeVars_0__Lean_IR_MaxIndex_visitExpr_spec__0(x_87, x_97, x_98, x_94, x_90); +lean_dec_ref(x_87); +return x_99; } } else { -size_t x_96; size_t x_97; lean_object* x_98; -lean_free_object(x_84); -x_96 = 0; -x_97 = lean_usize_of_nat(x_89); -x_98 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_FreeVars_0__Lean_IR_MaxIndex_visitExpr_spec__0(x_83, x_96, x_97, x_90, x_86); -lean_dec_ref(x_83); -return x_98; +size_t x_100; size_t x_101; lean_object* x_102; +lean_free_object(x_88); +x_100 = 0; +x_101 = lean_usize_of_nat(x_93); +x_102 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_FreeVars_0__Lean_IR_MaxIndex_visitExpr_spec__0(x_87, x_100, x_101, x_94, x_90); +lean_dec_ref(x_87); +return x_102; } } } else { -lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; uint8_t x_103; -x_99 = lean_ctor_get(x_84, 1); -lean_inc(x_99); -lean_dec(x_84); -x_100 = lean_unsigned_to_nat(0u); -x_101 = lean_array_get_size(x_83); -x_102 = lean_box(0); -x_103 = lean_nat_dec_lt(x_100, x_101); -if (x_103 == 0) +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; uint8_t x_107; +x_103 = lean_ctor_get(x_88, 1); +lean_inc(x_103); +lean_dec(x_88); +x_104 = lean_unsigned_to_nat(0u); +x_105 = lean_array_get_size(x_87); +x_106 = lean_box(0); +x_107 = lean_nat_dec_lt(x_104, x_105); +if (x_107 == 0) { -lean_object* x_104; -lean_dec_ref(x_83); -x_104 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_104, 0, x_102); -lean_ctor_set(x_104, 1, x_99); -return x_104; +lean_object* x_108; +lean_dec_ref(x_87); +x_108 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_108, 0, x_106); +lean_ctor_set(x_108, 1, x_103); +return x_108; } else { -uint8_t x_105; -x_105 = lean_nat_dec_le(x_101, x_101); -if (x_105 == 0) +uint8_t x_109; +x_109 = lean_nat_dec_le(x_105, x_105); +if (x_109 == 0) { -if (x_103 == 0) +if (x_107 == 0) { -lean_object* x_106; -lean_dec_ref(x_83); -x_106 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_106, 0, x_102); -lean_ctor_set(x_106, 1, x_99); -return x_106; +lean_object* x_110; +lean_dec_ref(x_87); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_106); +lean_ctor_set(x_110, 1, x_103); +return x_110; } else { -size_t x_107; size_t x_108; lean_object* x_109; -x_107 = 0; -x_108 = lean_usize_of_nat(x_101); -x_109 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_FreeVars_0__Lean_IR_MaxIndex_visitExpr_spec__0(x_83, x_107, x_108, x_102, x_99); -lean_dec_ref(x_83); -return x_109; +size_t x_111; size_t x_112; lean_object* x_113; +x_111 = 0; +x_112 = lean_usize_of_nat(x_105); +x_113 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_FreeVars_0__Lean_IR_MaxIndex_visitExpr_spec__0(x_87, x_111, x_112, x_106, x_103); +lean_dec_ref(x_87); +return x_113; } } else { -size_t x_110; size_t x_111; lean_object* x_112; -x_110 = 0; -x_111 = lean_usize_of_nat(x_101); -x_112 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_FreeVars_0__Lean_IR_MaxIndex_visitExpr_spec__0(x_83, x_110, x_111, x_102, x_99); -lean_dec_ref(x_83); -return x_112; +size_t x_114; size_t x_115; lean_object* x_116; +x_114 = 0; +x_115 = lean_usize_of_nat(x_105); +x_116 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_FreeVars_0__Lean_IR_MaxIndex_visitExpr_spec__0(x_87, x_114, x_115, x_106, x_103); +lean_dec_ref(x_87); +return x_116; } } } } case 10: { -lean_object* x_113; lean_object* x_114; -x_113 = lean_ctor_get(x_1, 0); -lean_inc(x_113); +lean_object* x_117; lean_object* x_118; +x_117 = lean_ctor_get(x_1, 0); +lean_inc(x_117); lean_dec_ref(x_1); -x_114 = l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_MaxIndex_visitIndex(x_113, x_2); -return x_114; +x_118 = l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_MaxIndex_visitIndex(x_117, x_2); +return x_118; } case 11: { -lean_object* x_115; lean_object* x_116; +lean_object* x_119; lean_object* x_120; lean_dec_ref(x_1); -x_115 = lean_box(0); -x_116 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_116, 0, x_115); -lean_ctor_set(x_116, 1, x_2); -return x_116; +x_119 = lean_box(0); +x_120 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_120, 0, x_119); +lean_ctor_set(x_120, 1, x_2); +return x_120; } case 12: { -lean_object* x_117; lean_object* x_118; -x_117 = lean_ctor_get(x_1, 0); -lean_inc(x_117); +lean_object* x_121; lean_object* x_122; +x_121 = lean_ctor_get(x_1, 0); +lean_inc(x_121); lean_dec_ref(x_1); -x_118 = l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_MaxIndex_visitIndex(x_117, x_2); -return x_118; +x_122 = l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_MaxIndex_visitIndex(x_121, x_2); +return x_122; } default: { -lean_object* x_119; lean_object* x_120; -x_119 = lean_ctor_get(x_1, 1); -lean_inc(x_119); +lean_object* x_123; lean_object* x_124; +x_123 = lean_ctor_get(x_1, 1); +lean_inc(x_123); lean_dec_ref(x_1); -x_120 = l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_MaxIndex_visitIndex(x_119, x_2); -return x_120; +x_124 = l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_MaxIndex_visitIndex(x_123, x_2); +return x_124; } } block_18: @@ -861,9 +879,9 @@ case 4: lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; x_56 = lean_ctor_get(x_1, 0); lean_inc(x_56); -x_57 = lean_ctor_get(x_1, 2); +x_57 = lean_ctor_get(x_1, 3); lean_inc(x_57); -x_58 = lean_ctor_get(x_1, 3); +x_58 = lean_ctor_get(x_1, 4); lean_inc(x_58); lean_dec_ref(x_1); x_59 = l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_MaxIndex_visitIndex(x_56, x_2); @@ -883,9 +901,9 @@ case 5: lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; x_64 = lean_ctor_get(x_1, 0); lean_inc(x_64); -x_65 = lean_ctor_get(x_1, 3); +x_65 = lean_ctor_get(x_1, 4); lean_inc(x_65); -x_66 = lean_ctor_get(x_1, 5); +x_66 = lean_ctor_get(x_1, 6); lean_inc(x_66); lean_dec_ref(x_1); x_67 = l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_MaxIndex_visitIndex(x_64, x_2); @@ -1591,9 +1609,9 @@ goto block_52; block_45: { lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_nat_add(x_40, x_41); +x_42 = lean_nat_add(x_39, x_41); lean_dec(x_41); -lean_dec(x_40); +lean_dec(x_39); if (lean_is_scalar(x_36)) { x_43 = lean_alloc_ctor(0, 5, 0); } else { @@ -1612,7 +1630,7 @@ if (lean_is_scalar(x_26)) { lean_ctor_set(x_44, 0, x_38); lean_ctor_set(x_44, 1, x_28); lean_ctor_set(x_44, 2, x_29); -lean_ctor_set(x_44, 3, x_39); +lean_ctor_set(x_44, 3, x_40); lean_ctor_set(x_44, 4, x_43); return x_44; } @@ -1638,8 +1656,8 @@ if (lean_obj_tag(x_31) == 0) lean_object* x_50; x_50 = lean_ctor_get(x_31, 0); lean_inc(x_50); -x_39 = x_48; -x_40 = x_49; +x_39 = x_49; +x_40 = x_48; x_41 = x_50; goto block_45; } @@ -1647,8 +1665,8 @@ else { lean_object* x_51; x_51 = lean_unsigned_to_nat(0u); -x_39 = x_48; -x_40 = x_49; +x_39 = x_49; +x_40 = x_48; x_41 = x_51; goto block_45; } @@ -2799,7 +2817,7 @@ return x_77; } } } -case 5: +case 3: { lean_object* x_78; lean_object* x_79; x_78 = lean_ctor_get(x_1, 2); @@ -2808,179 +2826,197 @@ lean_dec_ref(x_1); x_79 = l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_FreeIndices_visitIndex(x_78, x_2); return x_79; } +case 4: +{ +lean_object* x_80; lean_object* x_81; +x_80 = lean_ctor_get(x_1, 2); +lean_inc(x_80); +lean_dec_ref(x_1); +x_81 = l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_FreeIndices_visitIndex(x_80, x_2); +return x_81; +} +case 5: +{ +lean_object* x_82; lean_object* x_83; +x_82 = lean_ctor_get(x_1, 3); +lean_inc(x_82); +lean_dec_ref(x_1); +x_83 = l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_FreeIndices_visitIndex(x_82, x_2); +return x_83; +} case 6: { -lean_object* x_80; -x_80 = lean_ctor_get(x_1, 1); -lean_inc_ref(x_80); +lean_object* x_84; +x_84 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_84); lean_dec_ref(x_1); -x_3 = x_80; +x_3 = x_84; x_4 = x_2; goto block_18; } case 7: { -lean_object* x_81; -x_81 = lean_ctor_get(x_1, 1); -lean_inc_ref(x_81); +lean_object* x_85; +x_85 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_85); lean_dec_ref(x_1); -x_3 = x_81; +x_3 = x_85; x_4 = x_2; goto block_18; } case 8: { -lean_object* x_82; lean_object* x_83; lean_object* x_84; uint8_t x_85; -x_82 = lean_ctor_get(x_1, 0); -lean_inc(x_82); -x_83 = lean_ctor_get(x_1, 1); -lean_inc_ref(x_83); +lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; +x_86 = lean_ctor_get(x_1, 0); +lean_inc(x_86); +x_87 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_87); lean_dec_ref(x_1); -x_84 = l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_FreeIndices_visitIndex(x_82, x_2); -x_85 = !lean_is_exclusive(x_84); -if (x_85 == 0) -{ -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; uint8_t x_91; -x_86 = lean_ctor_get(x_84, 1); -x_87 = lean_ctor_get(x_84, 0); -lean_dec(x_87); -x_88 = lean_unsigned_to_nat(0u); -x_89 = lean_array_get_size(x_83); -x_90 = lean_box(0); -x_91 = lean_nat_dec_lt(x_88, x_89); -if (x_91 == 0) +x_88 = l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_FreeIndices_visitIndex(x_86, x_2); +x_89 = !lean_is_exclusive(x_88); +if (x_89 == 0) +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; uint8_t x_95; +x_90 = lean_ctor_get(x_88, 1); +x_91 = lean_ctor_get(x_88, 0); +lean_dec(x_91); +x_92 = lean_unsigned_to_nat(0u); +x_93 = lean_array_get_size(x_87); +x_94 = lean_box(0); +x_95 = lean_nat_dec_lt(x_92, x_93); +if (x_95 == 0) { -lean_dec_ref(x_83); -lean_ctor_set(x_84, 0, x_90); -return x_84; +lean_dec_ref(x_87); +lean_ctor_set(x_88, 0, x_94); +return x_88; } else { -uint8_t x_92; -x_92 = lean_nat_dec_le(x_89, x_89); -if (x_92 == 0) +uint8_t x_96; +x_96 = lean_nat_dec_le(x_93, x_93); +if (x_96 == 0) { -if (x_91 == 0) +if (x_95 == 0) { -lean_dec_ref(x_83); -lean_ctor_set(x_84, 0, x_90); -return x_84; +lean_dec_ref(x_87); +lean_ctor_set(x_88, 0, x_94); +return x_88; } else { -size_t x_93; size_t x_94; lean_object* x_95; -lean_free_object(x_84); -x_93 = 0; -x_94 = lean_usize_of_nat(x_89); -x_95 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_FreeVars_0__Lean_IR_FreeIndices_visitExpr_spec__0(x_83, x_93, x_94, x_90, x_86); -lean_dec_ref(x_83); -return x_95; +size_t x_97; size_t x_98; lean_object* x_99; +lean_free_object(x_88); +x_97 = 0; +x_98 = lean_usize_of_nat(x_93); +x_99 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_FreeVars_0__Lean_IR_FreeIndices_visitExpr_spec__0(x_87, x_97, x_98, x_94, x_90); +lean_dec_ref(x_87); +return x_99; } } else { -size_t x_96; size_t x_97; lean_object* x_98; -lean_free_object(x_84); -x_96 = 0; -x_97 = lean_usize_of_nat(x_89); -x_98 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_FreeVars_0__Lean_IR_FreeIndices_visitExpr_spec__0(x_83, x_96, x_97, x_90, x_86); -lean_dec_ref(x_83); -return x_98; +size_t x_100; size_t x_101; lean_object* x_102; +lean_free_object(x_88); +x_100 = 0; +x_101 = lean_usize_of_nat(x_93); +x_102 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_FreeVars_0__Lean_IR_FreeIndices_visitExpr_spec__0(x_87, x_100, x_101, x_94, x_90); +lean_dec_ref(x_87); +return x_102; } } } else { -lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; uint8_t x_103; -x_99 = lean_ctor_get(x_84, 1); -lean_inc(x_99); -lean_dec(x_84); -x_100 = lean_unsigned_to_nat(0u); -x_101 = lean_array_get_size(x_83); -x_102 = lean_box(0); -x_103 = lean_nat_dec_lt(x_100, x_101); -if (x_103 == 0) +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; uint8_t x_107; +x_103 = lean_ctor_get(x_88, 1); +lean_inc(x_103); +lean_dec(x_88); +x_104 = lean_unsigned_to_nat(0u); +x_105 = lean_array_get_size(x_87); +x_106 = lean_box(0); +x_107 = lean_nat_dec_lt(x_104, x_105); +if (x_107 == 0) { -lean_object* x_104; -lean_dec_ref(x_83); -x_104 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_104, 0, x_102); -lean_ctor_set(x_104, 1, x_99); -return x_104; +lean_object* x_108; +lean_dec_ref(x_87); +x_108 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_108, 0, x_106); +lean_ctor_set(x_108, 1, x_103); +return x_108; } else { -uint8_t x_105; -x_105 = lean_nat_dec_le(x_101, x_101); -if (x_105 == 0) +uint8_t x_109; +x_109 = lean_nat_dec_le(x_105, x_105); +if (x_109 == 0) { -if (x_103 == 0) +if (x_107 == 0) { -lean_object* x_106; -lean_dec_ref(x_83); -x_106 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_106, 0, x_102); -lean_ctor_set(x_106, 1, x_99); -return x_106; +lean_object* x_110; +lean_dec_ref(x_87); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_106); +lean_ctor_set(x_110, 1, x_103); +return x_110; } else { -size_t x_107; size_t x_108; lean_object* x_109; -x_107 = 0; -x_108 = lean_usize_of_nat(x_101); -x_109 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_FreeVars_0__Lean_IR_FreeIndices_visitExpr_spec__0(x_83, x_107, x_108, x_102, x_99); -lean_dec_ref(x_83); -return x_109; +size_t x_111; size_t x_112; lean_object* x_113; +x_111 = 0; +x_112 = lean_usize_of_nat(x_105); +x_113 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_FreeVars_0__Lean_IR_FreeIndices_visitExpr_spec__0(x_87, x_111, x_112, x_106, x_103); +lean_dec_ref(x_87); +return x_113; } } else { -size_t x_110; size_t x_111; lean_object* x_112; -x_110 = 0; -x_111 = lean_usize_of_nat(x_101); -x_112 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_FreeVars_0__Lean_IR_FreeIndices_visitExpr_spec__0(x_83, x_110, x_111, x_102, x_99); -lean_dec_ref(x_83); -return x_112; +size_t x_114; size_t x_115; lean_object* x_116; +x_114 = 0; +x_115 = lean_usize_of_nat(x_105); +x_116 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_FreeVars_0__Lean_IR_FreeIndices_visitExpr_spec__0(x_87, x_114, x_115, x_106, x_103); +lean_dec_ref(x_87); +return x_116; } } } } case 10: { -lean_object* x_113; lean_object* x_114; -x_113 = lean_ctor_get(x_1, 0); -lean_inc(x_113); +lean_object* x_117; lean_object* x_118; +x_117 = lean_ctor_get(x_1, 0); +lean_inc(x_117); lean_dec_ref(x_1); -x_114 = l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_FreeIndices_visitIndex(x_113, x_2); -return x_114; +x_118 = l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_FreeIndices_visitIndex(x_117, x_2); +return x_118; } case 11: { -lean_object* x_115; lean_object* x_116; +lean_object* x_119; lean_object* x_120; lean_dec_ref(x_1); -x_115 = lean_box(0); -x_116 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_116, 0, x_115); -lean_ctor_set(x_116, 1, x_2); -return x_116; +x_119 = lean_box(0); +x_120 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_120, 0, x_119); +lean_ctor_set(x_120, 1, x_2); +return x_120; } case 12: { -lean_object* x_117; lean_object* x_118; -x_117 = lean_ctor_get(x_1, 0); -lean_inc(x_117); +lean_object* x_121; lean_object* x_122; +x_121 = lean_ctor_get(x_1, 0); +lean_inc(x_121); lean_dec_ref(x_1); -x_118 = l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_FreeIndices_visitIndex(x_117, x_2); -return x_118; +x_122 = l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_FreeIndices_visitIndex(x_121, x_2); +return x_122; } default: { -lean_object* x_119; lean_object* x_120; -x_119 = lean_ctor_get(x_1, 1); -lean_inc(x_119); +lean_object* x_123; lean_object* x_124; +x_123 = lean_ctor_get(x_1, 1); +lean_inc(x_123); lean_dec_ref(x_1); -x_120 = l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_FreeIndices_visitIndex(x_119, x_2); -return x_120; +x_124 = l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_FreeIndices_visitIndex(x_123, x_2); +return x_124; } } block_18: @@ -3229,9 +3265,9 @@ case 4: lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; x_56 = lean_ctor_get(x_1, 0); lean_inc(x_56); -x_57 = lean_ctor_get(x_1, 2); +x_57 = lean_ctor_get(x_1, 3); lean_inc(x_57); -x_58 = lean_ctor_get(x_1, 3); +x_58 = lean_ctor_get(x_1, 4); lean_inc(x_58); lean_dec_ref(x_1); x_59 = l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_FreeIndices_visitIndex(x_56, x_2); @@ -3251,9 +3287,9 @@ case 5: lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; x_64 = lean_ctor_get(x_1, 0); lean_inc(x_64); -x_65 = lean_ctor_get(x_1, 3); +x_65 = lean_ctor_get(x_1, 4); lean_inc(x_65); -x_66 = lean_ctor_get(x_1, 5); +x_66 = lean_ctor_get(x_1, 6); lean_inc(x_66); lean_dec_ref(x_1); x_67 = l___private_Lean_Compiler_IR_FreeVars_0__Lean_IR_FreeIndices_visitIndex(x_64, x_2); @@ -4013,11 +4049,11 @@ LEAN_EXPORT uint8_t l_Lean_IR_HasIndex_visitExpr(lean_object* x_1, lean_object* _start: { switch (lean_obj_tag(x_2)) { -case 0: +case 1: { lean_object* x_3; uint8_t x_4; x_3 = lean_ctor_get(x_2, 1); -x_4 = l_Lean_IR_HasIndex_visitArgs(x_1, x_3); +x_4 = lean_nat_dec_eq(x_1, x_3); return x_4; } case 2: @@ -4037,25 +4073,25 @@ else return x_7; } } -case 5: +case 3: { lean_object* x_9; uint8_t x_10; x_9 = lean_ctor_get(x_2, 2); x_10 = lean_nat_dec_eq(x_1, x_9); return x_10; } -case 6: +case 4: { lean_object* x_11; uint8_t x_12; -x_11 = lean_ctor_get(x_2, 1); -x_12 = l_Lean_IR_HasIndex_visitArgs(x_1, x_11); +x_11 = lean_ctor_get(x_2, 2); +x_12 = lean_nat_dec_eq(x_1, x_11); return x_12; } -case 7: +case 5: { lean_object* x_13; uint8_t x_14; -x_13 = lean_ctor_get(x_2, 1); -x_14 = l_Lean_IR_HasIndex_visitArgs(x_1, x_13); +x_13 = lean_ctor_get(x_2, 3); +x_14 = lean_nat_dec_eq(x_1, x_13); return x_14; } case 8: @@ -4075,33 +4111,40 @@ else return x_17; } } -case 10: +case 9: { lean_object* x_19; uint8_t x_20; -x_19 = lean_ctor_get(x_2, 0); +x_19 = lean_ctor_get(x_2, 1); x_20 = lean_nat_dec_eq(x_1, x_19); return x_20; } -case 11: +case 10: { -uint8_t x_21; -x_21 = 0; -return x_21; +lean_object* x_21; uint8_t x_22; +x_21 = lean_ctor_get(x_2, 0); +x_22 = lean_nat_dec_eq(x_1, x_21); +return x_22; } -case 12: +case 11: { -lean_object* x_22; uint8_t x_23; -x_22 = lean_ctor_get(x_2, 0); -x_23 = lean_nat_dec_eq(x_1, x_22); +uint8_t x_23; +x_23 = 0; return x_23; } -default: +case 12: { lean_object* x_24; uint8_t x_25; -x_24 = lean_ctor_get(x_2, 1); +x_24 = lean_ctor_get(x_2, 0); x_25 = lean_nat_dec_eq(x_1, x_24); return x_25; } +default: +{ +lean_object* x_26; uint8_t x_27; +x_26 = lean_ctor_get(x_2, 1); +x_27 = l_Lean_IR_HasIndex_visitArgs(x_1, x_26); +return x_27; +} } } } @@ -4205,8 +4248,8 @@ case 4: { lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; uint8_t x_34; x_28 = lean_ctor_get(x_2, 0); -x_29 = lean_ctor_get(x_2, 2); -x_30 = lean_ctor_get(x_2, 3); +x_29 = lean_ctor_get(x_2, 3); +x_30 = lean_ctor_get(x_2, 4); x_34 = lean_nat_dec_eq(x_1, x_28); if (x_34 == 0) { @@ -4237,8 +4280,8 @@ case 5: { lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; uint8_t x_42; x_36 = lean_ctor_get(x_2, 0); -x_37 = lean_ctor_get(x_2, 3); -x_38 = lean_ctor_get(x_2, 5); +x_37 = lean_ctor_get(x_2, 4); +x_38 = lean_ctor_get(x_2, 6); x_42 = lean_nat_dec_eq(x_1, x_36); if (x_42 == 0) { diff --git a/stage0/stdlib/Lean/Compiler/IR/LiveVars.c b/stage0/stdlib/Lean/Compiler/IR/LiveVars.c index 2b860f39ef19..1977be6f75da 100644 --- a/stage0/stdlib/Lean/Compiler/IR/LiveVars.c +++ b/stage0/stdlib/Lean/Compiler/IR/LiveVars.c @@ -340,9 +340,9 @@ case 4: lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; x_41 = lean_ctor_get(x_2, 0); lean_inc(x_41); -x_42 = lean_ctor_get(x_2, 2); +x_42 = lean_ctor_get(x_2, 3); lean_inc(x_42); -x_43 = lean_ctor_get(x_2, 3); +x_43 = lean_ctor_get(x_2, 4); lean_inc(x_43); lean_dec_ref(x_2); x_44 = l_Lean_IR_HasIndex_visitVar(x_1, x_41); @@ -385,9 +385,9 @@ case 5: lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; x_51 = lean_ctor_get(x_2, 0); lean_inc(x_51); -x_52 = lean_ctor_get(x_2, 3); +x_52 = lean_ctor_get(x_2, 4); lean_inc(x_52); -x_53 = lean_ctor_get(x_2, 5); +x_53 = lean_ctor_get(x_2, 6); lean_inc(x_53); lean_dec_ref(x_2); x_54 = l_Lean_IR_HasIndex_visitVar(x_1, x_51); @@ -2046,9 +2046,9 @@ goto block_51; block_43: { lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_nat_add(x_38, x_39); +x_40 = lean_nat_add(x_37, x_39); lean_dec(x_39); -lean_dec(x_38); +lean_dec(x_37); if (lean_is_scalar(x_34)) { x_41 = lean_alloc_ctor(0, 5, 0); } else { @@ -2067,7 +2067,7 @@ if (lean_is_scalar(x_24)) { lean_ctor_set(x_42, 0, x_36); lean_ctor_set(x_42, 1, x_27); lean_ctor_set(x_42, 2, x_28); -lean_ctor_set(x_42, 3, x_37); +lean_ctor_set(x_42, 3, x_38); lean_ctor_set(x_42, 4, x_41); return x_42; } @@ -2094,8 +2094,8 @@ if (lean_obj_tag(x_30) == 0) lean_object* x_49; x_49 = lean_ctor_get(x_30, 0); lean_inc(x_49); -x_37 = x_47; -x_38 = x_48; +x_37 = x_48; +x_38 = x_47; x_39 = x_49; goto block_43; } @@ -2103,8 +2103,8 @@ else { lean_object* x_50; x_50 = lean_unsigned_to_nat(0u); -x_37 = x_47; -x_38 = x_48; +x_37 = x_48; +x_38 = x_47; x_39 = x_50; goto block_43; } @@ -5159,7 +5159,7 @@ return x_77; } } } -case 5: +case 3: { lean_object* x_78; lean_object* x_79; x_78 = lean_ctor_get(x_1, 2); @@ -5168,179 +5168,197 @@ lean_dec_ref(x_1); x_79 = l___private_Lean_Compiler_IR_LiveVars_0__Lean_IR_LiveVars_useVar(x_78, x_2); return x_79; } +case 4: +{ +lean_object* x_80; lean_object* x_81; +x_80 = lean_ctor_get(x_1, 2); +lean_inc(x_80); +lean_dec_ref(x_1); +x_81 = l___private_Lean_Compiler_IR_LiveVars_0__Lean_IR_LiveVars_useVar(x_80, x_2); +return x_81; +} +case 5: +{ +lean_object* x_82; lean_object* x_83; +x_82 = lean_ctor_get(x_1, 3); +lean_inc(x_82); +lean_dec_ref(x_1); +x_83 = l___private_Lean_Compiler_IR_LiveVars_0__Lean_IR_LiveVars_useVar(x_82, x_2); +return x_83; +} case 6: { -lean_object* x_80; -x_80 = lean_ctor_get(x_1, 1); -lean_inc_ref(x_80); +lean_object* x_84; +x_84 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_84); lean_dec_ref(x_1); -x_3 = x_80; +x_3 = x_84; x_4 = x_2; goto block_18; } case 7: { -lean_object* x_81; -x_81 = lean_ctor_get(x_1, 1); -lean_inc_ref(x_81); +lean_object* x_85; +x_85 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_85); lean_dec_ref(x_1); -x_3 = x_81; +x_3 = x_85; x_4 = x_2; goto block_18; } case 8: { -lean_object* x_82; lean_object* x_83; lean_object* x_84; uint8_t x_85; -x_82 = lean_ctor_get(x_1, 0); -lean_inc(x_82); -x_83 = lean_ctor_get(x_1, 1); -lean_inc_ref(x_83); +lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; +x_86 = lean_ctor_get(x_1, 0); +lean_inc(x_86); +x_87 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_87); lean_dec_ref(x_1); -x_84 = l___private_Lean_Compiler_IR_LiveVars_0__Lean_IR_LiveVars_useVar(x_82, x_2); -x_85 = !lean_is_exclusive(x_84); -if (x_85 == 0) +x_88 = l___private_Lean_Compiler_IR_LiveVars_0__Lean_IR_LiveVars_useVar(x_86, x_2); +x_89 = !lean_is_exclusive(x_88); +if (x_89 == 0) { -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; uint8_t x_91; -x_86 = lean_ctor_get(x_84, 1); -x_87 = lean_ctor_get(x_84, 0); -lean_dec(x_87); -x_88 = lean_unsigned_to_nat(0u); -x_89 = lean_array_get_size(x_83); -x_90 = lean_box(0); -x_91 = lean_nat_dec_lt(x_88, x_89); -if (x_91 == 0) +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; uint8_t x_95; +x_90 = lean_ctor_get(x_88, 1); +x_91 = lean_ctor_get(x_88, 0); +lean_dec(x_91); +x_92 = lean_unsigned_to_nat(0u); +x_93 = lean_array_get_size(x_87); +x_94 = lean_box(0); +x_95 = lean_nat_dec_lt(x_92, x_93); +if (x_95 == 0) { -lean_dec_ref(x_83); -lean_ctor_set(x_84, 0, x_90); -return x_84; +lean_dec_ref(x_87); +lean_ctor_set(x_88, 0, x_94); +return x_88; } else { -uint8_t x_92; -x_92 = lean_nat_dec_le(x_89, x_89); -if (x_92 == 0) +uint8_t x_96; +x_96 = lean_nat_dec_le(x_93, x_93); +if (x_96 == 0) { -if (x_91 == 0) +if (x_95 == 0) { -lean_dec_ref(x_83); -lean_ctor_set(x_84, 0, x_90); -return x_84; +lean_dec_ref(x_87); +lean_ctor_set(x_88, 0, x_94); +return x_88; } else { -size_t x_93; size_t x_94; lean_object* x_95; -lean_free_object(x_84); -x_93 = 0; -x_94 = lean_usize_of_nat(x_89); -x_95 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_LiveVars_0__Lean_IR_LiveVars_useExpr_spec__0(x_83, x_93, x_94, x_90, x_86); -lean_dec_ref(x_83); -return x_95; +size_t x_97; size_t x_98; lean_object* x_99; +lean_free_object(x_88); +x_97 = 0; +x_98 = lean_usize_of_nat(x_93); +x_99 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_LiveVars_0__Lean_IR_LiveVars_useExpr_spec__0(x_87, x_97, x_98, x_94, x_90); +lean_dec_ref(x_87); +return x_99; } } else { -size_t x_96; size_t x_97; lean_object* x_98; -lean_free_object(x_84); -x_96 = 0; -x_97 = lean_usize_of_nat(x_89); -x_98 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_LiveVars_0__Lean_IR_LiveVars_useExpr_spec__0(x_83, x_96, x_97, x_90, x_86); -lean_dec_ref(x_83); -return x_98; +size_t x_100; size_t x_101; lean_object* x_102; +lean_free_object(x_88); +x_100 = 0; +x_101 = lean_usize_of_nat(x_93); +x_102 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_LiveVars_0__Lean_IR_LiveVars_useExpr_spec__0(x_87, x_100, x_101, x_94, x_90); +lean_dec_ref(x_87); +return x_102; } } } else { -lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; uint8_t x_103; -x_99 = lean_ctor_get(x_84, 1); -lean_inc(x_99); -lean_dec(x_84); -x_100 = lean_unsigned_to_nat(0u); -x_101 = lean_array_get_size(x_83); -x_102 = lean_box(0); -x_103 = lean_nat_dec_lt(x_100, x_101); -if (x_103 == 0) +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; uint8_t x_107; +x_103 = lean_ctor_get(x_88, 1); +lean_inc(x_103); +lean_dec(x_88); +x_104 = lean_unsigned_to_nat(0u); +x_105 = lean_array_get_size(x_87); +x_106 = lean_box(0); +x_107 = lean_nat_dec_lt(x_104, x_105); +if (x_107 == 0) { -lean_object* x_104; -lean_dec_ref(x_83); -x_104 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_104, 0, x_102); -lean_ctor_set(x_104, 1, x_99); -return x_104; +lean_object* x_108; +lean_dec_ref(x_87); +x_108 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_108, 0, x_106); +lean_ctor_set(x_108, 1, x_103); +return x_108; } else { -uint8_t x_105; -x_105 = lean_nat_dec_le(x_101, x_101); -if (x_105 == 0) +uint8_t x_109; +x_109 = lean_nat_dec_le(x_105, x_105); +if (x_109 == 0) { -if (x_103 == 0) +if (x_107 == 0) { -lean_object* x_106; -lean_dec_ref(x_83); -x_106 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_106, 0, x_102); -lean_ctor_set(x_106, 1, x_99); -return x_106; +lean_object* x_110; +lean_dec_ref(x_87); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_106); +lean_ctor_set(x_110, 1, x_103); +return x_110; } else { -size_t x_107; size_t x_108; lean_object* x_109; -x_107 = 0; -x_108 = lean_usize_of_nat(x_101); -x_109 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_LiveVars_0__Lean_IR_LiveVars_useExpr_spec__0(x_83, x_107, x_108, x_102, x_99); -lean_dec_ref(x_83); -return x_109; +size_t x_111; size_t x_112; lean_object* x_113; +x_111 = 0; +x_112 = lean_usize_of_nat(x_105); +x_113 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_LiveVars_0__Lean_IR_LiveVars_useExpr_spec__0(x_87, x_111, x_112, x_106, x_103); +lean_dec_ref(x_87); +return x_113; } } else { -size_t x_110; size_t x_111; lean_object* x_112; -x_110 = 0; -x_111 = lean_usize_of_nat(x_101); -x_112 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_LiveVars_0__Lean_IR_LiveVars_useExpr_spec__0(x_83, x_110, x_111, x_102, x_99); -lean_dec_ref(x_83); -return x_112; +size_t x_114; size_t x_115; lean_object* x_116; +x_114 = 0; +x_115 = lean_usize_of_nat(x_105); +x_116 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_LiveVars_0__Lean_IR_LiveVars_useExpr_spec__0(x_87, x_114, x_115, x_106, x_103); +lean_dec_ref(x_87); +return x_116; } } } } case 10: { -lean_object* x_113; lean_object* x_114; -x_113 = lean_ctor_get(x_1, 0); -lean_inc(x_113); +lean_object* x_117; lean_object* x_118; +x_117 = lean_ctor_get(x_1, 0); +lean_inc(x_117); lean_dec_ref(x_1); -x_114 = l___private_Lean_Compiler_IR_LiveVars_0__Lean_IR_LiveVars_useVar(x_113, x_2); -return x_114; +x_118 = l___private_Lean_Compiler_IR_LiveVars_0__Lean_IR_LiveVars_useVar(x_117, x_2); +return x_118; } case 11: { -lean_object* x_115; lean_object* x_116; +lean_object* x_119; lean_object* x_120; lean_dec_ref(x_1); -x_115 = lean_box(0); -x_116 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_116, 0, x_115); -lean_ctor_set(x_116, 1, x_2); -return x_116; +x_119 = lean_box(0); +x_120 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_120, 0, x_119); +lean_ctor_set(x_120, 1, x_2); +return x_120; } case 12: { -lean_object* x_117; lean_object* x_118; -x_117 = lean_ctor_get(x_1, 0); -lean_inc(x_117); +lean_object* x_121; lean_object* x_122; +x_121 = lean_ctor_get(x_1, 0); +lean_inc(x_121); lean_dec_ref(x_1); -x_118 = l___private_Lean_Compiler_IR_LiveVars_0__Lean_IR_LiveVars_useVar(x_117, x_2); -return x_118; +x_122 = l___private_Lean_Compiler_IR_LiveVars_0__Lean_IR_LiveVars_useVar(x_121, x_2); +return x_122; } default: { -lean_object* x_119; lean_object* x_120; -x_119 = lean_ctor_get(x_1, 1); -lean_inc(x_119); +lean_object* x_123; lean_object* x_124; +x_123 = lean_ctor_get(x_1, 1); +lean_inc(x_123); lean_dec_ref(x_1); -x_120 = l___private_Lean_Compiler_IR_LiveVars_0__Lean_IR_LiveVars_useVar(x_119, x_2); -return x_120; +x_124 = l___private_Lean_Compiler_IR_LiveVars_0__Lean_IR_LiveVars_useVar(x_123, x_2); +return x_124; } } block_18: @@ -5552,9 +5570,9 @@ goto block_52; block_45: { lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_nat_add(x_39, x_41); +x_42 = lean_nat_add(x_40, x_41); lean_dec(x_41); -lean_dec(x_39); +lean_dec(x_40); if (lean_is_scalar(x_36)) { x_43 = lean_alloc_ctor(0, 5, 0); } else { @@ -5573,7 +5591,7 @@ if (lean_is_scalar(x_26)) { lean_ctor_set(x_44, 0, x_38); lean_ctor_set(x_44, 1, x_28); lean_ctor_set(x_44, 2, x_29); -lean_ctor_set(x_44, 3, x_40); +lean_ctor_set(x_44, 3, x_39); lean_ctor_set(x_44, 4, x_43); return x_44; } @@ -5599,8 +5617,8 @@ if (lean_obj_tag(x_31) == 0) lean_object* x_50; x_50 = lean_ctor_get(x_31, 0); lean_inc(x_50); -x_39 = x_49; -x_40 = x_48; +x_39 = x_48; +x_40 = x_49; x_41 = x_50; goto block_45; } @@ -5608,8 +5626,8 @@ else { lean_object* x_51; x_51 = lean_unsigned_to_nat(0u); -x_39 = x_49; -x_40 = x_48; +x_39 = x_48; +x_40 = x_49; x_41 = x_51; goto block_45; } @@ -6517,9 +6535,9 @@ case 4: lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; x_38 = lean_ctor_get(x_1, 0); lean_inc(x_38); -x_39 = lean_ctor_get(x_1, 2); +x_39 = lean_ctor_get(x_1, 3); lean_inc(x_39); -x_40 = lean_ctor_get(x_1, 3); +x_40 = lean_ctor_get(x_1, 4); lean_inc(x_40); lean_dec_ref(x_1); x_41 = l___private_Lean_Compiler_IR_LiveVars_0__Lean_IR_LiveVars_visitFnBody(x_40, x_2, x_3); @@ -6538,9 +6556,9 @@ case 5: lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; x_46 = lean_ctor_get(x_1, 0); lean_inc(x_46); -x_47 = lean_ctor_get(x_1, 3); +x_47 = lean_ctor_get(x_1, 4); lean_inc(x_47); -x_48 = lean_ctor_get(x_1, 5); +x_48 = lean_ctor_get(x_1, 6); lean_inc(x_48); lean_dec_ref(x_1); x_49 = l___private_Lean_Compiler_IR_LiveVars_0__Lean_IR_LiveVars_visitFnBody(x_48, x_2, x_3); diff --git a/stage0/stdlib/Lean/Compiler/IR/NormIds.c b/stage0/stdlib/Lean/Compiler/IR/NormIds.c index 4ca8ca6daa45..bf4b365257fd 100644 --- a/stage0/stdlib/Lean/Compiler/IR/NormIds.c +++ b/stage0/stdlib/Lean/Compiler/IR/NormIds.c @@ -2163,229 +2163,238 @@ x_29 = !lean_is_exclusive(x_1); if (x_29 == 0) { lean_object* x_30; lean_object* x_31; -x_30 = lean_ctor_get(x_1, 1); +x_30 = lean_ctor_get(x_1, 2); x_31 = l_Lean_IR_NormalizeIds_normIndex(x_30, x_2); lean_dec(x_30); -lean_ctor_set(x_1, 1, x_31); +lean_ctor_set(x_1, 2, x_31); return x_1; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; x_32 = lean_ctor_get(x_1, 0); x_33 = lean_ctor_get(x_1, 1); +x_34 = lean_ctor_get(x_1, 2); +lean_inc(x_34); lean_inc(x_33); lean_inc(x_32); lean_dec(x_1); -x_34 = l_Lean_IR_NormalizeIds_normIndex(x_33, x_2); -lean_dec(x_33); -x_35 = lean_alloc_ctor(3, 2, 0); -lean_ctor_set(x_35, 0, x_32); -lean_ctor_set(x_35, 1, x_34); -return x_35; +x_35 = l_Lean_IR_NormalizeIds_normIndex(x_34, x_2); +lean_dec(x_34); +x_36 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_36, 0, x_32); +lean_ctor_set(x_36, 1, x_33); +lean_ctor_set(x_36, 2, x_35); +return x_36; } } case 4: { -uint8_t x_36; -x_36 = !lean_is_exclusive(x_1); -if (x_36 == 0) +uint8_t x_37; +x_37 = !lean_is_exclusive(x_1); +if (x_37 == 0) { -lean_object* x_37; lean_object* x_38; -x_37 = lean_ctor_get(x_1, 1); -x_38 = l_Lean_IR_NormalizeIds_normIndex(x_37, x_2); -lean_dec(x_37); -lean_ctor_set(x_1, 1, x_38); +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_1, 2); +x_39 = l_Lean_IR_NormalizeIds_normIndex(x_38, x_2); +lean_dec(x_38); +lean_ctor_set(x_1, 2, x_39); return x_1; } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_39 = lean_ctor_get(x_1, 0); -x_40 = lean_ctor_get(x_1, 1); +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_40 = lean_ctor_get(x_1, 0); +x_41 = lean_ctor_get(x_1, 1); +x_42 = lean_ctor_get(x_1, 2); +lean_inc(x_42); +lean_inc(x_41); lean_inc(x_40); -lean_inc(x_39); lean_dec(x_1); -x_41 = l_Lean_IR_NormalizeIds_normIndex(x_40, x_2); -lean_dec(x_40); -x_42 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_42, 0, x_39); -lean_ctor_set(x_42, 1, x_41); -return x_42; +x_43 = l_Lean_IR_NormalizeIds_normIndex(x_42, x_2); +lean_dec(x_42); +x_44 = lean_alloc_ctor(4, 3, 0); +lean_ctor_set(x_44, 0, x_40); +lean_ctor_set(x_44, 1, x_41); +lean_ctor_set(x_44, 2, x_43); +return x_44; } } case 5: { -uint8_t x_43; -x_43 = !lean_is_exclusive(x_1); -if (x_43 == 0) +uint8_t x_45; +x_45 = !lean_is_exclusive(x_1); +if (x_45 == 0) { -lean_object* x_44; lean_object* x_45; -x_44 = lean_ctor_get(x_1, 2); -x_45 = l_Lean_IR_NormalizeIds_normIndex(x_44, x_2); -lean_dec(x_44); -lean_ctor_set(x_1, 2, x_45); +lean_object* x_46; lean_object* x_47; +x_46 = lean_ctor_get(x_1, 3); +x_47 = l_Lean_IR_NormalizeIds_normIndex(x_46, x_2); +lean_dec(x_46); +lean_ctor_set(x_1, 3, x_47); return x_1; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_46 = lean_ctor_get(x_1, 0); -x_47 = lean_ctor_get(x_1, 1); -x_48 = lean_ctor_get(x_1, 2); +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_48 = lean_ctor_get(x_1, 0); +x_49 = lean_ctor_get(x_1, 1); +x_50 = lean_ctor_get(x_1, 2); +x_51 = lean_ctor_get(x_1, 3); +lean_inc(x_51); +lean_inc(x_50); +lean_inc(x_49); lean_inc(x_48); -lean_inc(x_47); -lean_inc(x_46); lean_dec(x_1); -x_49 = l_Lean_IR_NormalizeIds_normIndex(x_48, x_2); -lean_dec(x_48); -x_50 = lean_alloc_ctor(5, 3, 0); -lean_ctor_set(x_50, 0, x_46); -lean_ctor_set(x_50, 1, x_47); -lean_ctor_set(x_50, 2, x_49); -return x_50; +x_52 = l_Lean_IR_NormalizeIds_normIndex(x_51, x_2); +lean_dec(x_51); +x_53 = lean_alloc_ctor(5, 4, 0); +lean_ctor_set(x_53, 0, x_48); +lean_ctor_set(x_53, 1, x_49); +lean_ctor_set(x_53, 2, x_50); +lean_ctor_set(x_53, 3, x_52); +return x_53; } } case 6: { -uint8_t x_51; -x_51 = !lean_is_exclusive(x_1); -if (x_51 == 0) +uint8_t x_54; +x_54 = !lean_is_exclusive(x_1); +if (x_54 == 0) { -lean_object* x_52; lean_object* x_53; -x_52 = lean_ctor_get(x_1, 1); -x_53 = l_Lean_IR_NormalizeIds_normArgs(x_52, x_2); -lean_ctor_set(x_1, 1, x_53); +lean_object* x_55; lean_object* x_56; +x_55 = lean_ctor_get(x_1, 1); +x_56 = l_Lean_IR_NormalizeIds_normArgs(x_55, x_2); +lean_ctor_set(x_1, 1, x_56); return x_1; } else { -lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_54 = lean_ctor_get(x_1, 0); -x_55 = lean_ctor_get(x_1, 1); -lean_inc(x_55); -lean_inc(x_54); +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_57 = lean_ctor_get(x_1, 0); +x_58 = lean_ctor_get(x_1, 1); +lean_inc(x_58); +lean_inc(x_57); lean_dec(x_1); -x_56 = l_Lean_IR_NormalizeIds_normArgs(x_55, x_2); -x_57 = lean_alloc_ctor(6, 2, 0); -lean_ctor_set(x_57, 0, x_54); -lean_ctor_set(x_57, 1, x_56); -return x_57; +x_59 = l_Lean_IR_NormalizeIds_normArgs(x_58, x_2); +x_60 = lean_alloc_ctor(6, 2, 0); +lean_ctor_set(x_60, 0, x_57); +lean_ctor_set(x_60, 1, x_59); +return x_60; } } case 7: { -uint8_t x_58; -x_58 = !lean_is_exclusive(x_1); -if (x_58 == 0) +uint8_t x_61; +x_61 = !lean_is_exclusive(x_1); +if (x_61 == 0) { -lean_object* x_59; lean_object* x_60; -x_59 = lean_ctor_get(x_1, 1); -x_60 = l_Lean_IR_NormalizeIds_normArgs(x_59, x_2); -lean_ctor_set(x_1, 1, x_60); +lean_object* x_62; lean_object* x_63; +x_62 = lean_ctor_get(x_1, 1); +x_63 = l_Lean_IR_NormalizeIds_normArgs(x_62, x_2); +lean_ctor_set(x_1, 1, x_63); return x_1; } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_61 = lean_ctor_get(x_1, 0); -x_62 = lean_ctor_get(x_1, 1); -lean_inc(x_62); -lean_inc(x_61); +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_64 = lean_ctor_get(x_1, 0); +x_65 = lean_ctor_get(x_1, 1); +lean_inc(x_65); +lean_inc(x_64); lean_dec(x_1); -x_63 = l_Lean_IR_NormalizeIds_normArgs(x_62, x_2); -x_64 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_64, 0, x_61); -lean_ctor_set(x_64, 1, x_63); -return x_64; +x_66 = l_Lean_IR_NormalizeIds_normArgs(x_65, x_2); +x_67 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_67, 0, x_64); +lean_ctor_set(x_67, 1, x_66); +return x_67; } } case 8: { -uint8_t x_65; -x_65 = !lean_is_exclusive(x_1); -if (x_65 == 0) +uint8_t x_68; +x_68 = !lean_is_exclusive(x_1); +if (x_68 == 0) { -lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_66 = lean_ctor_get(x_1, 0); -x_67 = lean_ctor_get(x_1, 1); -x_68 = l_Lean_IR_NormalizeIds_normIndex(x_66, x_2); -lean_dec(x_66); -x_69 = l_Lean_IR_NormalizeIds_normArgs(x_67, x_2); -lean_ctor_set(x_1, 1, x_69); -lean_ctor_set(x_1, 0, x_68); +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_69 = lean_ctor_get(x_1, 0); +x_70 = lean_ctor_get(x_1, 1); +x_71 = l_Lean_IR_NormalizeIds_normIndex(x_69, x_2); +lean_dec(x_69); +x_72 = l_Lean_IR_NormalizeIds_normArgs(x_70, x_2); +lean_ctor_set(x_1, 1, x_72); +lean_ctor_set(x_1, 0, x_71); return x_1; } else { -lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_70 = lean_ctor_get(x_1, 0); -x_71 = lean_ctor_get(x_1, 1); -lean_inc(x_71); -lean_inc(x_70); +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_73 = lean_ctor_get(x_1, 0); +x_74 = lean_ctor_get(x_1, 1); +lean_inc(x_74); +lean_inc(x_73); lean_dec(x_1); -x_72 = l_Lean_IR_NormalizeIds_normIndex(x_70, x_2); -lean_dec(x_70); -x_73 = l_Lean_IR_NormalizeIds_normArgs(x_71, x_2); -x_74 = lean_alloc_ctor(8, 2, 0); -lean_ctor_set(x_74, 0, x_72); -lean_ctor_set(x_74, 1, x_73); -return x_74; +x_75 = l_Lean_IR_NormalizeIds_normIndex(x_73, x_2); +lean_dec(x_73); +x_76 = l_Lean_IR_NormalizeIds_normArgs(x_74, x_2); +x_77 = lean_alloc_ctor(8, 2, 0); +lean_ctor_set(x_77, 0, x_75); +lean_ctor_set(x_77, 1, x_76); +return x_77; } } case 9: { -uint8_t x_75; -x_75 = !lean_is_exclusive(x_1); -if (x_75 == 0) +uint8_t x_78; +x_78 = !lean_is_exclusive(x_1); +if (x_78 == 0) { -lean_object* x_76; lean_object* x_77; -x_76 = lean_ctor_get(x_1, 1); -x_77 = l_Lean_IR_NormalizeIds_normIndex(x_76, x_2); -lean_dec(x_76); -lean_ctor_set(x_1, 1, x_77); +lean_object* x_79; lean_object* x_80; +x_79 = lean_ctor_get(x_1, 1); +x_80 = l_Lean_IR_NormalizeIds_normIndex(x_79, x_2); +lean_dec(x_79); +lean_ctor_set(x_1, 1, x_80); return x_1; } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_78 = lean_ctor_get(x_1, 0); -x_79 = lean_ctor_get(x_1, 1); -lean_inc(x_79); -lean_inc(x_78); +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_81 = lean_ctor_get(x_1, 0); +x_82 = lean_ctor_get(x_1, 1); +lean_inc(x_82); +lean_inc(x_81); lean_dec(x_1); -x_80 = l_Lean_IR_NormalizeIds_normIndex(x_79, x_2); -lean_dec(x_79); -x_81 = lean_alloc_ctor(9, 2, 0); -lean_ctor_set(x_81, 0, x_78); -lean_ctor_set(x_81, 1, x_80); -return x_81; +x_83 = l_Lean_IR_NormalizeIds_normIndex(x_82, x_2); +lean_dec(x_82); +x_84 = lean_alloc_ctor(9, 2, 0); +lean_ctor_set(x_84, 0, x_81); +lean_ctor_set(x_84, 1, x_83); +return x_84; } } case 10: { -uint8_t x_82; -x_82 = !lean_is_exclusive(x_1); -if (x_82 == 0) +uint8_t x_85; +x_85 = !lean_is_exclusive(x_1); +if (x_85 == 0) { -lean_object* x_83; lean_object* x_84; -x_83 = lean_ctor_get(x_1, 0); -x_84 = l_Lean_IR_NormalizeIds_normIndex(x_83, x_2); -lean_dec(x_83); -lean_ctor_set(x_1, 0, x_84); +lean_object* x_86; lean_object* x_87; +x_86 = lean_ctor_get(x_1, 0); +x_87 = l_Lean_IR_NormalizeIds_normIndex(x_86, x_2); +lean_dec(x_86); +lean_ctor_set(x_1, 0, x_87); return x_1; } else { -lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_85 = lean_ctor_get(x_1, 0); -lean_inc(x_85); +lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_88 = lean_ctor_get(x_1, 0); +lean_inc(x_88); lean_dec(x_1); -x_86 = l_Lean_IR_NormalizeIds_normIndex(x_85, x_2); -lean_dec(x_85); -x_87 = lean_alloc_ctor(10, 1, 0); -lean_ctor_set(x_87, 0, x_86); -return x_87; +x_89 = l_Lean_IR_NormalizeIds_normIndex(x_88, x_2); +lean_dec(x_88); +x_90 = lean_alloc_ctor(10, 1, 0); +lean_ctor_set(x_90, 0, x_89); +return x_90; } } case 11: @@ -2394,28 +2403,28 @@ return x_1; } default: { -uint8_t x_88; -x_88 = !lean_is_exclusive(x_1); -if (x_88 == 0) +uint8_t x_91; +x_91 = !lean_is_exclusive(x_1); +if (x_91 == 0) { -lean_object* x_89; lean_object* x_90; -x_89 = lean_ctor_get(x_1, 0); -x_90 = l_Lean_IR_NormalizeIds_normIndex(x_89, x_2); -lean_dec(x_89); -lean_ctor_set(x_1, 0, x_90); +lean_object* x_92; lean_object* x_93; +x_92 = lean_ctor_get(x_1, 0); +x_93 = l_Lean_IR_NormalizeIds_normIndex(x_92, x_2); +lean_dec(x_92); +lean_ctor_set(x_1, 0, x_93); return x_1; } else { -lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_91 = lean_ctor_get(x_1, 0); -lean_inc(x_91); +lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_94 = lean_ctor_get(x_1, 0); +lean_inc(x_94); lean_dec(x_1); -x_92 = l_Lean_IR_NormalizeIds_normIndex(x_91, x_2); -lean_dec(x_91); -x_93 = lean_alloc_ctor(12, 1, 0); -lean_ctor_set(x_93, 0, x_92); -return x_93; +x_95 = l_Lean_IR_NormalizeIds_normIndex(x_94, x_2); +lean_dec(x_94); +x_96 = lean_alloc_ctor(12, 1, 0); +lean_ctor_set(x_96, 0, x_95); +return x_96; } } } @@ -3501,8 +3510,8 @@ if (x_122 == 0) { lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; uint8_t x_127; x_123 = lean_ctor_get(x_1, 0); -x_124 = lean_ctor_get(x_1, 2); -x_125 = lean_ctor_get(x_1, 3); +x_124 = lean_ctor_get(x_1, 3); +x_125 = lean_ctor_get(x_1, 4); lean_inc(x_2); x_126 = l_Lean_IR_NormalizeIds_normFnBody(x_125, x_2, x_3); x_127 = !lean_is_exclusive(x_126); @@ -3515,8 +3524,8 @@ lean_dec(x_123); x_130 = l_Lean_IR_NormalizeIds_normIndex(x_124, x_2); lean_dec(x_2); lean_dec(x_124); -lean_ctor_set(x_1, 3, x_128); -lean_ctor_set(x_1, 2, x_130); +lean_ctor_set(x_1, 4, x_128); +lean_ctor_set(x_1, 3, x_130); lean_ctor_set(x_1, 0, x_129); lean_ctor_set(x_126, 0, x_1); return x_126; @@ -3534,8 +3543,8 @@ lean_dec(x_123); x_134 = l_Lean_IR_NormalizeIds_normIndex(x_124, x_2); lean_dec(x_2); lean_dec(x_124); -lean_ctor_set(x_1, 3, x_131); -lean_ctor_set(x_1, 2, x_134); +lean_ctor_set(x_1, 4, x_131); +lean_ctor_set(x_1, 3, x_134); lean_ctor_set(x_1, 0, x_133); x_135 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_135, 0, x_1); @@ -3545,581 +3554,587 @@ return x_135; } else { -lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; +lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; x_136 = lean_ctor_get(x_1, 0); x_137 = lean_ctor_get(x_1, 1); x_138 = lean_ctor_get(x_1, 2); x_139 = lean_ctor_get(x_1, 3); +x_140 = lean_ctor_get(x_1, 4); +lean_inc(x_140); lean_inc(x_139); lean_inc(x_138); lean_inc(x_137); lean_inc(x_136); lean_dec(x_1); lean_inc(x_2); -x_140 = l_Lean_IR_NormalizeIds_normFnBody(x_139, x_2, x_3); -x_141 = lean_ctor_get(x_140, 0); -lean_inc(x_141); -x_142 = lean_ctor_get(x_140, 1); +x_141 = l_Lean_IR_NormalizeIds_normFnBody(x_140, x_2, x_3); +x_142 = lean_ctor_get(x_141, 0); lean_inc(x_142); -if (lean_is_exclusive(x_140)) { - lean_ctor_release(x_140, 0); - lean_ctor_release(x_140, 1); - x_143 = x_140; +x_143 = lean_ctor_get(x_141, 1); +lean_inc(x_143); +if (lean_is_exclusive(x_141)) { + lean_ctor_release(x_141, 0); + lean_ctor_release(x_141, 1); + x_144 = x_141; } else { - lean_dec_ref(x_140); - x_143 = lean_box(0); + lean_dec_ref(x_141); + x_144 = lean_box(0); } -x_144 = l_Lean_IR_NormalizeIds_normIndex(x_136, x_2); +x_145 = l_Lean_IR_NormalizeIds_normIndex(x_136, x_2); lean_dec(x_136); -x_145 = l_Lean_IR_NormalizeIds_normIndex(x_138, x_2); +x_146 = l_Lean_IR_NormalizeIds_normIndex(x_139, x_2); lean_dec(x_2); -lean_dec(x_138); -x_146 = lean_alloc_ctor(4, 4, 0); -lean_ctor_set(x_146, 0, x_144); -lean_ctor_set(x_146, 1, x_137); -lean_ctor_set(x_146, 2, x_145); -lean_ctor_set(x_146, 3, x_141); -if (lean_is_scalar(x_143)) { - x_147 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_139); +x_147 = lean_alloc_ctor(4, 5, 0); +lean_ctor_set(x_147, 0, x_145); +lean_ctor_set(x_147, 1, x_137); +lean_ctor_set(x_147, 2, x_138); +lean_ctor_set(x_147, 3, x_146); +lean_ctor_set(x_147, 4, x_142); +if (lean_is_scalar(x_144)) { + x_148 = lean_alloc_ctor(0, 2, 0); } else { - x_147 = x_143; + x_148 = x_144; } -lean_ctor_set(x_147, 0, x_146); -lean_ctor_set(x_147, 1, x_142); -return x_147; +lean_ctor_set(x_148, 0, x_147); +lean_ctor_set(x_148, 1, x_143); +return x_148; } } case 5: { -uint8_t x_148; -x_148 = !lean_is_exclusive(x_1); -if (x_148 == 0) +uint8_t x_149; +x_149 = !lean_is_exclusive(x_1); +if (x_149 == 0) { -lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; uint8_t x_153; -x_149 = lean_ctor_get(x_1, 0); -x_150 = lean_ctor_get(x_1, 3); -x_151 = lean_ctor_get(x_1, 5); +lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; uint8_t x_154; +x_150 = lean_ctor_get(x_1, 0); +x_151 = lean_ctor_get(x_1, 4); +x_152 = lean_ctor_get(x_1, 6); lean_inc(x_2); -x_152 = l_Lean_IR_NormalizeIds_normFnBody(x_151, x_2, x_3); -x_153 = !lean_is_exclusive(x_152); -if (x_153 == 0) -{ -lean_object* x_154; lean_object* x_155; lean_object* x_156; -x_154 = lean_ctor_get(x_152, 0); -x_155 = l_Lean_IR_NormalizeIds_normIndex(x_149, x_2); -lean_dec(x_149); +x_153 = l_Lean_IR_NormalizeIds_normFnBody(x_152, x_2, x_3); +x_154 = !lean_is_exclusive(x_153); +if (x_154 == 0) +{ +lean_object* x_155; lean_object* x_156; lean_object* x_157; +x_155 = lean_ctor_get(x_153, 0); x_156 = l_Lean_IR_NormalizeIds_normIndex(x_150, x_2); -lean_dec(x_2); lean_dec(x_150); -lean_ctor_set(x_1, 5, x_154); -lean_ctor_set(x_1, 3, x_156); -lean_ctor_set(x_1, 0, x_155); -lean_ctor_set(x_152, 0, x_1); -return x_152; +x_157 = l_Lean_IR_NormalizeIds_normIndex(x_151, x_2); +lean_dec(x_2); +lean_dec(x_151); +lean_ctor_set(x_1, 6, x_155); +lean_ctor_set(x_1, 4, x_157); +lean_ctor_set(x_1, 0, x_156); +lean_ctor_set(x_153, 0, x_1); +return x_153; } else { -lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; -x_157 = lean_ctor_get(x_152, 0); -x_158 = lean_ctor_get(x_152, 1); +lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; +x_158 = lean_ctor_get(x_153, 0); +x_159 = lean_ctor_get(x_153, 1); +lean_inc(x_159); lean_inc(x_158); -lean_inc(x_157); -lean_dec(x_152); -x_159 = l_Lean_IR_NormalizeIds_normIndex(x_149, x_2); -lean_dec(x_149); +lean_dec(x_153); x_160 = l_Lean_IR_NormalizeIds_normIndex(x_150, x_2); -lean_dec(x_2); lean_dec(x_150); -lean_ctor_set(x_1, 5, x_157); -lean_ctor_set(x_1, 3, x_160); -lean_ctor_set(x_1, 0, x_159); -x_161 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_161, 0, x_1); -lean_ctor_set(x_161, 1, x_158); -return x_161; -} -} -else -{ -lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; -x_162 = lean_ctor_get(x_1, 0); -x_163 = lean_ctor_get(x_1, 1); -x_164 = lean_ctor_get(x_1, 2); -x_165 = lean_ctor_get(x_1, 3); -x_166 = lean_ctor_get(x_1, 4); -x_167 = lean_ctor_get(x_1, 5); +x_161 = l_Lean_IR_NormalizeIds_normIndex(x_151, x_2); +lean_dec(x_2); +lean_dec(x_151); +lean_ctor_set(x_1, 6, x_158); +lean_ctor_set(x_1, 4, x_161); +lean_ctor_set(x_1, 0, x_160); +x_162 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_162, 0, x_1); +lean_ctor_set(x_162, 1, x_159); +return x_162; +} +} +else +{ +lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; +x_163 = lean_ctor_get(x_1, 0); +x_164 = lean_ctor_get(x_1, 1); +x_165 = lean_ctor_get(x_1, 2); +x_166 = lean_ctor_get(x_1, 3); +x_167 = lean_ctor_get(x_1, 4); +x_168 = lean_ctor_get(x_1, 5); +x_169 = lean_ctor_get(x_1, 6); +lean_inc(x_169); +lean_inc(x_168); lean_inc(x_167); lean_inc(x_166); lean_inc(x_165); lean_inc(x_164); lean_inc(x_163); -lean_inc(x_162); lean_dec(x_1); lean_inc(x_2); -x_168 = l_Lean_IR_NormalizeIds_normFnBody(x_167, x_2, x_3); -x_169 = lean_ctor_get(x_168, 0); -lean_inc(x_169); -x_170 = lean_ctor_get(x_168, 1); -lean_inc(x_170); -if (lean_is_exclusive(x_168)) { - lean_ctor_release(x_168, 0); - lean_ctor_release(x_168, 1); - x_171 = x_168; +x_170 = l_Lean_IR_NormalizeIds_normFnBody(x_169, x_2, x_3); +x_171 = lean_ctor_get(x_170, 0); +lean_inc(x_171); +x_172 = lean_ctor_get(x_170, 1); +lean_inc(x_172); +if (lean_is_exclusive(x_170)) { + lean_ctor_release(x_170, 0); + lean_ctor_release(x_170, 1); + x_173 = x_170; } else { - lean_dec_ref(x_168); - x_171 = lean_box(0); + lean_dec_ref(x_170); + x_173 = lean_box(0); } -x_172 = l_Lean_IR_NormalizeIds_normIndex(x_162, x_2); -lean_dec(x_162); -x_173 = l_Lean_IR_NormalizeIds_normIndex(x_165, x_2); +x_174 = l_Lean_IR_NormalizeIds_normIndex(x_163, x_2); +lean_dec(x_163); +x_175 = l_Lean_IR_NormalizeIds_normIndex(x_167, x_2); lean_dec(x_2); -lean_dec(x_165); -x_174 = lean_alloc_ctor(5, 6, 0); -lean_ctor_set(x_174, 0, x_172); -lean_ctor_set(x_174, 1, x_163); -lean_ctor_set(x_174, 2, x_164); -lean_ctor_set(x_174, 3, x_173); -lean_ctor_set(x_174, 4, x_166); -lean_ctor_set(x_174, 5, x_169); -if (lean_is_scalar(x_171)) { - x_175 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_167); +x_176 = lean_alloc_ctor(5, 7, 0); +lean_ctor_set(x_176, 0, x_174); +lean_ctor_set(x_176, 1, x_164); +lean_ctor_set(x_176, 2, x_165); +lean_ctor_set(x_176, 3, x_166); +lean_ctor_set(x_176, 4, x_175); +lean_ctor_set(x_176, 5, x_168); +lean_ctor_set(x_176, 6, x_171); +if (lean_is_scalar(x_173)) { + x_177 = lean_alloc_ctor(0, 2, 0); } else { - x_175 = x_171; + x_177 = x_173; } -lean_ctor_set(x_175, 0, x_174); -lean_ctor_set(x_175, 1, x_170); -return x_175; +lean_ctor_set(x_177, 0, x_176); +lean_ctor_set(x_177, 1, x_172); +return x_177; } } case 6: { -uint8_t x_176; -x_176 = !lean_is_exclusive(x_1); -if (x_176 == 0) +uint8_t x_178; +x_178 = !lean_is_exclusive(x_1); +if (x_178 == 0) { -lean_object* x_177; lean_object* x_178; lean_object* x_179; uint8_t x_180; -x_177 = lean_ctor_get(x_1, 0); -x_178 = lean_ctor_get(x_1, 2); +lean_object* x_179; lean_object* x_180; lean_object* x_181; uint8_t x_182; +x_179 = lean_ctor_get(x_1, 0); +x_180 = lean_ctor_get(x_1, 2); lean_inc(x_2); -x_179 = l_Lean_IR_NormalizeIds_normFnBody(x_178, x_2, x_3); -x_180 = !lean_is_exclusive(x_179); -if (x_180 == 0) +x_181 = l_Lean_IR_NormalizeIds_normFnBody(x_180, x_2, x_3); +x_182 = !lean_is_exclusive(x_181); +if (x_182 == 0) { -lean_object* x_181; lean_object* x_182; -x_181 = lean_ctor_get(x_179, 0); -x_182 = l_Lean_IR_NormalizeIds_normIndex(x_177, x_2); +lean_object* x_183; lean_object* x_184; +x_183 = lean_ctor_get(x_181, 0); +x_184 = l_Lean_IR_NormalizeIds_normIndex(x_179, x_2); lean_dec(x_2); -lean_dec(x_177); -lean_ctor_set(x_1, 2, x_181); -lean_ctor_set(x_1, 0, x_182); -lean_ctor_set(x_179, 0, x_1); -return x_179; +lean_dec(x_179); +lean_ctor_set(x_1, 2, x_183); +lean_ctor_set(x_1, 0, x_184); +lean_ctor_set(x_181, 0, x_1); +return x_181; } else { -lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; -x_183 = lean_ctor_get(x_179, 0); -x_184 = lean_ctor_get(x_179, 1); -lean_inc(x_184); -lean_inc(x_183); -lean_dec(x_179); -x_185 = l_Lean_IR_NormalizeIds_normIndex(x_177, x_2); +lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; +x_185 = lean_ctor_get(x_181, 0); +x_186 = lean_ctor_get(x_181, 1); +lean_inc(x_186); +lean_inc(x_185); +lean_dec(x_181); +x_187 = l_Lean_IR_NormalizeIds_normIndex(x_179, x_2); lean_dec(x_2); -lean_dec(x_177); -lean_ctor_set(x_1, 2, x_183); -lean_ctor_set(x_1, 0, x_185); -x_186 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_186, 0, x_1); -lean_ctor_set(x_186, 1, x_184); -return x_186; +lean_dec(x_179); +lean_ctor_set(x_1, 2, x_185); +lean_ctor_set(x_1, 0, x_187); +x_188 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_188, 0, x_1); +lean_ctor_set(x_188, 1, x_186); +return x_188; } } else { -lean_object* x_187; lean_object* x_188; uint8_t x_189; uint8_t x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; -x_187 = lean_ctor_get(x_1, 0); -x_188 = lean_ctor_get(x_1, 1); -x_189 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); -x_190 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 1); -x_191 = lean_ctor_get(x_1, 2); -lean_inc(x_191); -lean_inc(x_188); -lean_inc(x_187); +lean_object* x_189; lean_object* x_190; uint8_t x_191; uint8_t x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; +x_189 = lean_ctor_get(x_1, 0); +x_190 = lean_ctor_get(x_1, 1); +x_191 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); +x_192 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 1); +x_193 = lean_ctor_get(x_1, 2); +lean_inc(x_193); +lean_inc(x_190); +lean_inc(x_189); lean_dec(x_1); lean_inc(x_2); -x_192 = l_Lean_IR_NormalizeIds_normFnBody(x_191, x_2, x_3); -x_193 = lean_ctor_get(x_192, 0); -lean_inc(x_193); -x_194 = lean_ctor_get(x_192, 1); -lean_inc(x_194); -if (lean_is_exclusive(x_192)) { - lean_ctor_release(x_192, 0); - lean_ctor_release(x_192, 1); - x_195 = x_192; +x_194 = l_Lean_IR_NormalizeIds_normFnBody(x_193, x_2, x_3); +x_195 = lean_ctor_get(x_194, 0); +lean_inc(x_195); +x_196 = lean_ctor_get(x_194, 1); +lean_inc(x_196); +if (lean_is_exclusive(x_194)) { + lean_ctor_release(x_194, 0); + lean_ctor_release(x_194, 1); + x_197 = x_194; } else { - lean_dec_ref(x_192); - x_195 = lean_box(0); + lean_dec_ref(x_194); + x_197 = lean_box(0); } -x_196 = l_Lean_IR_NormalizeIds_normIndex(x_187, x_2); +x_198 = l_Lean_IR_NormalizeIds_normIndex(x_189, x_2); lean_dec(x_2); -lean_dec(x_187); -x_197 = lean_alloc_ctor(6, 3, 2); -lean_ctor_set(x_197, 0, x_196); -lean_ctor_set(x_197, 1, x_188); -lean_ctor_set(x_197, 2, x_193); -lean_ctor_set_uint8(x_197, sizeof(void*)*3, x_189); -lean_ctor_set_uint8(x_197, sizeof(void*)*3 + 1, x_190); -if (lean_is_scalar(x_195)) { - x_198 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_189); +x_199 = lean_alloc_ctor(6, 3, 2); +lean_ctor_set(x_199, 0, x_198); +lean_ctor_set(x_199, 1, x_190); +lean_ctor_set(x_199, 2, x_195); +lean_ctor_set_uint8(x_199, sizeof(void*)*3, x_191); +lean_ctor_set_uint8(x_199, sizeof(void*)*3 + 1, x_192); +if (lean_is_scalar(x_197)) { + x_200 = lean_alloc_ctor(0, 2, 0); } else { - x_198 = x_195; + x_200 = x_197; } -lean_ctor_set(x_198, 0, x_197); -lean_ctor_set(x_198, 1, x_194); -return x_198; +lean_ctor_set(x_200, 0, x_199); +lean_ctor_set(x_200, 1, x_196); +return x_200; } } case 7: { -uint8_t x_199; -x_199 = !lean_is_exclusive(x_1); -if (x_199 == 0) +uint8_t x_201; +x_201 = !lean_is_exclusive(x_1); +if (x_201 == 0) { -lean_object* x_200; lean_object* x_201; lean_object* x_202; uint8_t x_203; -x_200 = lean_ctor_get(x_1, 0); -x_201 = lean_ctor_get(x_1, 2); +lean_object* x_202; lean_object* x_203; lean_object* x_204; uint8_t x_205; +x_202 = lean_ctor_get(x_1, 0); +x_203 = lean_ctor_get(x_1, 2); lean_inc(x_2); -x_202 = l_Lean_IR_NormalizeIds_normFnBody(x_201, x_2, x_3); -x_203 = !lean_is_exclusive(x_202); -if (x_203 == 0) +x_204 = l_Lean_IR_NormalizeIds_normFnBody(x_203, x_2, x_3); +x_205 = !lean_is_exclusive(x_204); +if (x_205 == 0) { -lean_object* x_204; lean_object* x_205; -x_204 = lean_ctor_get(x_202, 0); -x_205 = l_Lean_IR_NormalizeIds_normIndex(x_200, x_2); +lean_object* x_206; lean_object* x_207; +x_206 = lean_ctor_get(x_204, 0); +x_207 = l_Lean_IR_NormalizeIds_normIndex(x_202, x_2); lean_dec(x_2); -lean_dec(x_200); -lean_ctor_set(x_1, 2, x_204); -lean_ctor_set(x_1, 0, x_205); -lean_ctor_set(x_202, 0, x_1); -return x_202; +lean_dec(x_202); +lean_ctor_set(x_1, 2, x_206); +lean_ctor_set(x_1, 0, x_207); +lean_ctor_set(x_204, 0, x_1); +return x_204; } else { -lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; -x_206 = lean_ctor_get(x_202, 0); -x_207 = lean_ctor_get(x_202, 1); -lean_inc(x_207); -lean_inc(x_206); -lean_dec(x_202); -x_208 = l_Lean_IR_NormalizeIds_normIndex(x_200, x_2); +lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; +x_208 = lean_ctor_get(x_204, 0); +x_209 = lean_ctor_get(x_204, 1); +lean_inc(x_209); +lean_inc(x_208); +lean_dec(x_204); +x_210 = l_Lean_IR_NormalizeIds_normIndex(x_202, x_2); lean_dec(x_2); -lean_dec(x_200); -lean_ctor_set(x_1, 2, x_206); -lean_ctor_set(x_1, 0, x_208); -x_209 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_209, 0, x_1); -lean_ctor_set(x_209, 1, x_207); -return x_209; +lean_dec(x_202); +lean_ctor_set(x_1, 2, x_208); +lean_ctor_set(x_1, 0, x_210); +x_211 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_211, 0, x_1); +lean_ctor_set(x_211, 1, x_209); +return x_211; } } else { -lean_object* x_210; lean_object* x_211; uint8_t x_212; uint8_t x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; -x_210 = lean_ctor_get(x_1, 0); -x_211 = lean_ctor_get(x_1, 1); -x_212 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); -x_213 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 1); -x_214 = lean_ctor_get(x_1, 2); -lean_inc(x_214); -lean_inc(x_211); -lean_inc(x_210); +lean_object* x_212; lean_object* x_213; uint8_t x_214; uint8_t x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; +x_212 = lean_ctor_get(x_1, 0); +x_213 = lean_ctor_get(x_1, 1); +x_214 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); +x_215 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 1); +x_216 = lean_ctor_get(x_1, 2); +lean_inc(x_216); +lean_inc(x_213); +lean_inc(x_212); lean_dec(x_1); lean_inc(x_2); -x_215 = l_Lean_IR_NormalizeIds_normFnBody(x_214, x_2, x_3); -x_216 = lean_ctor_get(x_215, 0); -lean_inc(x_216); -x_217 = lean_ctor_get(x_215, 1); -lean_inc(x_217); -if (lean_is_exclusive(x_215)) { - lean_ctor_release(x_215, 0); - lean_ctor_release(x_215, 1); - x_218 = x_215; +x_217 = l_Lean_IR_NormalizeIds_normFnBody(x_216, x_2, x_3); +x_218 = lean_ctor_get(x_217, 0); +lean_inc(x_218); +x_219 = lean_ctor_get(x_217, 1); +lean_inc(x_219); +if (lean_is_exclusive(x_217)) { + lean_ctor_release(x_217, 0); + lean_ctor_release(x_217, 1); + x_220 = x_217; } else { - lean_dec_ref(x_215); - x_218 = lean_box(0); + lean_dec_ref(x_217); + x_220 = lean_box(0); } -x_219 = l_Lean_IR_NormalizeIds_normIndex(x_210, x_2); +x_221 = l_Lean_IR_NormalizeIds_normIndex(x_212, x_2); lean_dec(x_2); -lean_dec(x_210); -x_220 = lean_alloc_ctor(7, 3, 2); -lean_ctor_set(x_220, 0, x_219); -lean_ctor_set(x_220, 1, x_211); -lean_ctor_set(x_220, 2, x_216); -lean_ctor_set_uint8(x_220, sizeof(void*)*3, x_212); -lean_ctor_set_uint8(x_220, sizeof(void*)*3 + 1, x_213); -if (lean_is_scalar(x_218)) { - x_221 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_212); +x_222 = lean_alloc_ctor(7, 3, 2); +lean_ctor_set(x_222, 0, x_221); +lean_ctor_set(x_222, 1, x_213); +lean_ctor_set(x_222, 2, x_218); +lean_ctor_set_uint8(x_222, sizeof(void*)*3, x_214); +lean_ctor_set_uint8(x_222, sizeof(void*)*3 + 1, x_215); +if (lean_is_scalar(x_220)) { + x_223 = lean_alloc_ctor(0, 2, 0); } else { - x_221 = x_218; + x_223 = x_220; } -lean_ctor_set(x_221, 0, x_220); -lean_ctor_set(x_221, 1, x_217); -return x_221; +lean_ctor_set(x_223, 0, x_222); +lean_ctor_set(x_223, 1, x_219); +return x_223; } } case 8: { -uint8_t x_222; -x_222 = !lean_is_exclusive(x_1); -if (x_222 == 0) +uint8_t x_224; +x_224 = !lean_is_exclusive(x_1); +if (x_224 == 0) { -lean_object* x_223; lean_object* x_224; lean_object* x_225; uint8_t x_226; -x_223 = lean_ctor_get(x_1, 0); -x_224 = lean_ctor_get(x_1, 1); +lean_object* x_225; lean_object* x_226; lean_object* x_227; uint8_t x_228; +x_225 = lean_ctor_get(x_1, 0); +x_226 = lean_ctor_get(x_1, 1); lean_inc(x_2); -x_225 = l_Lean_IR_NormalizeIds_normFnBody(x_224, x_2, x_3); -x_226 = !lean_is_exclusive(x_225); -if (x_226 == 0) +x_227 = l_Lean_IR_NormalizeIds_normFnBody(x_226, x_2, x_3); +x_228 = !lean_is_exclusive(x_227); +if (x_228 == 0) { -lean_object* x_227; lean_object* x_228; -x_227 = lean_ctor_get(x_225, 0); -x_228 = l_Lean_IR_NormalizeIds_normIndex(x_223, x_2); +lean_object* x_229; lean_object* x_230; +x_229 = lean_ctor_get(x_227, 0); +x_230 = l_Lean_IR_NormalizeIds_normIndex(x_225, x_2); lean_dec(x_2); -lean_dec(x_223); -lean_ctor_set(x_1, 1, x_227); -lean_ctor_set(x_1, 0, x_228); -lean_ctor_set(x_225, 0, x_1); -return x_225; +lean_dec(x_225); +lean_ctor_set(x_1, 1, x_229); +lean_ctor_set(x_1, 0, x_230); +lean_ctor_set(x_227, 0, x_1); +return x_227; } else { -lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; -x_229 = lean_ctor_get(x_225, 0); -x_230 = lean_ctor_get(x_225, 1); -lean_inc(x_230); -lean_inc(x_229); -lean_dec(x_225); -x_231 = l_Lean_IR_NormalizeIds_normIndex(x_223, x_2); +lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; +x_231 = lean_ctor_get(x_227, 0); +x_232 = lean_ctor_get(x_227, 1); +lean_inc(x_232); +lean_inc(x_231); +lean_dec(x_227); +x_233 = l_Lean_IR_NormalizeIds_normIndex(x_225, x_2); lean_dec(x_2); -lean_dec(x_223); -lean_ctor_set(x_1, 1, x_229); -lean_ctor_set(x_1, 0, x_231); -x_232 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_232, 0, x_1); -lean_ctor_set(x_232, 1, x_230); -return x_232; +lean_dec(x_225); +lean_ctor_set(x_1, 1, x_231); +lean_ctor_set(x_1, 0, x_233); +x_234 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_234, 0, x_1); +lean_ctor_set(x_234, 1, x_232); +return x_234; } } else { -lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; -x_233 = lean_ctor_get(x_1, 0); -x_234 = lean_ctor_get(x_1, 1); -lean_inc(x_234); -lean_inc(x_233); +lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; +x_235 = lean_ctor_get(x_1, 0); +x_236 = lean_ctor_get(x_1, 1); +lean_inc(x_236); +lean_inc(x_235); lean_dec(x_1); lean_inc(x_2); -x_235 = l_Lean_IR_NormalizeIds_normFnBody(x_234, x_2, x_3); -x_236 = lean_ctor_get(x_235, 0); -lean_inc(x_236); -x_237 = lean_ctor_get(x_235, 1); -lean_inc(x_237); -if (lean_is_exclusive(x_235)) { - lean_ctor_release(x_235, 0); - lean_ctor_release(x_235, 1); - x_238 = x_235; +x_237 = l_Lean_IR_NormalizeIds_normFnBody(x_236, x_2, x_3); +x_238 = lean_ctor_get(x_237, 0); +lean_inc(x_238); +x_239 = lean_ctor_get(x_237, 1); +lean_inc(x_239); +if (lean_is_exclusive(x_237)) { + lean_ctor_release(x_237, 0); + lean_ctor_release(x_237, 1); + x_240 = x_237; } else { - lean_dec_ref(x_235); - x_238 = lean_box(0); + lean_dec_ref(x_237); + x_240 = lean_box(0); } -x_239 = l_Lean_IR_NormalizeIds_normIndex(x_233, x_2); +x_241 = l_Lean_IR_NormalizeIds_normIndex(x_235, x_2); lean_dec(x_2); -lean_dec(x_233); -x_240 = lean_alloc_ctor(8, 2, 0); -lean_ctor_set(x_240, 0, x_239); -lean_ctor_set(x_240, 1, x_236); -if (lean_is_scalar(x_238)) { - x_241 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_235); +x_242 = lean_alloc_ctor(8, 2, 0); +lean_ctor_set(x_242, 0, x_241); +lean_ctor_set(x_242, 1, x_238); +if (lean_is_scalar(x_240)) { + x_243 = lean_alloc_ctor(0, 2, 0); } else { - x_241 = x_238; + x_243 = x_240; } -lean_ctor_set(x_241, 0, x_240); -lean_ctor_set(x_241, 1, x_237); -return x_241; +lean_ctor_set(x_243, 0, x_242); +lean_ctor_set(x_243, 1, x_239); +return x_243; } } case 9: { -uint8_t x_242; -x_242 = !lean_is_exclusive(x_1); -if (x_242 == 0) +uint8_t x_244; +x_244 = !lean_is_exclusive(x_1); +if (x_244 == 0) { -lean_object* x_243; lean_object* x_244; size_t x_245; size_t x_246; lean_object* x_247; uint8_t x_248; -x_243 = lean_ctor_get(x_1, 1); -x_244 = lean_ctor_get(x_1, 3); -x_245 = lean_array_size(x_244); -x_246 = 0; +lean_object* x_245; lean_object* x_246; size_t x_247; size_t x_248; lean_object* x_249; uint8_t x_250; +x_245 = lean_ctor_get(x_1, 1); +x_246 = lean_ctor_get(x_1, 3); +x_247 = lean_array_size(x_246); +x_248 = 0; lean_inc(x_2); -x_247 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_NormalizeIds_normFnBody_spec__2(x_245, x_246, x_244, x_2, x_3); -x_248 = !lean_is_exclusive(x_247); -if (x_248 == 0) +x_249 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_NormalizeIds_normFnBody_spec__2(x_247, x_248, x_246, x_2, x_3); +x_250 = !lean_is_exclusive(x_249); +if (x_250 == 0) { -lean_object* x_249; lean_object* x_250; -x_249 = lean_ctor_get(x_247, 0); -x_250 = l_Lean_IR_NormalizeIds_normIndex(x_243, x_2); +lean_object* x_251; lean_object* x_252; +x_251 = lean_ctor_get(x_249, 0); +x_252 = l_Lean_IR_NormalizeIds_normIndex(x_245, x_2); lean_dec(x_2); -lean_dec(x_243); -lean_ctor_set(x_1, 3, x_249); -lean_ctor_set(x_1, 1, x_250); -lean_ctor_set(x_247, 0, x_1); -return x_247; +lean_dec(x_245); +lean_ctor_set(x_1, 3, x_251); +lean_ctor_set(x_1, 1, x_252); +lean_ctor_set(x_249, 0, x_1); +return x_249; } else { -lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; -x_251 = lean_ctor_get(x_247, 0); -x_252 = lean_ctor_get(x_247, 1); -lean_inc(x_252); -lean_inc(x_251); -lean_dec(x_247); -x_253 = l_Lean_IR_NormalizeIds_normIndex(x_243, x_2); +lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; +x_253 = lean_ctor_get(x_249, 0); +x_254 = lean_ctor_get(x_249, 1); +lean_inc(x_254); +lean_inc(x_253); +lean_dec(x_249); +x_255 = l_Lean_IR_NormalizeIds_normIndex(x_245, x_2); lean_dec(x_2); -lean_dec(x_243); -lean_ctor_set(x_1, 3, x_251); -lean_ctor_set(x_1, 1, x_253); -x_254 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_254, 0, x_1); -lean_ctor_set(x_254, 1, x_252); -return x_254; +lean_dec(x_245); +lean_ctor_set(x_1, 3, x_253); +lean_ctor_set(x_1, 1, x_255); +x_256 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_256, 0, x_1); +lean_ctor_set(x_256, 1, x_254); +return x_256; } } else { -lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; size_t x_259; size_t x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; -x_255 = lean_ctor_get(x_1, 0); -x_256 = lean_ctor_get(x_1, 1); -x_257 = lean_ctor_get(x_1, 2); -x_258 = lean_ctor_get(x_1, 3); +lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; size_t x_261; size_t x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; +x_257 = lean_ctor_get(x_1, 0); +x_258 = lean_ctor_get(x_1, 1); +x_259 = lean_ctor_get(x_1, 2); +x_260 = lean_ctor_get(x_1, 3); +lean_inc(x_260); +lean_inc(x_259); lean_inc(x_258); lean_inc(x_257); -lean_inc(x_256); -lean_inc(x_255); lean_dec(x_1); -x_259 = lean_array_size(x_258); -x_260 = 0; +x_261 = lean_array_size(x_260); +x_262 = 0; lean_inc(x_2); -x_261 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_NormalizeIds_normFnBody_spec__2(x_259, x_260, x_258, x_2, x_3); -x_262 = lean_ctor_get(x_261, 0); -lean_inc(x_262); -x_263 = lean_ctor_get(x_261, 1); -lean_inc(x_263); -if (lean_is_exclusive(x_261)) { - lean_ctor_release(x_261, 0); - lean_ctor_release(x_261, 1); - x_264 = x_261; +x_263 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_NormalizeIds_normFnBody_spec__2(x_261, x_262, x_260, x_2, x_3); +x_264 = lean_ctor_get(x_263, 0); +lean_inc(x_264); +x_265 = lean_ctor_get(x_263, 1); +lean_inc(x_265); +if (lean_is_exclusive(x_263)) { + lean_ctor_release(x_263, 0); + lean_ctor_release(x_263, 1); + x_266 = x_263; } else { - lean_dec_ref(x_261); - x_264 = lean_box(0); + lean_dec_ref(x_263); + x_266 = lean_box(0); } -x_265 = l_Lean_IR_NormalizeIds_normIndex(x_256, x_2); +x_267 = l_Lean_IR_NormalizeIds_normIndex(x_258, x_2); lean_dec(x_2); -lean_dec(x_256); -x_266 = lean_alloc_ctor(9, 4, 0); -lean_ctor_set(x_266, 0, x_255); -lean_ctor_set(x_266, 1, x_265); -lean_ctor_set(x_266, 2, x_257); -lean_ctor_set(x_266, 3, x_262); -if (lean_is_scalar(x_264)) { - x_267 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_258); +x_268 = lean_alloc_ctor(9, 4, 0); +lean_ctor_set(x_268, 0, x_257); +lean_ctor_set(x_268, 1, x_267); +lean_ctor_set(x_268, 2, x_259); +lean_ctor_set(x_268, 3, x_264); +if (lean_is_scalar(x_266)) { + x_269 = lean_alloc_ctor(0, 2, 0); } else { - x_267 = x_264; + x_269 = x_266; } -lean_ctor_set(x_267, 0, x_266); -lean_ctor_set(x_267, 1, x_263); -return x_267; +lean_ctor_set(x_269, 0, x_268); +lean_ctor_set(x_269, 1, x_265); +return x_269; } } case 10: { -uint8_t x_268; -x_268 = !lean_is_exclusive(x_1); -if (x_268 == 0) +uint8_t x_270; +x_270 = !lean_is_exclusive(x_1); +if (x_270 == 0) { -lean_object* x_269; lean_object* x_270; lean_object* x_271; -x_269 = lean_ctor_get(x_1, 0); -x_270 = l_Lean_IR_NormalizeIds_normArg(x_269, x_2); +lean_object* x_271; lean_object* x_272; lean_object* x_273; +x_271 = lean_ctor_get(x_1, 0); +x_272 = l_Lean_IR_NormalizeIds_normArg(x_271, x_2); lean_dec(x_2); -lean_ctor_set(x_1, 0, x_270); -x_271 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_271, 0, x_1); -lean_ctor_set(x_271, 1, x_3); -return x_271; +lean_ctor_set(x_1, 0, x_272); +x_273 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_273, 0, x_1); +lean_ctor_set(x_273, 1, x_3); +return x_273; } else { -lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; -x_272 = lean_ctor_get(x_1, 0); -lean_inc(x_272); +lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; +x_274 = lean_ctor_get(x_1, 0); +lean_inc(x_274); lean_dec(x_1); -x_273 = l_Lean_IR_NormalizeIds_normArg(x_272, x_2); +x_275 = l_Lean_IR_NormalizeIds_normArg(x_274, x_2); lean_dec(x_2); -x_274 = lean_alloc_ctor(10, 1, 0); -lean_ctor_set(x_274, 0, x_273); -x_275 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_275, 0, x_274); -lean_ctor_set(x_275, 1, x_3); -return x_275; +x_276 = lean_alloc_ctor(10, 1, 0); +lean_ctor_set(x_276, 0, x_275); +x_277 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_277, 0, x_276); +lean_ctor_set(x_277, 1, x_3); +return x_277; } } case 11: { -uint8_t x_276; -x_276 = !lean_is_exclusive(x_1); -if (x_276 == 0) +uint8_t x_278; +x_278 = !lean_is_exclusive(x_1); +if (x_278 == 0) { -lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; -x_277 = lean_ctor_get(x_1, 0); -x_278 = lean_ctor_get(x_1, 1); -x_279 = l_Lean_IR_NormalizeIds_normIndex(x_277, x_2); -lean_dec(x_277); -x_280 = l_Lean_IR_NormalizeIds_normArgs(x_278, x_2); +lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; +x_279 = lean_ctor_get(x_1, 0); +x_280 = lean_ctor_get(x_1, 1); +x_281 = l_Lean_IR_NormalizeIds_normIndex(x_279, x_2); +lean_dec(x_279); +x_282 = l_Lean_IR_NormalizeIds_normArgs(x_280, x_2); lean_dec(x_2); -lean_ctor_set(x_1, 1, x_280); -lean_ctor_set(x_1, 0, x_279); -x_281 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_281, 0, x_1); -lean_ctor_set(x_281, 1, x_3); -return x_281; +lean_ctor_set(x_1, 1, x_282); +lean_ctor_set(x_1, 0, x_281); +x_283 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_283, 0, x_1); +lean_ctor_set(x_283, 1, x_3); +return x_283; } else { -lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; -x_282 = lean_ctor_get(x_1, 0); -x_283 = lean_ctor_get(x_1, 1); -lean_inc(x_283); -lean_inc(x_282); +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_284 = lean_ctor_get(x_1, 0); +x_285 = lean_ctor_get(x_1, 1); +lean_inc(x_285); +lean_inc(x_284); lean_dec(x_1); -x_284 = l_Lean_IR_NormalizeIds_normIndex(x_282, x_2); -lean_dec(x_282); -x_285 = l_Lean_IR_NormalizeIds_normArgs(x_283, x_2); +x_286 = l_Lean_IR_NormalizeIds_normIndex(x_284, x_2); +lean_dec(x_284); +x_287 = l_Lean_IR_NormalizeIds_normArgs(x_285, x_2); lean_dec(x_2); -x_286 = lean_alloc_ctor(11, 2, 0); -lean_ctor_set(x_286, 0, x_284); -lean_ctor_set(x_286, 1, x_285); -x_287 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_287, 0, x_286); -lean_ctor_set(x_287, 1, x_3); -return x_287; +x_288 = lean_alloc_ctor(11, 2, 0); +lean_ctor_set(x_288, 0, x_286); +lean_ctor_set(x_288, 1, x_287); +x_289 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_289, 0, x_288); +lean_ctor_set(x_289, 1, x_3); +return x_289; } } default: { -lean_object* x_288; +lean_object* x_290; lean_dec(x_2); -x_288 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_288, 0, x_1); -lean_ctor_set(x_288, 1, x_3); -return x_288; +x_290 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_290, 0, x_1); +lean_ctor_set(x_290, 1, x_3); +return x_290; } } } @@ -4615,219 +4630,228 @@ x_29 = !lean_is_exclusive(x_2); if (x_29 == 0) { lean_object* x_30; lean_object* x_31; -x_30 = lean_ctor_get(x_2, 1); +x_30 = lean_ctor_get(x_2, 2); x_31 = lean_apply_1(x_1, x_30); -lean_ctor_set(x_2, 1, x_31); +lean_ctor_set(x_2, 2, x_31); return x_2; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; x_32 = lean_ctor_get(x_2, 0); x_33 = lean_ctor_get(x_2, 1); +x_34 = lean_ctor_get(x_2, 2); +lean_inc(x_34); lean_inc(x_33); lean_inc(x_32); lean_dec(x_2); -x_34 = lean_apply_1(x_1, x_33); -x_35 = lean_alloc_ctor(3, 2, 0); -lean_ctor_set(x_35, 0, x_32); -lean_ctor_set(x_35, 1, x_34); -return x_35; +x_35 = lean_apply_1(x_1, x_34); +x_36 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_36, 0, x_32); +lean_ctor_set(x_36, 1, x_33); +lean_ctor_set(x_36, 2, x_35); +return x_36; } } case 4: { -uint8_t x_36; -x_36 = !lean_is_exclusive(x_2); -if (x_36 == 0) +uint8_t x_37; +x_37 = !lean_is_exclusive(x_2); +if (x_37 == 0) { -lean_object* x_37; lean_object* x_38; -x_37 = lean_ctor_get(x_2, 1); -x_38 = lean_apply_1(x_1, x_37); -lean_ctor_set(x_2, 1, x_38); +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_2, 2); +x_39 = lean_apply_1(x_1, x_38); +lean_ctor_set(x_2, 2, x_39); return x_2; } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_39 = lean_ctor_get(x_2, 0); -x_40 = lean_ctor_get(x_2, 1); +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_40 = lean_ctor_get(x_2, 0); +x_41 = lean_ctor_get(x_2, 1); +x_42 = lean_ctor_get(x_2, 2); +lean_inc(x_42); +lean_inc(x_41); lean_inc(x_40); -lean_inc(x_39); lean_dec(x_2); -x_41 = lean_apply_1(x_1, x_40); -x_42 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_42, 0, x_39); -lean_ctor_set(x_42, 1, x_41); -return x_42; +x_43 = lean_apply_1(x_1, x_42); +x_44 = lean_alloc_ctor(4, 3, 0); +lean_ctor_set(x_44, 0, x_40); +lean_ctor_set(x_44, 1, x_41); +lean_ctor_set(x_44, 2, x_43); +return x_44; } } case 5: { -uint8_t x_43; -x_43 = !lean_is_exclusive(x_2); -if (x_43 == 0) +uint8_t x_45; +x_45 = !lean_is_exclusive(x_2); +if (x_45 == 0) { -lean_object* x_44; lean_object* x_45; -x_44 = lean_ctor_get(x_2, 2); -x_45 = lean_apply_1(x_1, x_44); -lean_ctor_set(x_2, 2, x_45); +lean_object* x_46; lean_object* x_47; +x_46 = lean_ctor_get(x_2, 3); +x_47 = lean_apply_1(x_1, x_46); +lean_ctor_set(x_2, 3, x_47); return x_2; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_46 = lean_ctor_get(x_2, 0); -x_47 = lean_ctor_get(x_2, 1); -x_48 = lean_ctor_get(x_2, 2); +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_48 = lean_ctor_get(x_2, 0); +x_49 = lean_ctor_get(x_2, 1); +x_50 = lean_ctor_get(x_2, 2); +x_51 = lean_ctor_get(x_2, 3); +lean_inc(x_51); +lean_inc(x_50); +lean_inc(x_49); lean_inc(x_48); -lean_inc(x_47); -lean_inc(x_46); lean_dec(x_2); -x_49 = lean_apply_1(x_1, x_48); -x_50 = lean_alloc_ctor(5, 3, 0); -lean_ctor_set(x_50, 0, x_46); -lean_ctor_set(x_50, 1, x_47); -lean_ctor_set(x_50, 2, x_49); -return x_50; +x_52 = lean_apply_1(x_1, x_51); +x_53 = lean_alloc_ctor(5, 4, 0); +lean_ctor_set(x_53, 0, x_48); +lean_ctor_set(x_53, 1, x_49); +lean_ctor_set(x_53, 2, x_50); +lean_ctor_set(x_53, 3, x_52); +return x_53; } } case 6: { -uint8_t x_51; -x_51 = !lean_is_exclusive(x_2); -if (x_51 == 0) +uint8_t x_54; +x_54 = !lean_is_exclusive(x_2); +if (x_54 == 0) { -lean_object* x_52; lean_object* x_53; -x_52 = lean_ctor_get(x_2, 1); -x_53 = l_Lean_IR_MapVars_mapArgs(x_1, x_52); -lean_ctor_set(x_2, 1, x_53); +lean_object* x_55; lean_object* x_56; +x_55 = lean_ctor_get(x_2, 1); +x_56 = l_Lean_IR_MapVars_mapArgs(x_1, x_55); +lean_ctor_set(x_2, 1, x_56); return x_2; } else { -lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_54 = lean_ctor_get(x_2, 0); -x_55 = lean_ctor_get(x_2, 1); -lean_inc(x_55); -lean_inc(x_54); +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_57 = lean_ctor_get(x_2, 0); +x_58 = lean_ctor_get(x_2, 1); +lean_inc(x_58); +lean_inc(x_57); lean_dec(x_2); -x_56 = l_Lean_IR_MapVars_mapArgs(x_1, x_55); -x_57 = lean_alloc_ctor(6, 2, 0); -lean_ctor_set(x_57, 0, x_54); -lean_ctor_set(x_57, 1, x_56); -return x_57; +x_59 = l_Lean_IR_MapVars_mapArgs(x_1, x_58); +x_60 = lean_alloc_ctor(6, 2, 0); +lean_ctor_set(x_60, 0, x_57); +lean_ctor_set(x_60, 1, x_59); +return x_60; } } case 7: { -uint8_t x_58; -x_58 = !lean_is_exclusive(x_2); -if (x_58 == 0) +uint8_t x_61; +x_61 = !lean_is_exclusive(x_2); +if (x_61 == 0) { -lean_object* x_59; lean_object* x_60; -x_59 = lean_ctor_get(x_2, 1); -x_60 = l_Lean_IR_MapVars_mapArgs(x_1, x_59); -lean_ctor_set(x_2, 1, x_60); +lean_object* x_62; lean_object* x_63; +x_62 = lean_ctor_get(x_2, 1); +x_63 = l_Lean_IR_MapVars_mapArgs(x_1, x_62); +lean_ctor_set(x_2, 1, x_63); return x_2; } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_61 = lean_ctor_get(x_2, 0); -x_62 = lean_ctor_get(x_2, 1); -lean_inc(x_62); -lean_inc(x_61); +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_64 = lean_ctor_get(x_2, 0); +x_65 = lean_ctor_get(x_2, 1); +lean_inc(x_65); +lean_inc(x_64); lean_dec(x_2); -x_63 = l_Lean_IR_MapVars_mapArgs(x_1, x_62); -x_64 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_64, 0, x_61); -lean_ctor_set(x_64, 1, x_63); -return x_64; +x_66 = l_Lean_IR_MapVars_mapArgs(x_1, x_65); +x_67 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_67, 0, x_64); +lean_ctor_set(x_67, 1, x_66); +return x_67; } } case 8: { -uint8_t x_65; -x_65 = !lean_is_exclusive(x_2); -if (x_65 == 0) +uint8_t x_68; +x_68 = !lean_is_exclusive(x_2); +if (x_68 == 0) { -lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_66 = lean_ctor_get(x_2, 0); -x_67 = lean_ctor_get(x_2, 1); +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_69 = lean_ctor_get(x_2, 0); +x_70 = lean_ctor_get(x_2, 1); lean_inc_ref(x_1); -x_68 = lean_apply_1(x_1, x_66); -x_69 = l_Lean_IR_MapVars_mapArgs(x_1, x_67); -lean_ctor_set(x_2, 1, x_69); -lean_ctor_set(x_2, 0, x_68); +x_71 = lean_apply_1(x_1, x_69); +x_72 = l_Lean_IR_MapVars_mapArgs(x_1, x_70); +lean_ctor_set(x_2, 1, x_72); +lean_ctor_set(x_2, 0, x_71); return x_2; } else { -lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_70 = lean_ctor_get(x_2, 0); -x_71 = lean_ctor_get(x_2, 1); -lean_inc(x_71); -lean_inc(x_70); +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_73 = lean_ctor_get(x_2, 0); +x_74 = lean_ctor_get(x_2, 1); +lean_inc(x_74); +lean_inc(x_73); lean_dec(x_2); lean_inc_ref(x_1); -x_72 = lean_apply_1(x_1, x_70); -x_73 = l_Lean_IR_MapVars_mapArgs(x_1, x_71); -x_74 = lean_alloc_ctor(8, 2, 0); -lean_ctor_set(x_74, 0, x_72); -lean_ctor_set(x_74, 1, x_73); -return x_74; +x_75 = lean_apply_1(x_1, x_73); +x_76 = l_Lean_IR_MapVars_mapArgs(x_1, x_74); +x_77 = lean_alloc_ctor(8, 2, 0); +lean_ctor_set(x_77, 0, x_75); +lean_ctor_set(x_77, 1, x_76); +return x_77; } } case 9: { -uint8_t x_75; -x_75 = !lean_is_exclusive(x_2); -if (x_75 == 0) +uint8_t x_78; +x_78 = !lean_is_exclusive(x_2); +if (x_78 == 0) { -lean_object* x_76; lean_object* x_77; -x_76 = lean_ctor_get(x_2, 1); -x_77 = lean_apply_1(x_1, x_76); -lean_ctor_set(x_2, 1, x_77); +lean_object* x_79; lean_object* x_80; +x_79 = lean_ctor_get(x_2, 1); +x_80 = lean_apply_1(x_1, x_79); +lean_ctor_set(x_2, 1, x_80); return x_2; } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_78 = lean_ctor_get(x_2, 0); -x_79 = lean_ctor_get(x_2, 1); -lean_inc(x_79); -lean_inc(x_78); +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_81 = lean_ctor_get(x_2, 0); +x_82 = lean_ctor_get(x_2, 1); +lean_inc(x_82); +lean_inc(x_81); lean_dec(x_2); -x_80 = lean_apply_1(x_1, x_79); -x_81 = lean_alloc_ctor(9, 2, 0); -lean_ctor_set(x_81, 0, x_78); -lean_ctor_set(x_81, 1, x_80); -return x_81; +x_83 = lean_apply_1(x_1, x_82); +x_84 = lean_alloc_ctor(9, 2, 0); +lean_ctor_set(x_84, 0, x_81); +lean_ctor_set(x_84, 1, x_83); +return x_84; } } case 10: { -uint8_t x_82; -x_82 = !lean_is_exclusive(x_2); -if (x_82 == 0) +uint8_t x_85; +x_85 = !lean_is_exclusive(x_2); +if (x_85 == 0) { -lean_object* x_83; lean_object* x_84; -x_83 = lean_ctor_get(x_2, 0); -x_84 = lean_apply_1(x_1, x_83); -lean_ctor_set(x_2, 0, x_84); +lean_object* x_86; lean_object* x_87; +x_86 = lean_ctor_get(x_2, 0); +x_87 = lean_apply_1(x_1, x_86); +lean_ctor_set(x_2, 0, x_87); return x_2; } else { -lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_85 = lean_ctor_get(x_2, 0); -lean_inc(x_85); +lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_88 = lean_ctor_get(x_2, 0); +lean_inc(x_88); lean_dec(x_2); -x_86 = lean_apply_1(x_1, x_85); -x_87 = lean_alloc_ctor(10, 1, 0); -lean_ctor_set(x_87, 0, x_86); -return x_87; +x_89 = lean_apply_1(x_1, x_88); +x_90 = lean_alloc_ctor(10, 1, 0); +lean_ctor_set(x_90, 0, x_89); +return x_90; } } case 11: @@ -4837,26 +4861,26 @@ return x_2; } default: { -uint8_t x_88; -x_88 = !lean_is_exclusive(x_2); -if (x_88 == 0) +uint8_t x_91; +x_91 = !lean_is_exclusive(x_2); +if (x_91 == 0) { -lean_object* x_89; lean_object* x_90; -x_89 = lean_ctor_get(x_2, 0); -x_90 = lean_apply_1(x_1, x_89); -lean_ctor_set(x_2, 0, x_90); +lean_object* x_92; lean_object* x_93; +x_92 = lean_ctor_get(x_2, 0); +x_93 = lean_apply_1(x_1, x_92); +lean_ctor_set(x_2, 0, x_93); return x_2; } else { -lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_91 = lean_ctor_get(x_2, 0); -lean_inc(x_91); +lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_94 = lean_ctor_get(x_2, 0); +lean_inc(x_94); lean_dec(x_2); -x_92 = lean_apply_1(x_1, x_91); -x_93 = lean_alloc_ctor(12, 1, 0); -lean_ctor_set(x_93, 0, x_92); -return x_93; +x_95 = lean_apply_1(x_1, x_94); +x_96 = lean_alloc_ctor(12, 1, 0); +lean_ctor_set(x_96, 0, x_95); +return x_96; } } } @@ -5060,309 +5084,315 @@ if (x_54 == 0) { lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; x_55 = lean_ctor_get(x_2, 0); -x_56 = lean_ctor_get(x_2, 2); -x_57 = lean_ctor_get(x_2, 3); +x_56 = lean_ctor_get(x_2, 3); +x_57 = lean_ctor_get(x_2, 4); lean_inc_ref(x_1); x_58 = lean_apply_1(x_1, x_55); lean_inc_ref(x_1); x_59 = lean_apply_1(x_1, x_56); x_60 = l_Lean_IR_MapVars_mapFnBody(x_1, x_57); -lean_ctor_set(x_2, 3, x_60); -lean_ctor_set(x_2, 2, x_59); +lean_ctor_set(x_2, 4, x_60); +lean_ctor_set(x_2, 3, x_59); lean_ctor_set(x_2, 0, x_58); return x_2; } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; x_61 = lean_ctor_get(x_2, 0); x_62 = lean_ctor_get(x_2, 1); x_63 = lean_ctor_get(x_2, 2); x_64 = lean_ctor_get(x_2, 3); +x_65 = lean_ctor_get(x_2, 4); +lean_inc(x_65); lean_inc(x_64); lean_inc(x_63); lean_inc(x_62); lean_inc(x_61); lean_dec(x_2); lean_inc_ref(x_1); -x_65 = lean_apply_1(x_1, x_61); +x_66 = lean_apply_1(x_1, x_61); lean_inc_ref(x_1); -x_66 = lean_apply_1(x_1, x_63); -x_67 = l_Lean_IR_MapVars_mapFnBody(x_1, x_64); -x_68 = lean_alloc_ctor(4, 4, 0); -lean_ctor_set(x_68, 0, x_65); -lean_ctor_set(x_68, 1, x_62); -lean_ctor_set(x_68, 2, x_66); -lean_ctor_set(x_68, 3, x_67); -return x_68; +x_67 = lean_apply_1(x_1, x_64); +x_68 = l_Lean_IR_MapVars_mapFnBody(x_1, x_65); +x_69 = lean_alloc_ctor(4, 5, 0); +lean_ctor_set(x_69, 0, x_66); +lean_ctor_set(x_69, 1, x_62); +lean_ctor_set(x_69, 2, x_63); +lean_ctor_set(x_69, 3, x_67); +lean_ctor_set(x_69, 4, x_68); +return x_69; } } case 5: { -uint8_t x_69; -x_69 = !lean_is_exclusive(x_2); -if (x_69 == 0) +uint8_t x_70; +x_70 = !lean_is_exclusive(x_2); +if (x_70 == 0) { -lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_70 = lean_ctor_get(x_2, 0); -x_71 = lean_ctor_get(x_2, 3); -x_72 = lean_ctor_get(x_2, 5); -lean_inc_ref(x_1); -x_73 = lean_apply_1(x_1, x_70); +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_71 = lean_ctor_get(x_2, 0); +x_72 = lean_ctor_get(x_2, 4); +x_73 = lean_ctor_get(x_2, 6); lean_inc_ref(x_1); x_74 = lean_apply_1(x_1, x_71); -x_75 = l_Lean_IR_MapVars_mapFnBody(x_1, x_72); -lean_ctor_set(x_2, 5, x_75); -lean_ctor_set(x_2, 3, x_74); -lean_ctor_set(x_2, 0, x_73); +lean_inc_ref(x_1); +x_75 = lean_apply_1(x_1, x_72); +x_76 = l_Lean_IR_MapVars_mapFnBody(x_1, x_73); +lean_ctor_set(x_2, 6, x_76); +lean_ctor_set(x_2, 4, x_75); +lean_ctor_set(x_2, 0, x_74); return x_2; } else { -lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_76 = lean_ctor_get(x_2, 0); -x_77 = lean_ctor_get(x_2, 1); -x_78 = lean_ctor_get(x_2, 2); -x_79 = lean_ctor_get(x_2, 3); -x_80 = lean_ctor_get(x_2, 4); -x_81 = lean_ctor_get(x_2, 5); +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_77 = lean_ctor_get(x_2, 0); +x_78 = lean_ctor_get(x_2, 1); +x_79 = lean_ctor_get(x_2, 2); +x_80 = lean_ctor_get(x_2, 3); +x_81 = lean_ctor_get(x_2, 4); +x_82 = lean_ctor_get(x_2, 5); +x_83 = lean_ctor_get(x_2, 6); +lean_inc(x_83); +lean_inc(x_82); lean_inc(x_81); lean_inc(x_80); lean_inc(x_79); lean_inc(x_78); lean_inc(x_77); -lean_inc(x_76); lean_dec(x_2); lean_inc_ref(x_1); -x_82 = lean_apply_1(x_1, x_76); +x_84 = lean_apply_1(x_1, x_77); lean_inc_ref(x_1); -x_83 = lean_apply_1(x_1, x_79); -x_84 = l_Lean_IR_MapVars_mapFnBody(x_1, x_81); -x_85 = lean_alloc_ctor(5, 6, 0); -lean_ctor_set(x_85, 0, x_82); -lean_ctor_set(x_85, 1, x_77); -lean_ctor_set(x_85, 2, x_78); -lean_ctor_set(x_85, 3, x_83); -lean_ctor_set(x_85, 4, x_80); -lean_ctor_set(x_85, 5, x_84); -return x_85; +x_85 = lean_apply_1(x_1, x_81); +x_86 = l_Lean_IR_MapVars_mapFnBody(x_1, x_83); +x_87 = lean_alloc_ctor(5, 7, 0); +lean_ctor_set(x_87, 0, x_84); +lean_ctor_set(x_87, 1, x_78); +lean_ctor_set(x_87, 2, x_79); +lean_ctor_set(x_87, 3, x_80); +lean_ctor_set(x_87, 4, x_85); +lean_ctor_set(x_87, 5, x_82); +lean_ctor_set(x_87, 6, x_86); +return x_87; } } case 6: { -uint8_t x_86; -x_86 = !lean_is_exclusive(x_2); -if (x_86 == 0) +uint8_t x_88; +x_88 = !lean_is_exclusive(x_2); +if (x_88 == 0) { -lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; -x_87 = lean_ctor_get(x_2, 0); -x_88 = lean_ctor_get(x_2, 2); +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_89 = lean_ctor_get(x_2, 0); +x_90 = lean_ctor_get(x_2, 2); lean_inc_ref(x_1); -x_89 = lean_apply_1(x_1, x_87); -x_90 = l_Lean_IR_MapVars_mapFnBody(x_1, x_88); -lean_ctor_set(x_2, 2, x_90); -lean_ctor_set(x_2, 0, x_89); +x_91 = lean_apply_1(x_1, x_89); +x_92 = l_Lean_IR_MapVars_mapFnBody(x_1, x_90); +lean_ctor_set(x_2, 2, x_92); +lean_ctor_set(x_2, 0, x_91); return x_2; } else { -lean_object* x_91; lean_object* x_92; uint8_t x_93; uint8_t x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; -x_91 = lean_ctor_get(x_2, 0); -x_92 = lean_ctor_get(x_2, 1); -x_93 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); -x_94 = lean_ctor_get_uint8(x_2, sizeof(void*)*3 + 1); -x_95 = lean_ctor_get(x_2, 2); -lean_inc(x_95); -lean_inc(x_92); -lean_inc(x_91); +lean_object* x_93; lean_object* x_94; uint8_t x_95; uint8_t x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; +x_93 = lean_ctor_get(x_2, 0); +x_94 = lean_ctor_get(x_2, 1); +x_95 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); +x_96 = lean_ctor_get_uint8(x_2, sizeof(void*)*3 + 1); +x_97 = lean_ctor_get(x_2, 2); +lean_inc(x_97); +lean_inc(x_94); +lean_inc(x_93); lean_dec(x_2); lean_inc_ref(x_1); -x_96 = lean_apply_1(x_1, x_91); -x_97 = l_Lean_IR_MapVars_mapFnBody(x_1, x_95); -x_98 = lean_alloc_ctor(6, 3, 2); -lean_ctor_set(x_98, 0, x_96); -lean_ctor_set(x_98, 1, x_92); -lean_ctor_set(x_98, 2, x_97); -lean_ctor_set_uint8(x_98, sizeof(void*)*3, x_93); -lean_ctor_set_uint8(x_98, sizeof(void*)*3 + 1, x_94); -return x_98; +x_98 = lean_apply_1(x_1, x_93); +x_99 = l_Lean_IR_MapVars_mapFnBody(x_1, x_97); +x_100 = lean_alloc_ctor(6, 3, 2); +lean_ctor_set(x_100, 0, x_98); +lean_ctor_set(x_100, 1, x_94); +lean_ctor_set(x_100, 2, x_99); +lean_ctor_set_uint8(x_100, sizeof(void*)*3, x_95); +lean_ctor_set_uint8(x_100, sizeof(void*)*3 + 1, x_96); +return x_100; } } case 7: { -uint8_t x_99; -x_99 = !lean_is_exclusive(x_2); -if (x_99 == 0) +uint8_t x_101; +x_101 = !lean_is_exclusive(x_2); +if (x_101 == 0) { -lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; -x_100 = lean_ctor_get(x_2, 0); -x_101 = lean_ctor_get(x_2, 2); +lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; +x_102 = lean_ctor_get(x_2, 0); +x_103 = lean_ctor_get(x_2, 2); lean_inc_ref(x_1); -x_102 = lean_apply_1(x_1, x_100); -x_103 = l_Lean_IR_MapVars_mapFnBody(x_1, x_101); -lean_ctor_set(x_2, 2, x_103); -lean_ctor_set(x_2, 0, x_102); +x_104 = lean_apply_1(x_1, x_102); +x_105 = l_Lean_IR_MapVars_mapFnBody(x_1, x_103); +lean_ctor_set(x_2, 2, x_105); +lean_ctor_set(x_2, 0, x_104); return x_2; } else { -lean_object* x_104; lean_object* x_105; uint8_t x_106; uint8_t x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; -x_104 = lean_ctor_get(x_2, 0); -x_105 = lean_ctor_get(x_2, 1); -x_106 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); -x_107 = lean_ctor_get_uint8(x_2, sizeof(void*)*3 + 1); -x_108 = lean_ctor_get(x_2, 2); -lean_inc(x_108); -lean_inc(x_105); -lean_inc(x_104); +lean_object* x_106; lean_object* x_107; uint8_t x_108; uint8_t x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; +x_106 = lean_ctor_get(x_2, 0); +x_107 = lean_ctor_get(x_2, 1); +x_108 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); +x_109 = lean_ctor_get_uint8(x_2, sizeof(void*)*3 + 1); +x_110 = lean_ctor_get(x_2, 2); +lean_inc(x_110); +lean_inc(x_107); +lean_inc(x_106); lean_dec(x_2); lean_inc_ref(x_1); -x_109 = lean_apply_1(x_1, x_104); -x_110 = l_Lean_IR_MapVars_mapFnBody(x_1, x_108); -x_111 = lean_alloc_ctor(7, 3, 2); -lean_ctor_set(x_111, 0, x_109); -lean_ctor_set(x_111, 1, x_105); -lean_ctor_set(x_111, 2, x_110); -lean_ctor_set_uint8(x_111, sizeof(void*)*3, x_106); -lean_ctor_set_uint8(x_111, sizeof(void*)*3 + 1, x_107); -return x_111; +x_111 = lean_apply_1(x_1, x_106); +x_112 = l_Lean_IR_MapVars_mapFnBody(x_1, x_110); +x_113 = lean_alloc_ctor(7, 3, 2); +lean_ctor_set(x_113, 0, x_111); +lean_ctor_set(x_113, 1, x_107); +lean_ctor_set(x_113, 2, x_112); +lean_ctor_set_uint8(x_113, sizeof(void*)*3, x_108); +lean_ctor_set_uint8(x_113, sizeof(void*)*3 + 1, x_109); +return x_113; } } case 8: { -uint8_t x_112; -x_112 = !lean_is_exclusive(x_2); -if (x_112 == 0) +uint8_t x_114; +x_114 = !lean_is_exclusive(x_2); +if (x_114 == 0) { -lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; -x_113 = lean_ctor_get(x_2, 0); -x_114 = lean_ctor_get(x_2, 1); +lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; +x_115 = lean_ctor_get(x_2, 0); +x_116 = lean_ctor_get(x_2, 1); lean_inc_ref(x_1); -x_115 = lean_apply_1(x_1, x_113); -x_116 = l_Lean_IR_MapVars_mapFnBody(x_1, x_114); -lean_ctor_set(x_2, 1, x_116); -lean_ctor_set(x_2, 0, x_115); +x_117 = lean_apply_1(x_1, x_115); +x_118 = l_Lean_IR_MapVars_mapFnBody(x_1, x_116); +lean_ctor_set(x_2, 1, x_118); +lean_ctor_set(x_2, 0, x_117); return x_2; } else { -lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; -x_117 = lean_ctor_get(x_2, 0); -x_118 = lean_ctor_get(x_2, 1); -lean_inc(x_118); -lean_inc(x_117); +lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_119 = lean_ctor_get(x_2, 0); +x_120 = lean_ctor_get(x_2, 1); +lean_inc(x_120); +lean_inc(x_119); lean_dec(x_2); lean_inc_ref(x_1); -x_119 = lean_apply_1(x_1, x_117); -x_120 = l_Lean_IR_MapVars_mapFnBody(x_1, x_118); -x_121 = lean_alloc_ctor(8, 2, 0); -lean_ctor_set(x_121, 0, x_119); -lean_ctor_set(x_121, 1, x_120); -return x_121; +x_121 = lean_apply_1(x_1, x_119); +x_122 = l_Lean_IR_MapVars_mapFnBody(x_1, x_120); +x_123 = lean_alloc_ctor(8, 2, 0); +lean_ctor_set(x_123, 0, x_121); +lean_ctor_set(x_123, 1, x_122); +return x_123; } } case 9: { -uint8_t x_122; -x_122 = !lean_is_exclusive(x_2); -if (x_122 == 0) +uint8_t x_124; +x_124 = !lean_is_exclusive(x_2); +if (x_124 == 0) { -lean_object* x_123; lean_object* x_124; lean_object* x_125; size_t x_126; size_t x_127; lean_object* x_128; -x_123 = lean_ctor_get(x_2, 1); -x_124 = lean_ctor_get(x_2, 3); +lean_object* x_125; lean_object* x_126; lean_object* x_127; size_t x_128; size_t x_129; lean_object* x_130; +x_125 = lean_ctor_get(x_2, 1); +x_126 = lean_ctor_get(x_2, 3); lean_inc_ref(x_1); -x_125 = lean_apply_1(x_1, x_123); -x_126 = lean_array_size(x_124); -x_127 = 0; -x_128 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_MapVars_mapFnBody_spec__0(x_1, x_126, x_127, x_124); -lean_ctor_set(x_2, 3, x_128); -lean_ctor_set(x_2, 1, x_125); +x_127 = lean_apply_1(x_1, x_125); +x_128 = lean_array_size(x_126); +x_129 = 0; +x_130 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_MapVars_mapFnBody_spec__0(x_1, x_128, x_129, x_126); +lean_ctor_set(x_2, 3, x_130); +lean_ctor_set(x_2, 1, x_127); return x_2; } else { -lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; size_t x_134; size_t x_135; lean_object* x_136; lean_object* x_137; -x_129 = lean_ctor_get(x_2, 0); -x_130 = lean_ctor_get(x_2, 1); -x_131 = lean_ctor_get(x_2, 2); -x_132 = lean_ctor_get(x_2, 3); +lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; size_t x_136; size_t x_137; lean_object* x_138; lean_object* x_139; +x_131 = lean_ctor_get(x_2, 0); +x_132 = lean_ctor_get(x_2, 1); +x_133 = lean_ctor_get(x_2, 2); +x_134 = lean_ctor_get(x_2, 3); +lean_inc(x_134); +lean_inc(x_133); lean_inc(x_132); lean_inc(x_131); -lean_inc(x_130); -lean_inc(x_129); lean_dec(x_2); lean_inc_ref(x_1); -x_133 = lean_apply_1(x_1, x_130); -x_134 = lean_array_size(x_132); -x_135 = 0; -x_136 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_MapVars_mapFnBody_spec__0(x_1, x_134, x_135, x_132); -x_137 = lean_alloc_ctor(9, 4, 0); -lean_ctor_set(x_137, 0, x_129); -lean_ctor_set(x_137, 1, x_133); -lean_ctor_set(x_137, 2, x_131); -lean_ctor_set(x_137, 3, x_136); -return x_137; +x_135 = lean_apply_1(x_1, x_132); +x_136 = lean_array_size(x_134); +x_137 = 0; +x_138 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_MapVars_mapFnBody_spec__0(x_1, x_136, x_137, x_134); +x_139 = lean_alloc_ctor(9, 4, 0); +lean_ctor_set(x_139, 0, x_131); +lean_ctor_set(x_139, 1, x_135); +lean_ctor_set(x_139, 2, x_133); +lean_ctor_set(x_139, 3, x_138); +return x_139; } } case 10: { -lean_object* x_138; -x_138 = lean_ctor_get(x_2, 0); -lean_inc(x_138); -if (lean_obj_tag(x_138) == 0) -{ -uint8_t x_139; -x_139 = !lean_is_exclusive(x_2); -if (x_139 == 0) -{ -lean_object* x_140; uint8_t x_141; +lean_object* x_140; x_140 = lean_ctor_get(x_2, 0); -lean_dec(x_140); -x_141 = !lean_is_exclusive(x_138); +lean_inc(x_140); +if (lean_obj_tag(x_140) == 0) +{ +uint8_t x_141; +x_141 = !lean_is_exclusive(x_2); if (x_141 == 0) { -lean_object* x_142; lean_object* x_143; -x_142 = lean_ctor_get(x_138, 0); -x_143 = lean_apply_1(x_1, x_142); -lean_ctor_set(x_138, 0, x_143); +lean_object* x_142; uint8_t x_143; +x_142 = lean_ctor_get(x_2, 0); +lean_dec(x_142); +x_143 = !lean_is_exclusive(x_140); +if (x_143 == 0) +{ +lean_object* x_144; lean_object* x_145; +x_144 = lean_ctor_get(x_140, 0); +x_145 = lean_apply_1(x_1, x_144); +lean_ctor_set(x_140, 0, x_145); return x_2; } else { -lean_object* x_144; lean_object* x_145; lean_object* x_146; -x_144 = lean_ctor_get(x_138, 0); -lean_inc(x_144); -lean_dec(x_138); -x_145 = lean_apply_1(x_1, x_144); -x_146 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_146, 0, x_145); -lean_ctor_set(x_2, 0, x_146); +lean_object* x_146; lean_object* x_147; lean_object* x_148; +x_146 = lean_ctor_get(x_140, 0); +lean_inc(x_146); +lean_dec(x_140); +x_147 = lean_apply_1(x_1, x_146); +x_148 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_148, 0, x_147); +lean_ctor_set(x_2, 0, x_148); return x_2; } } else { -lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; +lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_dec(x_2); -x_147 = lean_ctor_get(x_138, 0); -lean_inc(x_147); -if (lean_is_exclusive(x_138)) { - lean_ctor_release(x_138, 0); - x_148 = x_138; +x_149 = lean_ctor_get(x_140, 0); +lean_inc(x_149); +if (lean_is_exclusive(x_140)) { + lean_ctor_release(x_140, 0); + x_150 = x_140; } else { - lean_dec_ref(x_138); - x_148 = lean_box(0); + lean_dec_ref(x_140); + x_150 = lean_box(0); } -x_149 = lean_apply_1(x_1, x_147); -if (lean_is_scalar(x_148)) { - x_150 = lean_alloc_ctor(0, 1, 0); +x_151 = lean_apply_1(x_1, x_149); +if (lean_is_scalar(x_150)) { + x_152 = lean_alloc_ctor(0, 1, 0); } else { - x_150 = x_148; + x_152 = x_150; } -lean_ctor_set(x_150, 0, x_149); -x_151 = lean_alloc_ctor(10, 1, 0); -lean_ctor_set(x_151, 0, x_150); -return x_151; +lean_ctor_set(x_152, 0, x_151); +x_153 = lean_alloc_ctor(10, 1, 0); +lean_ctor_set(x_153, 0, x_152); +return x_153; } } else @@ -5373,29 +5403,29 @@ return x_2; } case 11: { -uint8_t x_152; -x_152 = !lean_is_exclusive(x_2); -if (x_152 == 0) +uint8_t x_154; +x_154 = !lean_is_exclusive(x_2); +if (x_154 == 0) { -lean_object* x_153; lean_object* x_154; -x_153 = lean_ctor_get(x_2, 1); -x_154 = l_Lean_IR_MapVars_mapArgs(x_1, x_153); -lean_ctor_set(x_2, 1, x_154); +lean_object* x_155; lean_object* x_156; +x_155 = lean_ctor_get(x_2, 1); +x_156 = l_Lean_IR_MapVars_mapArgs(x_1, x_155); +lean_ctor_set(x_2, 1, x_156); return x_2; } else { -lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; -x_155 = lean_ctor_get(x_2, 0); -x_156 = lean_ctor_get(x_2, 1); -lean_inc(x_156); -lean_inc(x_155); +lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; +x_157 = lean_ctor_get(x_2, 0); +x_158 = lean_ctor_get(x_2, 1); +lean_inc(x_158); +lean_inc(x_157); lean_dec(x_2); -x_157 = l_Lean_IR_MapVars_mapArgs(x_1, x_156); -x_158 = lean_alloc_ctor(11, 2, 0); -lean_ctor_set(x_158, 0, x_155); -lean_ctor_set(x_158, 1, x_157); -return x_158; +x_159 = l_Lean_IR_MapVars_mapArgs(x_1, x_158); +x_160 = lean_alloc_ctor(11, 2, 0); +lean_ctor_set(x_160, 0, x_157); +lean_ctor_set(x_160, 1, x_159); +return x_160; } } default: diff --git a/stage0/stdlib/Lean/Compiler/IR/RC.c b/stage0/stdlib/Lean/Compiler/IR/RC.c index 814c1b17f53f..a8cbbe5905ce 100644 --- a/stage0/stdlib/Lean/Compiler/IR/RC.c +++ b/stage0/stdlib/Lean/Compiler/IR/RC.c @@ -313,6 +313,7 @@ LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_m static lean_object* l___private_Lean_Compiler_IR_RC_0__Lean_IR_initFn___closed__25_00___x40_Lean_Compiler_IR_RC_3281171570____hygCtx___hyg_2_; size_t lean_usize_sub(size_t, size_t); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useArgs_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_IR_IRType_isStruct(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useVar___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useArgs_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); lean_object* l_mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -5242,141 +5243,156 @@ x_3 = lean_ctor_get(x_1, 2); switch (lean_obj_tag(x_3)) { case 3: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_inc_ref(x_3); x_4 = lean_ctor_get(x_1, 0); lean_inc(x_4); -x_5 = lean_ctor_get(x_1, 3); +x_5 = lean_ctor_get(x_1, 1); lean_inc(x_5); -lean_dec_ref(x_1); -x_6 = lean_ctor_get(x_3, 1); +x_6 = lean_ctor_get(x_1, 3); lean_inc(x_6); +lean_dec_ref(x_1); +x_7 = lean_ctor_get(x_3, 2); +lean_inc(x_7); lean_dec_ref(x_3); -x_7 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_addDerivedValue(x_6, x_4, x_2); -x_8 = lean_ctor_get(x_7, 1); -lean_inc(x_8); -lean_dec_ref(x_7); -x_1 = x_5; -x_2 = x_8; +x_8 = l_Lean_IR_IRType_isScalar(x_5); +lean_dec(x_5); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; +x_9 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_addDerivedValue(x_7, x_4, x_2); +x_10 = lean_ctor_get(x_9, 1); +lean_inc(x_10); +lean_dec_ref(x_9); +x_1 = x_6; +x_2 = x_10; +goto _start; +} +else +{ +lean_dec(x_7); +lean_dec(x_4); +x_1 = x_6; goto _start; } +} case 6: { -lean_object* x_10; -x_10 = lean_ctor_get(x_3, 0); -if (lean_obj_tag(x_10) == 1) +lean_object* x_13; +x_13 = lean_ctor_get(x_3, 0); +if (lean_obj_tag(x_13) == 1) { -lean_object* x_11; -x_11 = lean_ctor_get(x_10, 0); -if (lean_obj_tag(x_11) == 1) +lean_object* x_14; +x_14 = lean_ctor_get(x_13, 0); +if (lean_obj_tag(x_14) == 1) { -lean_object* x_12; -x_12 = lean_ctor_get(x_11, 0); -if (lean_obj_tag(x_12) == 0) +lean_object* x_15; +x_15 = lean_ctor_get(x_14, 0); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -lean_inc_ref(x_11); -lean_inc_ref(x_10); +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +lean_inc_ref(x_14); +lean_inc_ref(x_13); lean_inc_ref(x_3); -x_13 = lean_ctor_get(x_1, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_1, 3); -lean_inc(x_14); +x_16 = lean_ctor_get(x_1, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_1, 3); +lean_inc(x_17); lean_dec_ref(x_1); -x_15 = lean_ctor_get(x_3, 1); -lean_inc_ref(x_15); +x_18 = lean_ctor_get(x_3, 1); +lean_inc_ref(x_18); lean_dec_ref(x_3); -x_16 = lean_ctor_get(x_10, 1); -lean_inc_ref(x_16); -lean_dec_ref(x_10); -x_17 = lean_ctor_get(x_11, 1); -lean_inc_ref(x_17); -lean_dec_ref(x_11); -x_18 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_visitFnBody___closed__0; -x_19 = lean_string_dec_eq(x_17, x_18); -lean_dec_ref(x_17); -if (x_19 == 0) +x_19 = lean_ctor_get(x_13, 1); +lean_inc_ref(x_19); +lean_dec_ref(x_13); +x_20 = lean_ctor_get(x_14, 1); +lean_inc_ref(x_20); +lean_dec_ref(x_14); +x_21 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_visitFnBody___closed__0; +x_22 = lean_string_dec_eq(x_20, x_21); +lean_dec_ref(x_20); +if (x_22 == 0) { -lean_dec_ref(x_16); -lean_dec_ref(x_15); -lean_dec(x_13); -x_1 = x_14; +lean_dec_ref(x_19); +lean_dec_ref(x_18); +lean_dec(x_16); +x_1 = x_17; goto _start; } else { -lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_21 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_visitFnBody___closed__1; -x_22 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_visitFnBody___closed__2; -x_23 = lean_string_dec_eq(x_16, x_22); -if (x_23 == 0) +lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_24 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_visitFnBody___closed__1; +x_25 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_visitFnBody___closed__2; +x_26 = lean_string_dec_eq(x_19, x_25); +if (x_26 == 0) { -lean_object* x_24; uint8_t x_25; -x_24 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_visitFnBody___closed__3; -x_25 = lean_string_dec_eq(x_16, x_24); -lean_dec_ref(x_16); -if (x_25 == 0) +lean_object* x_27; uint8_t x_28; +x_27 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_visitFnBody___closed__3; +x_28 = lean_string_dec_eq(x_19, x_27); +lean_dec_ref(x_19); +if (x_28 == 0) { -lean_dec_ref(x_15); -lean_dec(x_13); -x_1 = x_14; +lean_dec_ref(x_18); +lean_dec(x_16); +x_1 = x_17; goto _start; } else { -lean_object* x_27; lean_object* x_28; -x_27 = lean_unsigned_to_nat(2u); -x_28 = lean_array_get(x_21, x_15, x_27); -lean_dec_ref(x_15); -if (lean_obj_tag(x_28) == 0) +lean_object* x_30; lean_object* x_31; +x_30 = lean_unsigned_to_nat(2u); +x_31 = lean_array_get(x_24, x_18, x_30); +lean_dec_ref(x_18); +if (lean_obj_tag(x_31) == 0) { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -lean_dec_ref(x_28); -x_30 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_addDerivedValue(x_29, x_13, x_2); -x_31 = lean_ctor_get(x_30, 1); -lean_inc(x_31); -lean_dec_ref(x_30); -x_1 = x_14; -x_2 = x_31; +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +lean_dec_ref(x_31); +x_33 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_addDerivedValue(x_32, x_16, x_2); +x_34 = lean_ctor_get(x_33, 1); +lean_inc(x_34); +lean_dec_ref(x_33); +x_1 = x_17; +x_2 = x_34; goto _start; } else { -lean_dec(x_28); -lean_dec(x_13); -x_1 = x_14; +lean_dec(x_31); +lean_dec(x_16); +x_1 = x_17; goto _start; } } } else { -lean_object* x_34; lean_object* x_35; -lean_dec_ref(x_16); -x_34 = lean_unsigned_to_nat(1u); -x_35 = lean_array_get(x_21, x_15, x_34); -lean_dec_ref(x_15); -if (lean_obj_tag(x_35) == 0) +lean_object* x_37; lean_object* x_38; +lean_dec_ref(x_19); +x_37 = lean_unsigned_to_nat(1u); +x_38 = lean_array_get(x_24, x_18, x_37); +lean_dec_ref(x_18); +if (lean_obj_tag(x_38) == 0) { -lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_36 = lean_ctor_get(x_35, 0); -lean_inc(x_36); -lean_dec_ref(x_35); -x_37 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_addDerivedValue(x_36, x_13, x_2); -x_38 = lean_ctor_get(x_37, 1); -lean_inc(x_38); -lean_dec_ref(x_37); -x_1 = x_14; -x_2 = x_38; +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +lean_dec_ref(x_38); +x_40 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_addDerivedValue(x_39, x_16, x_2); +x_41 = lean_ctor_get(x_40, 1); +lean_inc(x_41); +lean_dec_ref(x_40); +x_1 = x_17; +x_2 = x_41; goto _start; } else { -lean_dec(x_35); -lean_dec(x_13); -x_1 = x_14; +lean_dec(x_38); +lean_dec(x_16); +x_1 = x_17; goto _start; } } @@ -5384,215 +5400,250 @@ goto _start; } else { -lean_object* x_41; -x_41 = lean_ctor_get(x_1, 3); -lean_inc(x_41); +lean_object* x_44; +x_44 = lean_ctor_get(x_1, 3); +lean_inc(x_44); lean_dec_ref(x_1); -x_1 = x_41; +x_1 = x_44; goto _start; } } else { -lean_object* x_43; -x_43 = lean_ctor_get(x_1, 3); -lean_inc(x_43); +lean_object* x_46; +x_46 = lean_ctor_get(x_1, 3); +lean_inc(x_46); lean_dec_ref(x_1); -x_1 = x_43; +x_1 = x_46; goto _start; } } else { -lean_object* x_45; -x_45 = lean_ctor_get(x_1, 3); -lean_inc(x_45); +lean_object* x_48; +x_48 = lean_ctor_get(x_1, 3); +lean_inc(x_48); lean_dec_ref(x_1); -x_1 = x_45; +x_1 = x_48; goto _start; } } case 1: { -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_inc_ref(x_3); -x_47 = lean_ctor_get(x_1, 3); -lean_inc(x_47); +x_50 = lean_ctor_get(x_1, 3); +lean_inc(x_50); lean_dec_ref(x_1); -x_48 = lean_ctor_get(x_3, 1); -lean_inc(x_48); +x_51 = lean_ctor_get(x_3, 1); +lean_inc(x_51); lean_dec_ref(x_3); -x_49 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_removeFromParent(x_48, x_2); -lean_dec(x_48); -x_50 = lean_ctor_get(x_49, 1); -lean_inc(x_50); -lean_dec_ref(x_49); -x_1 = x_47; -x_2 = x_50; +x_52 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_removeFromParent(x_51, x_2); +lean_dec(x_51); +x_53 = lean_ctor_get(x_52, 1); +lean_inc(x_53); +lean_dec_ref(x_52); +x_1 = x_50; +x_2 = x_53; goto _start; } +case 10: +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; +lean_inc_ref(x_3); +x_55 = lean_ctor_get(x_1, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_1, 1); +lean_inc(x_56); +x_57 = lean_ctor_get(x_1, 3); +lean_inc(x_57); +lean_dec_ref(x_1); +x_58 = lean_ctor_get(x_3, 0); +lean_inc(x_58); +lean_dec_ref(x_3); +x_59 = l_Lean_IR_IRType_isStruct(x_56); +lean_dec(x_56); +if (x_59 == 0) +{ +lean_dec(x_58); +lean_dec(x_55); +x_1 = x_57; +goto _start; +} +else +{ +lean_object* x_61; lean_object* x_62; +x_61 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_addDerivedValue(x_58, x_55, x_2); +x_62 = lean_ctor_get(x_61, 1); +lean_inc(x_62); +lean_dec_ref(x_61); +x_1 = x_57; +x_2 = x_62; +goto _start; +} +} default: { -lean_object* x_52; -x_52 = lean_ctor_get(x_1, 3); -lean_inc(x_52); +lean_object* x_64; +x_64 = lean_ctor_get(x_1, 3); +lean_inc(x_64); lean_dec_ref(x_1); -x_1 = x_52; +x_1 = x_64; goto _start; } } } case 1: { -lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_62; lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_54 = lean_ctor_get(x_1, 1); -lean_inc_ref(x_54); -x_55 = lean_ctor_get(x_1, 2); -lean_inc(x_55); -x_56 = lean_ctor_get(x_1, 3); -lean_inc(x_56); +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_74; lean_object* x_77; lean_object* x_78; uint8_t x_79; +x_66 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_66); +x_67 = lean_ctor_get(x_1, 2); +lean_inc(x_67); +x_68 = lean_ctor_get(x_1, 3); +lean_inc(x_68); lean_dec_ref(x_1); -x_65 = lean_unsigned_to_nat(0u); -x_66 = lean_array_get_size(x_54); -x_67 = lean_nat_dec_lt(x_65, x_66); -if (x_67 == 0) +x_77 = lean_unsigned_to_nat(0u); +x_78 = lean_array_get_size(x_66); +x_79 = lean_nat_dec_lt(x_77, x_78); +if (x_79 == 0) { -lean_dec_ref(x_54); -x_57 = x_2; -goto block_61; +lean_dec_ref(x_66); +x_69 = x_2; +goto block_73; } else { -lean_object* x_68; uint8_t x_69; -x_68 = lean_box(0); -x_69 = lean_nat_dec_le(x_66, x_66); -if (x_69 == 0) +lean_object* x_80; uint8_t x_81; +x_80 = lean_box(0); +x_81 = lean_nat_dec_le(x_78, x_78); +if (x_81 == 0) { -if (x_67 == 0) +if (x_79 == 0) { -lean_dec_ref(x_54); -x_57 = x_2; -goto block_61; +lean_dec_ref(x_66); +x_69 = x_2; +goto block_73; } else { -size_t x_70; size_t x_71; lean_object* x_72; -x_70 = 0; -x_71 = lean_usize_of_nat(x_66); -x_72 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_visitFnBody_spec__0(x_54, x_70, x_71, x_68, x_2); -lean_dec_ref(x_54); -x_62 = x_72; -goto block_64; +size_t x_82; size_t x_83; lean_object* x_84; +x_82 = 0; +x_83 = lean_usize_of_nat(x_78); +x_84 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_visitFnBody_spec__0(x_66, x_82, x_83, x_80, x_2); +lean_dec_ref(x_66); +x_74 = x_84; +goto block_76; } } else { -size_t x_73; size_t x_74; lean_object* x_75; -x_73 = 0; -x_74 = lean_usize_of_nat(x_66); -x_75 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_visitFnBody_spec__0(x_54, x_73, x_74, x_68, x_2); -lean_dec_ref(x_54); -x_62 = x_75; -goto block_64; +size_t x_85; size_t x_86; lean_object* x_87; +x_85 = 0; +x_86 = lean_usize_of_nat(x_78); +x_87 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_visitFnBody_spec__0(x_66, x_85, x_86, x_80, x_2); +lean_dec_ref(x_66); +x_74 = x_87; +goto block_76; } } -block_61: +block_73: { -lean_object* x_58; lean_object* x_59; -x_58 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_visitFnBody(x_55, x_57); -x_59 = lean_ctor_get(x_58, 1); -lean_inc(x_59); -lean_dec_ref(x_58); -x_1 = x_56; -x_2 = x_59; +lean_object* x_70; lean_object* x_71; +x_70 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_visitFnBody(x_67, x_69); +x_71 = lean_ctor_get(x_70, 1); +lean_inc(x_71); +lean_dec_ref(x_70); +x_1 = x_68; +x_2 = x_71; goto _start; } -block_64: +block_76: { -lean_object* x_63; -x_63 = lean_ctor_get(x_62, 1); -lean_inc(x_63); -lean_dec_ref(x_62); -x_57 = x_63; -goto block_61; +lean_object* x_75; +x_75 = lean_ctor_get(x_74, 1); +lean_inc(x_75); +lean_dec_ref(x_74); +x_69 = x_75; +goto block_73; } } case 9: { -lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; uint8_t x_80; -x_76 = lean_ctor_get(x_1, 3); -lean_inc_ref(x_76); +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; uint8_t x_92; +x_88 = lean_ctor_get(x_1, 3); +lean_inc_ref(x_88); lean_dec_ref(x_1); -x_77 = lean_unsigned_to_nat(0u); -x_78 = lean_array_get_size(x_76); -x_79 = lean_box(0); -x_80 = lean_nat_dec_lt(x_77, x_78); -if (x_80 == 0) -{ -lean_object* x_81; -lean_dec_ref(x_76); -x_81 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_81, 0, x_79); -lean_ctor_set(x_81, 1, x_2); -return x_81; +x_89 = lean_unsigned_to_nat(0u); +x_90 = lean_array_get_size(x_88); +x_91 = lean_box(0); +x_92 = lean_nat_dec_lt(x_89, x_90); +if (x_92 == 0) +{ +lean_object* x_93; +lean_dec_ref(x_88); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_2); +return x_93; } else { -uint8_t x_82; -x_82 = lean_nat_dec_le(x_78, x_78); -if (x_82 == 0) +uint8_t x_94; +x_94 = lean_nat_dec_le(x_90, x_90); +if (x_94 == 0) { -if (x_80 == 0) +if (x_92 == 0) { -lean_object* x_83; -lean_dec_ref(x_76); -x_83 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_83, 0, x_79); -lean_ctor_set(x_83, 1, x_2); -return x_83; +lean_object* x_95; +lean_dec_ref(x_88); +x_95 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_95, 0, x_91); +lean_ctor_set(x_95, 1, x_2); +return x_95; } else { -size_t x_84; size_t x_85; lean_object* x_86; -x_84 = 0; -x_85 = lean_usize_of_nat(x_78); -x_86 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_visitFnBody_spec__1(x_76, x_84, x_85, x_79, x_2); -lean_dec_ref(x_76); -return x_86; +size_t x_96; size_t x_97; lean_object* x_98; +x_96 = 0; +x_97 = lean_usize_of_nat(x_90); +x_98 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_visitFnBody_spec__1(x_88, x_96, x_97, x_91, x_2); +lean_dec_ref(x_88); +return x_98; } } else { -size_t x_87; size_t x_88; lean_object* x_89; -x_87 = 0; -x_88 = lean_usize_of_nat(x_78); -x_89 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_visitFnBody_spec__1(x_76, x_87, x_88, x_79, x_2); -lean_dec_ref(x_76); -return x_89; +size_t x_99; size_t x_100; lean_object* x_101; +x_99 = 0; +x_100 = lean_usize_of_nat(x_90); +x_101 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_visitFnBody_spec__1(x_88, x_99, x_100, x_91, x_2); +lean_dec_ref(x_88); +return x_101; } } } default: { -uint8_t x_90; -x_90 = l_Lean_IR_FnBody_isTerminal(x_1); -if (x_90 == 0) +uint8_t x_102; +x_102 = l_Lean_IR_FnBody_isTerminal(x_1); +if (x_102 == 0) { -lean_object* x_91; -x_91 = l_Lean_IR_FnBody_body(x_1); +lean_object* x_103; +x_103 = l_Lean_IR_FnBody_body(x_1); lean_dec(x_1); -x_1 = x_91; +x_1 = x_103; goto _start; } else { -lean_object* x_93; lean_object* x_94; +lean_object* x_105; lean_object* x_106; lean_dec(x_1); -x_93 = lean_box(0); -x_94 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_94, 0, x_93); -lean_ctor_set(x_94, 1, x_2); -return x_94; +x_105 = lean_box(0); +x_106 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_106, 0, x_105); +lean_ctor_set(x_106, 1, x_2); +return x_106; } } } @@ -7596,14 +7647,13 @@ LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_u _start: { switch (lean_obj_tag(x_2)) { -case 0: +case 1: { lean_object* x_4; lean_object* x_5; x_4 = lean_ctor_get(x_2, 1); -lean_inc_ref(x_4); +lean_inc(x_4); lean_dec_ref(x_2); -x_5 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useArgs(x_1, x_4, x_3); -lean_dec_ref(x_4); +x_5 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useVar___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useExpr_spec__0(x_1, x_4, x_3); return x_5; } case 2: @@ -7619,7 +7669,7 @@ x_9 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useArgs(x_1, x_7, x_ lean_dec_ref(x_7); return x_9; } -case 5: +case 3: { lean_object* x_10; lean_object* x_11; x_10 = lean_ctor_get(x_2, 2); @@ -7628,24 +7678,22 @@ lean_dec_ref(x_2); x_11 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useVar___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useExpr_spec__0(x_1, x_10, x_3); return x_11; } -case 6: +case 4: { lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_2, 1); -lean_inc_ref(x_12); +x_12 = lean_ctor_get(x_2, 2); +lean_inc(x_12); lean_dec_ref(x_2); -x_13 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useArgs(x_1, x_12, x_3); -lean_dec_ref(x_12); +x_13 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useVar___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useExpr_spec__0(x_1, x_12, x_3); return x_13; } -case 7: +case 5: { lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_2, 1); -lean_inc_ref(x_14); +x_14 = lean_ctor_get(x_2, 3); +lean_inc(x_14); lean_dec_ref(x_2); -x_15 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useArgs(x_1, x_14, x_3); -lean_dec_ref(x_14); +x_15 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useVar___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useExpr_spec__0(x_1, x_14, x_3); return x_15; } case 8: @@ -7661,21 +7709,16 @@ x_19 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useArgs(x_1, x_17, lean_dec_ref(x_17); return x_19; } -case 10: +case 9: { lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_2, 0); +x_20 = lean_ctor_get(x_2, 1); lean_inc(x_20); lean_dec_ref(x_2); x_21 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useVar___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useExpr_spec__0(x_1, x_20, x_3); return x_21; } -case 11: -{ -lean_dec_ref(x_2); -return x_3; -} -case 12: +case 10: { lean_object* x_22; lean_object* x_23; x_22 = lean_ctor_get(x_2, 0); @@ -7684,15 +7727,30 @@ lean_dec_ref(x_2); x_23 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useVar___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useExpr_spec__0(x_1, x_22, x_3); return x_23; } -default: +case 11: +{ +lean_dec_ref(x_2); +return x_3; +} +case 12: { lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_2, 1); +x_24 = lean_ctor_get(x_2, 0); lean_inc(x_24); lean_dec_ref(x_2); x_25 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useVar___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useExpr_spec__0(x_1, x_24, x_3); return x_25; } +default: +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_2, 1); +lean_inc_ref(x_26); +lean_dec_ref(x_2); +x_27 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useArgs(x_1, x_26, x_3); +lean_dec_ref(x_26); +return x_27; +} } } } @@ -9647,218 +9705,227 @@ goto block_21; } case 3: { -lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; -x_29 = lean_ctor_get(x_4, 1); -x_30 = lean_ctor_get(x_6, 1); +lean_object* x_29; lean_object* x_30; lean_object* x_33; lean_object* x_34; uint8_t x_35; uint8_t x_40; uint8_t x_46; +x_29 = lean_ctor_get(x_4, 2); lean_inc(x_29); -x_31 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_addDecIfNeeded(x_1, x_29, x_22, x_6); -x_32 = l_Std_DTreeMap_Internal_Impl_contains___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_visitParam_spec__1___redArg(x_2, x_30); -if (x_32 == 0) +x_30 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_addDecIfNeeded(x_1, x_29, x_22, x_6); +x_46 = l_Lean_IR_IRType_isPossibleRef(x_3); +if (x_46 == 0) { -lean_object* x_33; uint8_t x_34; lean_object* x_35; uint8_t x_36; -x_33 = l_Lean_IR_ExplicitRC_getVarInfo(x_1, x_2); -x_34 = lean_ctor_get_uint8(x_33, 1); -x_35 = lean_unsigned_to_nat(1u); -if (x_34 == 0) +x_40 = x_46; +goto block_45; +} +else { -uint8_t x_41; -x_41 = 1; -x_36 = x_41; -goto block_40; +lean_object* x_47; uint8_t x_48; +x_47 = lean_ctor_get(x_6, 1); +x_48 = l_Std_DTreeMap_Internal_Impl_contains___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_visitParam_spec__1___redArg(x_2, x_47); +if (x_48 == 0) +{ +x_40 = x_46; +goto block_45; } else { -x_36 = x_32; -goto block_40; +goto block_32; +} } -block_40: +block_32: { -uint8_t x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get_uint8(x_33, 2); -lean_dec_ref(x_33); -lean_inc(x_2); -x_38 = lean_alloc_ctor(6, 3, 2); -lean_ctor_set(x_38, 0, x_2); -lean_ctor_set(x_38, 1, x_35); -lean_ctor_set(x_38, 2, x_31); -lean_ctor_set_uint8(x_38, sizeof(void*)*3, x_36); -lean_ctor_set_uint8(x_38, sizeof(void*)*3 + 1, x_37); +lean_object* x_31; lean_inc_ref(x_4); lean_inc(x_2); -x_39 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_39, 0, x_2); -lean_ctor_set(x_39, 1, x_3); -lean_ctor_set(x_39, 2, x_4); -lean_ctor_set(x_39, 3, x_38); -x_7 = x_39; +x_31 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_31, 0, x_2); +lean_ctor_set(x_31, 1, x_3); +lean_ctor_set(x_31, 2, x_4); +lean_ctor_set(x_31, 3, x_30); +x_7 = x_31; goto block_21; } -} -else +block_39: { -lean_object* x_42; +uint8_t x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get_uint8(x_34, 2); +lean_dec_ref(x_34); +lean_inc(x_2); +x_37 = lean_alloc_ctor(6, 3, 2); +lean_ctor_set(x_37, 0, x_2); +lean_ctor_set(x_37, 1, x_33); +lean_ctor_set(x_37, 2, x_30); +lean_ctor_set_uint8(x_37, sizeof(void*)*3, x_35); +lean_ctor_set_uint8(x_37, sizeof(void*)*3 + 1, x_36); lean_inc_ref(x_4); lean_inc(x_2); -x_42 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_42, 0, x_2); -lean_ctor_set(x_42, 1, x_3); -lean_ctor_set(x_42, 2, x_4); -lean_ctor_set(x_42, 3, x_31); -x_7 = x_42; +x_38 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_38, 0, x_2); +lean_ctor_set(x_38, 1, x_3); +lean_ctor_set(x_38, 2, x_4); +lean_ctor_set(x_38, 3, x_37); +x_7 = x_38; goto block_21; } +block_45: +{ +if (x_40 == 0) +{ +goto block_32; +} +else +{ +lean_object* x_41; uint8_t x_42; lean_object* x_43; +x_41 = l_Lean_IR_ExplicitRC_getVarInfo(x_1, x_2); +x_42 = lean_ctor_get_uint8(x_41, 1); +x_43 = lean_unsigned_to_nat(1u); +if (x_42 == 0) +{ +x_33 = x_43; +x_34 = x_41; +x_35 = x_40; +goto block_39; +} +else +{ +uint8_t x_44; +x_44 = 0; +x_33 = x_43; +x_34 = x_41; +x_35 = x_44; +goto block_39; +} +} +} } case 4: { -lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_43 = lean_ctor_get(x_4, 1); -lean_inc(x_43); -x_44 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_addDecIfNeeded(x_1, x_43, x_22, x_6); +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_4, 2); +lean_inc(x_49); +x_50 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_addDecIfNeeded(x_1, x_49, x_22, x_6); lean_inc_ref(x_4); lean_inc(x_2); -x_45 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_45, 0, x_2); -lean_ctor_set(x_45, 1, x_3); -lean_ctor_set(x_45, 2, x_4); -lean_ctor_set(x_45, 3, x_44); -x_7 = x_45; +x_51 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_51, 0, x_2); +lean_ctor_set(x_51, 1, x_3); +lean_ctor_set(x_51, 2, x_4); +lean_ctor_set(x_51, 3, x_50); +x_7 = x_51; goto block_21; } case 5: { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_4, 2); -lean_inc(x_46); -x_47 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_addDecIfNeeded(x_1, x_46, x_22, x_6); +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_4, 3); +lean_inc(x_52); +x_53 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_addDecIfNeeded(x_1, x_52, x_22, x_6); lean_inc_ref(x_4); lean_inc(x_2); -x_48 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_48, 0, x_2); -lean_ctor_set(x_48, 1, x_3); -lean_ctor_set(x_48, 2, x_4); -lean_ctor_set(x_48, 3, x_47); -x_7 = x_48; +x_54 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_54, 0, x_2); +lean_ctor_set(x_54, 1, x_3); +lean_ctor_set(x_54, 2, x_4); +lean_ctor_set(x_54, 3, x_53); +x_7 = x_54; goto block_21; } case 6: { -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_58; uint8_t x_62; lean_object* x_70; uint8_t x_71; -x_49 = lean_ctor_get(x_4, 0); -x_50 = lean_ctor_get(x_4, 1); -lean_inc(x_49); +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_64; uint8_t x_68; lean_object* x_76; uint8_t x_77; +x_55 = lean_ctor_get(x_4, 0); +x_56 = lean_ctor_get(x_4, 1); +lean_inc(x_55); lean_inc_ref(x_1); -x_51 = l_Lean_IR_ExplicitRC_getDecl(x_1, x_49); -x_52 = l_Lean_IR_Decl_params(x_51); -lean_dec_ref(x_51); -lean_inc_ref(x_52); -x_53 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_addDecAfterFullApp(x_1, x_50, x_52, x_22, x_6); -x_70 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_processVDecl___closed__5; -x_71 = lean_name_eq(x_49, x_70); -if (x_71 == 0) +x_57 = l_Lean_IR_ExplicitRC_getDecl(x_1, x_55); +x_58 = l_Lean_IR_Decl_params(x_57); +lean_dec_ref(x_57); +lean_inc_ref(x_58); +x_59 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_addDecAfterFullApp(x_1, x_56, x_58, x_22, x_6); +x_76 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_processVDecl___closed__5; +x_77 = lean_name_eq(x_55, x_76); +if (x_77 == 0) { -x_62 = x_71; -goto block_69; +x_68 = x_77; +goto block_75; } else { -lean_object* x_72; uint8_t x_73; -x_72 = lean_box(0); -x_73 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_processVDecl___lam__0(x_6, x_2, x_72); -x_62 = x_73; -goto block_69; +lean_object* x_78; uint8_t x_79; +x_78 = lean_box(0); +x_79 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_processVDecl___lam__0(x_6, x_2, x_78); +x_68 = x_79; +goto block_75; } -block_57: +block_63: { -lean_object* x_55; lean_object* x_56; +lean_object* x_61; lean_object* x_62; lean_inc(x_2); -x_55 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_55, 0, x_2); -lean_ctor_set(x_55, 1, x_3); -lean_ctor_set(x_55, 2, x_54); -lean_ctor_set(x_55, 3, x_53); -x_56 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_addIncBefore(x_1, x_50, x_52, x_55, x_6); -x_7 = x_56; +x_61 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_61, 0, x_2); +lean_ctor_set(x_61, 1, x_3); +lean_ctor_set(x_61, 2, x_60); +lean_ctor_set(x_61, 3, x_59); +x_62 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_addIncBefore(x_1, x_56, x_58, x_61, x_6); +x_7 = x_62; goto block_21; } -block_61: +block_67: { -if (x_58 == 0) +if (x_64 == 0) { lean_inc_ref(x_4); -x_54 = x_4; -goto block_57; +x_60 = x_4; +goto block_63; } else { -lean_object* x_59; lean_object* x_60; -x_59 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_processVDecl___closed__1; -lean_inc_ref(x_50); -x_60 = lean_alloc_ctor(6, 2, 0); -lean_ctor_set(x_60, 0, x_59); -lean_ctor_set(x_60, 1, x_50); -x_54 = x_60; -goto block_57; +lean_object* x_65; lean_object* x_66; +x_65 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_processVDecl___closed__1; +lean_inc_ref(x_56); +x_66 = lean_alloc_ctor(6, 2, 0); +lean_ctor_set(x_66, 0, x_65); +lean_ctor_set(x_66, 1, x_56); +x_60 = x_66; +goto block_63; } } -block_69: +block_75: { -if (x_62 == 0) +if (x_68 == 0) { -lean_object* x_63; uint8_t x_64; -x_63 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_processVDecl___closed__2; -x_64 = lean_name_eq(x_49, x_63); -if (x_64 == 0) +lean_object* x_69; uint8_t x_70; +x_69 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_processVDecl___closed__2; +x_70 = lean_name_eq(x_55, x_69); +if (x_70 == 0) { -x_58 = x_64; -goto block_61; +x_64 = x_70; +goto block_67; } else { -lean_object* x_65; uint8_t x_66; -x_65 = lean_box(0); -x_66 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_processVDecl___lam__0(x_6, x_2, x_65); -x_58 = x_66; -goto block_61; +lean_object* x_71; uint8_t x_72; +x_71 = lean_box(0); +x_72 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_processVDecl___lam__0(x_6, x_2, x_71); +x_64 = x_72; +goto block_67; } } else { -lean_object* x_67; lean_object* x_68; -x_67 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_processVDecl___closed__4; -lean_inc_ref(x_50); -x_68 = lean_alloc_ctor(6, 2, 0); -lean_ctor_set(x_68, 0, x_67); -lean_ctor_set(x_68, 1, x_50); -x_54 = x_68; -goto block_57; +lean_object* x_73; lean_object* x_74; +x_73 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_processVDecl___closed__4; +lean_inc_ref(x_56); +x_74 = lean_alloc_ctor(6, 2, 0); +lean_ctor_set(x_74, 0, x_73); +lean_ctor_set(x_74, 1, x_56); +x_60 = x_74; +goto block_63; } } } case 7: { -lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_74 = lean_ctor_get(x_4, 1); -lean_inc_ref(x_4); -lean_inc(x_2); -x_75 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_75, 0, x_2); -lean_ctor_set(x_75, 1, x_3); -lean_ctor_set(x_75, 2, x_4); -lean_ctor_set(x_75, 3, x_22); -x_76 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_addIncBeforeConsumeAll(x_1, x_74, x_75, x_6); -x_7 = x_76; -goto block_21; -} -case 8: -{ -lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_77 = lean_ctor_get(x_4, 0); -x_78 = lean_ctor_get(x_4, 1); -lean_inc(x_77); -x_79 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_79, 0, x_77); -lean_inc_ref(x_78); -x_80 = lean_array_push(x_78, x_79); +lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_80 = lean_ctor_get(x_4, 1); lean_inc_ref(x_4); lean_inc(x_2); x_81 = lean_alloc_ctor(0, 4, 0); @@ -9867,37 +9934,227 @@ lean_ctor_set(x_81, 1, x_3); lean_ctor_set(x_81, 2, x_4); lean_ctor_set(x_81, 3, x_22); x_82 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_addIncBeforeConsumeAll(x_1, x_80, x_81, x_6); -lean_dec_ref(x_80); x_7 = x_82; goto block_21; } -case 10: +case 8: { -lean_object* x_83; lean_object* x_84; lean_object* x_85; +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; x_83 = lean_ctor_get(x_4, 0); +x_84 = lean_ctor_get(x_4, 1); lean_inc(x_83); -x_84 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_addDecIfNeeded(x_1, x_83, x_22, x_6); +x_85 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_85, 0, x_83); +lean_inc_ref(x_84); +x_86 = lean_array_push(x_84, x_85); +lean_inc_ref(x_4); +lean_inc(x_2); +x_87 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_87, 0, x_2); +lean_ctor_set(x_87, 1, x_3); +lean_ctor_set(x_87, 2, x_4); +lean_ctor_set(x_87, 3, x_22); +x_88 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_addIncBeforeConsumeAll(x_1, x_86, x_87, x_6); +lean_dec_ref(x_86); +x_7 = x_88; +goto block_21; +} +case 9: +{ +lean_object* x_89; lean_object* x_90; uint8_t x_91; +x_89 = lean_ctor_get(x_4, 0); +x_90 = lean_ctor_get(x_4, 1); +x_91 = l_Lean_IR_IRType_isStruct(x_89); +if (x_91 == 0) +{ +lean_object* x_92; +lean_inc_ref(x_4); +lean_inc(x_2); +x_92 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_92, 0, x_2); +lean_ctor_set(x_92, 1, x_3); +lean_ctor_set(x_92, 2, x_4); +lean_ctor_set(x_92, 3, x_22); +x_7 = x_92; +goto block_21; +} +else +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; uint8_t x_98; uint8_t x_102; uint8_t x_108; +x_93 = lean_ctor_get(x_6, 0); +x_94 = lean_ctor_get(x_6, 1); +lean_inc_ref(x_4); +lean_inc(x_2); +x_95 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_95, 0, x_2); +lean_ctor_set(x_95, 1, x_3); +lean_ctor_set(x_95, 2, x_4); +lean_ctor_set(x_95, 3, x_22); +x_108 = l_Std_DTreeMap_Internal_Impl_contains___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_visitParam_spec__1___redArg(x_90, x_93); +if (x_108 == 0) +{ +uint8_t x_109; +x_109 = l_Std_DTreeMap_Internal_Impl_contains___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_visitParam_spec__1___redArg(x_90, x_94); +if (x_109 == 0) +{ +x_7 = x_95; +goto block_21; +} +else +{ +x_102 = x_109; +goto block_107; +} +} +else +{ +x_102 = x_108; +goto block_107; +} +block_101: +{ +uint8_t x_99; lean_object* x_100; +x_99 = lean_ctor_get_uint8(x_96, 2); +lean_dec_ref(x_96); +lean_inc(x_90); +x_100 = lean_alloc_ctor(6, 3, 2); +lean_ctor_set(x_100, 0, x_90); +lean_ctor_set(x_100, 1, x_97); +lean_ctor_set(x_100, 2, x_95); +lean_ctor_set_uint8(x_100, sizeof(void*)*3, x_98); +lean_ctor_set_uint8(x_100, sizeof(void*)*3 + 1, x_99); +x_7 = x_100; +goto block_21; +} +block_107: +{ +lean_object* x_103; uint8_t x_104; lean_object* x_105; +x_103 = l_Lean_IR_ExplicitRC_getVarInfo(x_1, x_90); +x_104 = lean_ctor_get_uint8(x_103, 1); +x_105 = lean_unsigned_to_nat(1u); +if (x_104 == 0) +{ +x_96 = x_103; +x_97 = x_105; +x_98 = x_102; +goto block_101; +} +else +{ +uint8_t x_106; +x_106 = 0; +x_96 = x_103; +x_97 = x_105; +x_98 = x_106; +goto block_101; +} +} +} +} +case 10: +{ +lean_object* x_110; lean_object* x_111; lean_object* x_114; lean_object* x_115; uint8_t x_116; uint8_t x_121; uint8_t x_127; +x_110 = lean_ctor_get(x_4, 0); +lean_inc(x_110); +x_111 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_addDecIfNeeded(x_1, x_110, x_22, x_6); +x_127 = l_Lean_IR_IRType_isPossibleRef(x_3); +if (x_127 == 0) +{ +x_121 = x_127; +goto block_126; +} +else +{ +lean_object* x_128; uint8_t x_129; +x_128 = lean_ctor_get(x_6, 1); +x_129 = l_Std_DTreeMap_Internal_Impl_contains___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_visitParam_spec__1___redArg(x_2, x_128); +if (x_129 == 0) +{ +x_121 = x_127; +goto block_126; +} +else +{ +goto block_113; +} +} +block_113: +{ +lean_object* x_112; +lean_inc_ref(x_4); +lean_inc(x_2); +x_112 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_112, 0, x_2); +lean_ctor_set(x_112, 1, x_3); +lean_ctor_set(x_112, 2, x_4); +lean_ctor_set(x_112, 3, x_111); +x_7 = x_112; +goto block_21; +} +block_120: +{ +uint8_t x_117; lean_object* x_118; lean_object* x_119; +x_117 = lean_ctor_get_uint8(x_115, 2); +lean_dec_ref(x_115); +lean_inc(x_2); +x_118 = lean_alloc_ctor(6, 3, 2); +lean_ctor_set(x_118, 0, x_2); +lean_ctor_set(x_118, 1, x_114); +lean_ctor_set(x_118, 2, x_111); +lean_ctor_set_uint8(x_118, sizeof(void*)*3, x_116); +lean_ctor_set_uint8(x_118, sizeof(void*)*3 + 1, x_117); lean_inc_ref(x_4); lean_inc(x_2); -x_85 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_85, 0, x_2); -lean_ctor_set(x_85, 1, x_3); -lean_ctor_set(x_85, 2, x_4); -lean_ctor_set(x_85, 3, x_84); -x_7 = x_85; +x_119 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_119, 0, x_2); +lean_ctor_set(x_119, 1, x_3); +lean_ctor_set(x_119, 2, x_4); +lean_ctor_set(x_119, 3, x_118); +x_7 = x_119; goto block_21; } +block_126: +{ +if (x_121 == 0) +{ +goto block_113; +} +else +{ +lean_object* x_122; uint8_t x_123; lean_object* x_124; +x_122 = l_Lean_IR_ExplicitRC_getVarInfo(x_1, x_2); +x_123 = lean_ctor_get_uint8(x_122, 1); +x_124 = lean_unsigned_to_nat(1u); +if (x_123 == 0) +{ +x_114 = x_124; +x_115 = x_122; +x_116 = x_121; +goto block_120; +} +else +{ +uint8_t x_125; +x_125 = 0; +x_114 = x_124; +x_115 = x_122; +x_116 = x_125; +goto block_120; +} +} +} +} default: { -lean_object* x_86; +lean_object* x_130; lean_inc_ref(x_4); lean_inc(x_2); -x_86 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_86, 0, x_2); -lean_ctor_set(x_86, 1, x_3); -lean_ctor_set(x_86, 2, x_4); -lean_ctor_set(x_86, 3, x_22); -x_7 = x_86; +x_130 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_130, 0, x_2); +lean_ctor_set(x_130, 1, x_3); +lean_ctor_set(x_130, 2, x_4); +lean_ctor_set(x_130, 3, x_22); +x_7 = x_130; goto block_21; } } @@ -10936,9 +11193,9 @@ goto block_152; block_144: { lean_object* x_141; lean_object* x_142; lean_object* x_143; -x_141 = lean_nat_add(x_138, x_140); +x_141 = lean_nat_add(x_139, x_140); lean_dec(x_140); -lean_dec(x_138); +lean_dec(x_139); if (lean_is_scalar(x_135)) { x_142 = lean_alloc_ctor(0, 5, 0); } else { @@ -10957,7 +11214,7 @@ if (lean_is_scalar(x_125)) { lean_ctor_set(x_143, 0, x_137); lean_ctor_set(x_143, 1, x_128); lean_ctor_set(x_143, 2, x_129); -lean_ctor_set(x_143, 3, x_139); +lean_ctor_set(x_143, 3, x_138); lean_ctor_set(x_143, 4, x_142); return x_143; } @@ -10983,8 +11240,8 @@ if (lean_obj_tag(x_131) == 0) lean_object* x_150; x_150 = lean_ctor_get(x_131, 0); lean_inc(x_150); -x_138 = x_149; -x_139 = x_148; +x_138 = x_148; +x_139 = x_149; x_140 = x_150; goto block_144; } @@ -10992,8 +11249,8 @@ else { lean_object* x_151; x_151 = lean_unsigned_to_nat(0u); -x_138 = x_149; -x_139 = x_148; +x_138 = x_148; +x_139 = x_149; x_140 = x_151; goto block_144; } @@ -11519,7 +11776,7 @@ static lean_object* _init_l_Lean_IR_ExplicitRC_visitFnBody___closed__3() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_IR_ExplicitRC_visitFnBody___closed__2; x_2 = lean_unsigned_to_nat(58u); -x_3 = lean_unsigned_to_nat(460u); +x_3 = lean_unsigned_to_nat(473u); x_4 = l_Lean_IR_ExplicitRC_visitFnBody___closed__1; x_5 = l_Lean_IR_ExplicitRC_visitFnBody___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -11788,7 +12045,7 @@ if (x_81 == 0) { lean_object* x_82; lean_object* x_83; lean_object* x_84; uint8_t x_85; x_82 = lean_ctor_get(x_1, 0); -x_83 = lean_ctor_get(x_1, 3); +x_83 = lean_ctor_get(x_1, 4); lean_inc_ref(x_2); x_84 = l_Lean_IR_ExplicitRC_visitFnBody(x_83, x_2); x_85 = !lean_is_exclusive(x_84); @@ -11800,7 +12057,7 @@ x_87 = lean_ctor_get(x_84, 1); lean_inc(x_82); x_88 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useVar___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useExpr_spec__0(x_2, x_82, x_87); lean_dec_ref(x_2); -lean_ctor_set(x_1, 3, x_86); +lean_ctor_set(x_1, 4, x_86); lean_ctor_set(x_84, 1, x_88); lean_ctor_set(x_84, 0, x_1); return x_84; @@ -11816,7 +12073,7 @@ lean_dec(x_84); lean_inc(x_82); x_91 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useVar___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useExpr_spec__0(x_2, x_82, x_90); lean_dec_ref(x_2); -lean_ctor_set(x_1, 3, x_89); +lean_ctor_set(x_1, 4, x_89); x_92 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_92, 0, x_1); lean_ctor_set(x_92, 1, x_91); @@ -11825,139 +12082,145 @@ return x_92; } else { -lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; x_93 = lean_ctor_get(x_1, 0); x_94 = lean_ctor_get(x_1, 1); x_95 = lean_ctor_get(x_1, 2); x_96 = lean_ctor_get(x_1, 3); +x_97 = lean_ctor_get(x_1, 4); +lean_inc(x_97); lean_inc(x_96); lean_inc(x_95); lean_inc(x_94); lean_inc(x_93); lean_dec(x_1); lean_inc_ref(x_2); -x_97 = l_Lean_IR_ExplicitRC_visitFnBody(x_96, x_2); -x_98 = lean_ctor_get(x_97, 0); -lean_inc(x_98); -x_99 = lean_ctor_get(x_97, 1); +x_98 = l_Lean_IR_ExplicitRC_visitFnBody(x_97, x_2); +x_99 = lean_ctor_get(x_98, 0); lean_inc(x_99); -if (lean_is_exclusive(x_97)) { - lean_ctor_release(x_97, 0); - lean_ctor_release(x_97, 1); - x_100 = x_97; +x_100 = lean_ctor_get(x_98, 1); +lean_inc(x_100); +if (lean_is_exclusive(x_98)) { + lean_ctor_release(x_98, 0); + lean_ctor_release(x_98, 1); + x_101 = x_98; } else { - lean_dec_ref(x_97); - x_100 = lean_box(0); + lean_dec_ref(x_98); + x_101 = lean_box(0); } lean_inc(x_93); -x_101 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useVar___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useExpr_spec__0(x_2, x_93, x_99); +x_102 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useVar___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useExpr_spec__0(x_2, x_93, x_100); lean_dec_ref(x_2); -x_102 = lean_alloc_ctor(4, 4, 0); -lean_ctor_set(x_102, 0, x_93); -lean_ctor_set(x_102, 1, x_94); -lean_ctor_set(x_102, 2, x_95); -lean_ctor_set(x_102, 3, x_98); -if (lean_is_scalar(x_100)) { - x_103 = lean_alloc_ctor(0, 2, 0); +x_103 = lean_alloc_ctor(4, 5, 0); +lean_ctor_set(x_103, 0, x_93); +lean_ctor_set(x_103, 1, x_94); +lean_ctor_set(x_103, 2, x_95); +lean_ctor_set(x_103, 3, x_96); +lean_ctor_set(x_103, 4, x_99); +if (lean_is_scalar(x_101)) { + x_104 = lean_alloc_ctor(0, 2, 0); } else { - x_103 = x_100; + x_104 = x_101; } -lean_ctor_set(x_103, 0, x_102); -lean_ctor_set(x_103, 1, x_101); -return x_103; +lean_ctor_set(x_104, 0, x_103); +lean_ctor_set(x_104, 1, x_102); +return x_104; } } case 5: { -uint8_t x_104; -x_104 = !lean_is_exclusive(x_1); -if (x_104 == 0) +uint8_t x_105; +x_105 = !lean_is_exclusive(x_1); +if (x_105 == 0) { -lean_object* x_105; lean_object* x_106; lean_object* x_107; uint8_t x_108; -x_105 = lean_ctor_get(x_1, 0); -x_106 = lean_ctor_get(x_1, 5); +lean_object* x_106; lean_object* x_107; lean_object* x_108; uint8_t x_109; +x_106 = lean_ctor_get(x_1, 0); +x_107 = lean_ctor_get(x_1, 6); lean_inc_ref(x_2); -x_107 = l_Lean_IR_ExplicitRC_visitFnBody(x_106, x_2); -x_108 = !lean_is_exclusive(x_107); -if (x_108 == 0) +x_108 = l_Lean_IR_ExplicitRC_visitFnBody(x_107, x_2); +x_109 = !lean_is_exclusive(x_108); +if (x_109 == 0) { -lean_object* x_109; lean_object* x_110; lean_object* x_111; -x_109 = lean_ctor_get(x_107, 0); -x_110 = lean_ctor_get(x_107, 1); -lean_inc(x_105); -x_111 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useVar___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useExpr_spec__0(x_2, x_105, x_110); +lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_110 = lean_ctor_get(x_108, 0); +x_111 = lean_ctor_get(x_108, 1); +lean_inc(x_106); +x_112 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useVar___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useExpr_spec__0(x_2, x_106, x_111); lean_dec_ref(x_2); -lean_ctor_set(x_1, 5, x_109); -lean_ctor_set(x_107, 1, x_111); -lean_ctor_set(x_107, 0, x_1); -return x_107; +lean_ctor_set(x_1, 6, x_110); +lean_ctor_set(x_108, 1, x_112); +lean_ctor_set(x_108, 0, x_1); +return x_108; } else { -lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; -x_112 = lean_ctor_get(x_107, 0); -x_113 = lean_ctor_get(x_107, 1); +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_113 = lean_ctor_get(x_108, 0); +x_114 = lean_ctor_get(x_108, 1); +lean_inc(x_114); lean_inc(x_113); -lean_inc(x_112); -lean_dec(x_107); -lean_inc(x_105); -x_114 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useVar___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useExpr_spec__0(x_2, x_105, x_113); +lean_dec(x_108); +lean_inc(x_106); +x_115 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useVar___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useExpr_spec__0(x_2, x_106, x_114); lean_dec_ref(x_2); -lean_ctor_set(x_1, 5, x_112); -x_115 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_115, 0, x_1); -lean_ctor_set(x_115, 1, x_114); -return x_115; +lean_ctor_set(x_1, 6, x_113); +x_116 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_116, 0, x_1); +lean_ctor_set(x_116, 1, x_115); +return x_116; } } else { -lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; -x_116 = lean_ctor_get(x_1, 0); -x_117 = lean_ctor_get(x_1, 1); -x_118 = lean_ctor_get(x_1, 2); -x_119 = lean_ctor_get(x_1, 3); -x_120 = lean_ctor_get(x_1, 4); -x_121 = lean_ctor_get(x_1, 5); +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; +x_117 = lean_ctor_get(x_1, 0); +x_118 = lean_ctor_get(x_1, 1); +x_119 = lean_ctor_get(x_1, 2); +x_120 = lean_ctor_get(x_1, 3); +x_121 = lean_ctor_get(x_1, 4); +x_122 = lean_ctor_get(x_1, 5); +x_123 = lean_ctor_get(x_1, 6); +lean_inc(x_123); +lean_inc(x_122); lean_inc(x_121); lean_inc(x_120); lean_inc(x_119); lean_inc(x_118); lean_inc(x_117); -lean_inc(x_116); lean_dec(x_1); lean_inc_ref(x_2); -x_122 = l_Lean_IR_ExplicitRC_visitFnBody(x_121, x_2); -x_123 = lean_ctor_get(x_122, 0); -lean_inc(x_123); -x_124 = lean_ctor_get(x_122, 1); -lean_inc(x_124); -if (lean_is_exclusive(x_122)) { - lean_ctor_release(x_122, 0); - lean_ctor_release(x_122, 1); - x_125 = x_122; +x_124 = l_Lean_IR_ExplicitRC_visitFnBody(x_123, x_2); +x_125 = lean_ctor_get(x_124, 0); +lean_inc(x_125); +x_126 = lean_ctor_get(x_124, 1); +lean_inc(x_126); +if (lean_is_exclusive(x_124)) { + lean_ctor_release(x_124, 0); + lean_ctor_release(x_124, 1); + x_127 = x_124; } else { - lean_dec_ref(x_122); - x_125 = lean_box(0); + lean_dec_ref(x_124); + x_127 = lean_box(0); } -lean_inc(x_116); -x_126 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useVar___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useExpr_spec__0(x_2, x_116, x_124); +lean_inc(x_117); +x_128 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useVar___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useExpr_spec__0(x_2, x_117, x_126); lean_dec_ref(x_2); -x_127 = lean_alloc_ctor(5, 6, 0); -lean_ctor_set(x_127, 0, x_116); -lean_ctor_set(x_127, 1, x_117); -lean_ctor_set(x_127, 2, x_118); -lean_ctor_set(x_127, 3, x_119); -lean_ctor_set(x_127, 4, x_120); -lean_ctor_set(x_127, 5, x_123); -if (lean_is_scalar(x_125)) { - x_128 = lean_alloc_ctor(0, 2, 0); -} else { - x_128 = x_125; -} -lean_ctor_set(x_128, 0, x_127); -lean_ctor_set(x_128, 1, x_126); -return x_128; +x_129 = lean_alloc_ctor(5, 7, 0); +lean_ctor_set(x_129, 0, x_117); +lean_ctor_set(x_129, 1, x_118); +lean_ctor_set(x_129, 2, x_119); +lean_ctor_set(x_129, 3, x_120); +lean_ctor_set(x_129, 4, x_121); +lean_ctor_set(x_129, 5, x_122); +lean_ctor_set(x_129, 6, x_125); +if (lean_is_scalar(x_127)) { + x_130 = lean_alloc_ctor(0, 2, 0); +} else { + x_130 = x_127; +} +lean_ctor_set(x_130, 0, x_129); +lean_ctor_set(x_130, 1, x_128); +return x_130; } } case 6: @@ -11974,216 +12237,216 @@ goto block_5; } case 9: { -lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; size_t x_134; size_t x_135; lean_object* x_136; lean_object* x_137; lean_object* x_144; lean_object* x_145; lean_object* x_146; uint8_t x_147; -x_129 = lean_ctor_get(x_1, 0); -lean_inc(x_129); -x_130 = lean_ctor_get(x_1, 1); -lean_inc(x_130); -x_131 = lean_ctor_get(x_1, 2); +lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; size_t x_136; size_t x_137; lean_object* x_138; lean_object* x_139; lean_object* x_146; lean_object* x_147; lean_object* x_148; uint8_t x_149; +x_131 = lean_ctor_get(x_1, 0); lean_inc(x_131); -x_132 = lean_ctor_get(x_1, 3); -lean_inc_ref(x_132); +x_132 = lean_ctor_get(x_1, 1); +lean_inc(x_132); +x_133 = lean_ctor_get(x_1, 2); +lean_inc(x_133); +x_134 = lean_ctor_get(x_1, 3); +lean_inc_ref(x_134); if (lean_is_exclusive(x_1)) { lean_ctor_release(x_1, 0); lean_ctor_release(x_1, 1); lean_ctor_release(x_1, 2); lean_ctor_release(x_1, 3); - x_133 = x_1; + x_135 = x_1; } else { lean_dec_ref(x_1); - x_133 = lean_box(0); + x_135 = lean_box(0); } -x_134 = lean_array_size(x_132); -x_135 = 0; -lean_inc(x_130); +x_136 = lean_array_size(x_134); +x_137 = 0; +lean_inc(x_132); lean_inc_ref(x_2); -x_136 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitRC_visitFnBody_spec__4(x_2, x_130, x_134, x_135, x_132); -x_144 = l_Lean_IR_ExplicitRC_instInhabitedLiveVars_default___closed__0; -x_145 = lean_unsigned_to_nat(0u); -x_146 = lean_array_get_size(x_136); -x_147 = lean_nat_dec_lt(x_145, x_146); -if (x_147 == 0) +x_138 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitRC_visitFnBody_spec__4(x_2, x_132, x_136, x_137, x_134); +x_146 = l_Lean_IR_ExplicitRC_instInhabitedLiveVars_default___closed__0; +x_147 = lean_unsigned_to_nat(0u); +x_148 = lean_array_get_size(x_138); +x_149 = lean_nat_dec_lt(x_147, x_148); +if (x_149 == 0) { -x_137 = x_144; -goto block_143; +x_139 = x_146; +goto block_145; } else { -uint8_t x_148; -x_148 = lean_nat_dec_le(x_146, x_146); -if (x_148 == 0) +uint8_t x_150; +x_150 = lean_nat_dec_le(x_148, x_148); +if (x_150 == 0) { -if (x_147 == 0) +if (x_149 == 0) { -x_137 = x_144; -goto block_143; +x_139 = x_146; +goto block_145; } else { -size_t x_149; lean_object* x_150; -x_149 = lean_usize_of_nat(x_146); -x_150 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ExplicitRC_visitFnBody_spec__6(x_136, x_135, x_149, x_144); -x_137 = x_150; -goto block_143; +size_t x_151; lean_object* x_152; +x_151 = lean_usize_of_nat(x_148); +x_152 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ExplicitRC_visitFnBody_spec__6(x_138, x_137, x_151, x_146); +x_139 = x_152; +goto block_145; } } else { -size_t x_151; lean_object* x_152; -x_151 = lean_usize_of_nat(x_146); -x_152 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ExplicitRC_visitFnBody_spec__6(x_136, x_135, x_151, x_144); -x_137 = x_152; -goto block_143; +size_t x_153; lean_object* x_154; +x_153 = lean_usize_of_nat(x_148); +x_154 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ExplicitRC_visitFnBody_spec__6(x_138, x_137, x_153, x_146); +x_139 = x_154; +goto block_145; } } -block_143: +block_145: { -lean_object* x_138; size_t x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; -lean_inc(x_130); -x_138 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useVar___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useExpr_spec__0(x_2, x_130, x_137); -x_139 = lean_array_size(x_136); -lean_inc(x_130); -x_140 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitRC_visitFnBody_spec__5(x_2, x_130, x_138, x_139, x_135, x_136); -if (lean_is_scalar(x_133)) { - x_141 = lean_alloc_ctor(9, 4, 0); +lean_object* x_140; size_t x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; +lean_inc(x_132); +x_140 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useVar___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useExpr_spec__0(x_2, x_132, x_139); +x_141 = lean_array_size(x_138); +lean_inc(x_132); +x_142 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitRC_visitFnBody_spec__5(x_2, x_132, x_140, x_141, x_137, x_138); +if (lean_is_scalar(x_135)) { + x_143 = lean_alloc_ctor(9, 4, 0); } else { - x_141 = x_133; + x_143 = x_135; } -lean_ctor_set(x_141, 0, x_129); -lean_ctor_set(x_141, 1, x_130); -lean_ctor_set(x_141, 2, x_131); -lean_ctor_set(x_141, 3, x_140); -x_142 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_142, 0, x_141); -lean_ctor_set(x_142, 1, x_138); -return x_142; +lean_ctor_set(x_143, 0, x_131); +lean_ctor_set(x_143, 1, x_132); +lean_ctor_set(x_143, 2, x_133); +lean_ctor_set(x_143, 3, x_142); +x_144 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_144, 0, x_143); +lean_ctor_set(x_144, 1, x_140); +return x_144; } } case 10: { -lean_object* x_153; lean_object* x_154; -x_153 = lean_ctor_get(x_1, 0); +lean_object* x_155; lean_object* x_156; +x_155 = lean_ctor_get(x_1, 0); lean_inc_ref(x_2); -x_154 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_mkRetLiveVars(x_2); -if (lean_obj_tag(x_153) == 0) -{ -lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; uint8_t x_159; lean_object* x_160; -x_155 = lean_ctor_get(x_153, 0); -x_156 = lean_ctor_get(x_2, 4); -x_157 = l_Lean_IR_ExplicitRC_instInhabitedVarInfo_default; -x_158 = l_Std_DTreeMap_Internal_Impl_Const_get_x21___at___00Lean_IR_ExplicitRC_getVarInfo_spec__0___redArg(x_157, x_156, x_155); -x_159 = lean_ctor_get_uint8(x_158, 0); -lean_dec(x_158); -lean_inc(x_155); -x_160 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useVar___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useExpr_spec__0(x_2, x_155, x_154); -if (x_159 == 0) -{ -lean_object* x_161; +x_156 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_mkRetLiveVars(x_2); +if (lean_obj_tag(x_155) == 0) +{ +lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; uint8_t x_161; lean_object* x_162; +x_157 = lean_ctor_get(x_155, 0); +x_158 = lean_ctor_get(x_2, 4); +x_159 = l_Lean_IR_ExplicitRC_instInhabitedVarInfo_default; +x_160 = l_Std_DTreeMap_Internal_Impl_Const_get_x21___at___00Lean_IR_ExplicitRC_getVarInfo_spec__0___redArg(x_159, x_158, x_157); +x_161 = lean_ctor_get_uint8(x_160, 0); +lean_dec(x_160); +lean_inc(x_157); +x_162 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useVar___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useExpr_spec__0(x_2, x_157, x_156); +if (x_161 == 0) +{ +lean_object* x_163; lean_dec_ref(x_2); -x_161 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_161, 0, x_1); -lean_ctor_set(x_161, 1, x_160); -return x_161; +x_163 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_163, 0, x_1); +lean_ctor_set(x_163, 1, x_162); +return x_163; } else { -lean_object* x_162; uint8_t x_163; -x_162 = lean_ctor_get(x_160, 1); -lean_inc(x_162); -x_163 = l_Std_DTreeMap_Internal_Impl_contains___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_visitParam_spec__1___redArg(x_155, x_162); -lean_dec(x_162); -if (x_163 == 0) +lean_object* x_164; uint8_t x_165; +x_164 = lean_ctor_get(x_162, 1); +lean_inc(x_164); +x_165 = l_Std_DTreeMap_Internal_Impl_contains___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_visitParam_spec__1___redArg(x_157, x_164); +lean_dec(x_164); +if (x_165 == 0) { -lean_object* x_164; +lean_object* x_166; lean_dec_ref(x_2); -x_164 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_164, 0, x_1); -lean_ctor_set(x_164, 1, x_160); -return x_164; +x_166 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_166, 0, x_1); +lean_ctor_set(x_166, 1, x_162); +return x_166; } else { -lean_object* x_165; uint8_t x_166; lean_object* x_167; uint8_t x_168; -lean_inc(x_155); -x_165 = l_Lean_IR_ExplicitRC_getVarInfo(x_2, x_155); +lean_object* x_167; uint8_t x_168; lean_object* x_169; uint8_t x_170; +lean_inc(x_157); +x_167 = l_Lean_IR_ExplicitRC_getVarInfo(x_2, x_157); lean_dec_ref(x_2); -x_166 = lean_ctor_get_uint8(x_165, 1); -x_167 = lean_unsigned_to_nat(1u); -if (x_166 == 0) +x_168 = lean_ctor_get_uint8(x_167, 1); +x_169 = lean_unsigned_to_nat(1u); +if (x_168 == 0) { -x_168 = x_163; -goto block_172; +x_170 = x_165; +goto block_174; } else { -uint8_t x_173; -x_173 = 0; -x_168 = x_173; -goto block_172; +uint8_t x_175; +x_175 = 0; +x_170 = x_175; +goto block_174; } -block_172: +block_174: { -uint8_t x_169; lean_object* x_170; lean_object* x_171; -x_169 = lean_ctor_get_uint8(x_165, 2); -lean_dec_ref(x_165); -x_170 = lean_alloc_ctor(6, 3, 2); -lean_ctor_set(x_170, 0, x_155); -lean_ctor_set(x_170, 1, x_167); -lean_ctor_set(x_170, 2, x_1); -lean_ctor_set_uint8(x_170, sizeof(void*)*3, x_168); -lean_ctor_set_uint8(x_170, sizeof(void*)*3 + 1, x_169); -x_171 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_171, 0, x_170); -lean_ctor_set(x_171, 1, x_160); -return x_171; +uint8_t x_171; lean_object* x_172; lean_object* x_173; +x_171 = lean_ctor_get_uint8(x_167, 2); +lean_dec_ref(x_167); +x_172 = lean_alloc_ctor(6, 3, 2); +lean_ctor_set(x_172, 0, x_157); +lean_ctor_set(x_172, 1, x_169); +lean_ctor_set(x_172, 2, x_1); +lean_ctor_set_uint8(x_172, sizeof(void*)*3, x_170); +lean_ctor_set_uint8(x_172, sizeof(void*)*3 + 1, x_171); +x_173 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_173, 0, x_172); +lean_ctor_set(x_173, 1, x_162); +return x_173; } } } } else { -lean_object* x_174; +lean_object* x_176; lean_dec_ref(x_2); -x_174 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_174, 0, x_1); -lean_ctor_set(x_174, 1, x_154); -return x_174; +x_176 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_176, 0, x_1); +lean_ctor_set(x_176, 1, x_156); +return x_176; } } case 11: { -lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; -x_175 = lean_ctor_get(x_1, 0); -x_176 = lean_ctor_get(x_1, 1); -lean_inc_ref(x_176); -x_177 = l_Lean_IR_ExplicitRC_getJPLiveVars(x_2, x_175); -x_178 = l_Lean_IR_ExplicitRC_getJPParams(x_2, x_175); -x_179 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_addIncBefore(x_2, x_176, x_178, x_1, x_177); -x_180 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useArgs(x_2, x_176, x_177); -lean_dec_ref(x_176); +lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; +x_177 = lean_ctor_get(x_1, 0); +x_178 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_178); +x_179 = l_Lean_IR_ExplicitRC_getJPLiveVars(x_2, x_177); +x_180 = l_Lean_IR_ExplicitRC_getJPParams(x_2, x_177); +x_181 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_addIncBefore(x_2, x_178, x_180, x_1, x_179); +x_182 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useArgs(x_2, x_178, x_179); +lean_dec_ref(x_178); lean_dec_ref(x_2); -x_181 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_181, 0, x_179); -lean_ctor_set(x_181, 1, x_180); -return x_181; -} -case 12: -{ -lean_object* x_182; lean_object* x_183; -x_182 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_mkRetLiveVars(x_2); x_183 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_183, 0, x_1); +lean_ctor_set(x_183, 0, x_181); lean_ctor_set(x_183, 1, x_182); return x_183; } -default: +case 12: { lean_object* x_184; lean_object* x_185; +x_184 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_mkRetLiveVars(x_2); +x_185 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_185, 0, x_1); +lean_ctor_set(x_185, 1, x_184); +return x_185; +} +default: +{ +lean_object* x_186; lean_object* x_187; lean_dec_ref(x_2); lean_dec(x_1); -x_184 = l_Lean_IR_ExplicitRC_visitFnBody___closed__3; -x_185 = l_panic___at___00Lean_IR_ExplicitRC_visitFnBody_spec__2(x_184); -return x_185; +x_186 = l_Lean_IR_ExplicitRC_visitFnBody___closed__3; +x_187 = l_panic___at___00Lean_IR_ExplicitRC_visitFnBody_spec__2(x_186); +return x_187; } } block_5: diff --git a/stage0/stdlib/Lean/Compiler/IR/ResetReuse.c b/stage0/stdlib/Lean/Compiler/IR/ResetReuse.c index b2c38414f5e9..946de84ec61e 100644 --- a/stage0/stdlib/Lean/Compiler/IR/ResetReuse.c +++ b/stage0/stdlib/Lean/Compiler/IR/ResetReuse.c @@ -65,7 +65,6 @@ static lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_initFn___ LEAN_EXPORT lean_object* l_Lean_IR_ResetReuse_UseClassification_none_elim___redArg(lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___00Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_IR_ResetReuse_R_spec__1_spec__2_spec__4_spec__6___redArg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_initFn___closed__19_00___x40_Lean_Compiler_IR_ResetReuse_1484512311____hygCtx___hyg_2_; static lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_initFn___closed__5_00___x40_Lean_Compiler_IR_ResetReuse_1484512311____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_IR_ResetReuse_R_spec__1_spec__2___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -85,6 +84,7 @@ lean_object* l_Lean_Name_getPrefix(lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_isCtorUsing(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at___00Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_IR_ResetReuse_R_spec__0_spec__0_spec__1___redArg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___00Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_IR_ResetReuse_R_spec__1_spec__2_spec__4_spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain_spec__0___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ResetReuse_UseClassification_toCtorIdx(uint8_t); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_mkFresh___boxed(lean_object*, lean_object*); @@ -93,6 +93,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_Rese LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_IR_ResetReuse_R_spec__1_spec__2_spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_initFn___closed__8_00___x40_Lean_Compiler_IR_ResetReuse_1484512311____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_IR_ResetReuse_R_spec__1_spec__2_spec__4___redArg(lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_IR_IRType_isObj(lean_object*); lean_object* l_Lean_IR_FnBody_body(lean_object*); static lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_initFn___closed__24_00___x40_Lean_Compiler_IR_ResetReuse_1484512311____hygCtx___hyg_2_; static size_t l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_IR_ResetReuse_R_spec__0_spec__0___redArg___closed__1; @@ -115,6 +116,7 @@ static lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_initFn___ LEAN_EXPORT lean_object* l_Lean_IR_ResetReuse_UseClassification_none_elim___redArg___boxed(lean_object*); LEAN_EXPORT uint8_t l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00__private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_argsContainsVar_spec__0(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dfinalize___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___00Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_IR_ResetReuse_R_spec__1_spec__2_spec__4_spec__7___redArg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_usize_to_nat(size_t); static lean_object* l_Lean_PersistentHashMap_empty___at___00Lean_IR_Decl_insertResetReuseCore_spec__0___closed__1; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ResetReuse_collectResets_spec__0(lean_object*, size_t, size_t, lean_object*, lean_object*); @@ -144,7 +146,6 @@ static lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_initFn___ static lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_initFn___closed__16_00___x40_Lean_Compiler_IR_ResetReuse_1484512311____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_IR_ResetReuse_R_spec__1_spec__2_spec__5___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___00Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_IR_ResetReuse_R_spec__1_spec__2_spec__4_spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_IR_Alt_body(lean_object*); lean_object* l_Array_toSubarray___redArg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_initFn_00___x40_Lean_Compiler_IR_ResetReuse_1484512311____hygCtx___hyg_2_(); @@ -159,6 +160,7 @@ static lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_initFn___ LEAN_EXPORT lean_object* l_Lean_IR_ResetReuse_UseClassification_ownedArg_elim___redArg(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ResetReuse_UseClassification_ctorElim___redArg___boxed(lean_object*); size_t lean_usize_sub(size_t, size_t); +uint8_t l_Lean_IR_IRType_isStruct(lean_object*); lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___redArg(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ResetReuse_UseClassification_none_elim___boxed(lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); @@ -182,10 +184,12 @@ LEAN_EXPORT lean_object* l_Lean_IR_ResetReuse_UseClassification_other_elim___red uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S(lean_object*, lean_object*, uint8_t, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ResetReuse_R_spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t l_Lean_IR_instHashableVarId_hash(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ResetReuse_UseClassification_ctorElim___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_mayReuse(lean_object*, lean_object*, uint8_t); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ResetReuse_R_spec__3(size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ResetReuse_UseClassification_other_elim___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); static lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_initFn___closed__6_00___x40_Lean_Compiler_IR_ResetReuse_1484512311____hygCtx___hyg_2_; @@ -277,7 +281,7 @@ lean_object* x_24; x_24 = lean_ctor_get(x_4, 2); if (lean_obj_tag(x_24) == 0) { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; uint8_t x_35; +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; uint8_t x_35; uint8_t x_43; lean_inc_ref(x_24); x_25 = lean_ctor_get(x_4, 0); lean_inc(x_25); @@ -297,7 +301,40 @@ if (lean_is_exclusive(x_4)) { } x_29 = lean_ctor_get(x_24, 0); x_30 = lean_ctor_get(x_24, 1); -x_35 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_mayReuse(x_2, x_29, x_3); +x_43 = l_Lean_IR_IRType_isObj(x_26); +if (x_43 == 0) +{ +x_35 = x_43; +goto block_42; +} +else +{ +uint8_t x_44; +x_44 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_mayReuse(x_2, x_29, x_3); +x_35 = x_44; +goto block_42; +} +block_34: +{ +lean_object* x_32; lean_object* x_33; +x_32 = lean_alloc_ctor(2, 3, 1); +lean_ctor_set(x_32, 0, x_1); +lean_ctor_set(x_32, 1, x_29); +lean_ctor_set(x_32, 2, x_30); +lean_ctor_set_uint8(x_32, sizeof(void*)*3, x_31); +if (lean_is_scalar(x_28)) { + x_33 = lean_alloc_ctor(0, 4, 0); +} else { + x_33 = x_28; +} +lean_ctor_set(x_33, 0, x_25); +lean_ctor_set(x_33, 1, x_26); +lean_ctor_set(x_33, 2, x_32); +lean_ctor_set(x_33, 3, x_27); +return x_33; +} +block_42: +{ if (x_35 == 0) { lean_object* x_36; lean_object* x_37; @@ -332,24 +369,6 @@ x_31 = x_41; goto block_34; } } -block_34: -{ -lean_object* x_32; lean_object* x_33; -x_32 = lean_alloc_ctor(2, 3, 1); -lean_ctor_set(x_32, 0, x_1); -lean_ctor_set(x_32, 1, x_29); -lean_ctor_set(x_32, 2, x_30); -lean_ctor_set_uint8(x_32, sizeof(void*)*3, x_31); -if (lean_is_scalar(x_28)) { - x_33 = lean_alloc_ctor(0, 4, 0); -} else { - x_33 = x_28; -} -lean_ctor_set(x_33, 0, x_25); -lean_ctor_set(x_33, 1, x_26); -lean_ctor_set(x_33, 2, x_32); -lean_ctor_set(x_33, 3, x_27); -return x_33; } } else @@ -360,114 +379,114 @@ goto block_23; } case 1: { -uint8_t x_42; -x_42 = !lean_is_exclusive(x_4); -if (x_42 == 0) +uint8_t x_45; +x_45 = !lean_is_exclusive(x_4); +if (x_45 == 0) { -lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; -x_43 = lean_ctor_get(x_4, 2); -x_44 = lean_ctor_get(x_4, 3); -lean_inc(x_43); +lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; +x_46 = lean_ctor_get(x_4, 2); +x_47 = lean_ctor_get(x_4, 3); +lean_inc(x_46); lean_inc(x_1); -x_45 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go(x_1, x_2, x_3, x_43); -lean_inc(x_45); -lean_inc(x_43); -x_46 = l_Lean_IR_FnBody_beq(x_43, x_45); -if (x_46 == 0) +x_48 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go(x_1, x_2, x_3, x_46); +lean_inc(x_48); +lean_inc(x_46); +x_49 = l_Lean_IR_FnBody_beq(x_46, x_48); +if (x_49 == 0) { -lean_dec(x_43); +lean_dec(x_46); lean_dec(x_1); -lean_ctor_set(x_4, 2, x_45); +lean_ctor_set(x_4, 2, x_48); return x_4; } else { -lean_object* x_47; -lean_dec(x_45); -x_47 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go(x_1, x_2, x_3, x_44); -lean_ctor_set(x_4, 3, x_47); +lean_object* x_50; +lean_dec(x_48); +x_50 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go(x_1, x_2, x_3, x_47); +lean_ctor_set(x_4, 3, x_50); return x_4; } } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; -x_48 = lean_ctor_get(x_4, 0); -x_49 = lean_ctor_get(x_4, 1); -x_50 = lean_ctor_get(x_4, 2); -x_51 = lean_ctor_get(x_4, 3); +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; +x_51 = lean_ctor_get(x_4, 0); +x_52 = lean_ctor_get(x_4, 1); +x_53 = lean_ctor_get(x_4, 2); +x_54 = lean_ctor_get(x_4, 3); +lean_inc(x_54); +lean_inc(x_53); +lean_inc(x_52); lean_inc(x_51); -lean_inc(x_50); -lean_inc(x_49); -lean_inc(x_48); lean_dec(x_4); -lean_inc(x_50); +lean_inc(x_53); lean_inc(x_1); -x_52 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go(x_1, x_2, x_3, x_50); -lean_inc(x_52); -lean_inc(x_50); -x_53 = l_Lean_IR_FnBody_beq(x_50, x_52); -if (x_53 == 0) +x_55 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go(x_1, x_2, x_3, x_53); +lean_inc(x_55); +lean_inc(x_53); +x_56 = l_Lean_IR_FnBody_beq(x_53, x_55); +if (x_56 == 0) { -lean_object* x_54; -lean_dec(x_50); +lean_object* x_57; +lean_dec(x_53); lean_dec(x_1); -x_54 = lean_alloc_ctor(1, 4, 0); -lean_ctor_set(x_54, 0, x_48); -lean_ctor_set(x_54, 1, x_49); -lean_ctor_set(x_54, 2, x_52); -lean_ctor_set(x_54, 3, x_51); -return x_54; +x_57 = lean_alloc_ctor(1, 4, 0); +lean_ctor_set(x_57, 0, x_51); +lean_ctor_set(x_57, 1, x_52); +lean_ctor_set(x_57, 2, x_55); +lean_ctor_set(x_57, 3, x_54); +return x_57; } else { -lean_object* x_55; lean_object* x_56; -lean_dec(x_52); -x_55 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go(x_1, x_2, x_3, x_51); -x_56 = lean_alloc_ctor(1, 4, 0); -lean_ctor_set(x_56, 0, x_48); -lean_ctor_set(x_56, 1, x_49); -lean_ctor_set(x_56, 2, x_50); -lean_ctor_set(x_56, 3, x_55); -return x_56; +lean_object* x_58; lean_object* x_59; +lean_dec(x_55); +x_58 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go(x_1, x_2, x_3, x_54); +x_59 = lean_alloc_ctor(1, 4, 0); +lean_ctor_set(x_59, 0, x_51); +lean_ctor_set(x_59, 1, x_52); +lean_ctor_set(x_59, 2, x_53); +lean_ctor_set(x_59, 3, x_58); +return x_59; } } } case 9: { -uint8_t x_57; -x_57 = !lean_is_exclusive(x_4); -if (x_57 == 0) +uint8_t x_60; +x_60 = !lean_is_exclusive(x_4); +if (x_60 == 0) { -lean_object* x_58; size_t x_59; size_t x_60; lean_object* x_61; -x_58 = lean_ctor_get(x_4, 3); -x_59 = lean_array_size(x_58); -x_60 = 0; -x_61 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go_spec__0(x_1, x_2, x_3, x_59, x_60, x_58); -lean_ctor_set(x_4, 3, x_61); +lean_object* x_61; size_t x_62; size_t x_63; lean_object* x_64; +x_61 = lean_ctor_get(x_4, 3); +x_62 = lean_array_size(x_61); +x_63 = 0; +x_64 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go_spec__0(x_1, x_2, x_3, x_62, x_63, x_61); +lean_ctor_set(x_4, 3, x_64); return x_4; } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; size_t x_66; size_t x_67; lean_object* x_68; lean_object* x_69; -x_62 = lean_ctor_get(x_4, 0); -x_63 = lean_ctor_get(x_4, 1); -x_64 = lean_ctor_get(x_4, 2); -x_65 = lean_ctor_get(x_4, 3); +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; size_t x_69; size_t x_70; lean_object* x_71; lean_object* x_72; +x_65 = lean_ctor_get(x_4, 0); +x_66 = lean_ctor_get(x_4, 1); +x_67 = lean_ctor_get(x_4, 2); +x_68 = lean_ctor_get(x_4, 3); +lean_inc(x_68); +lean_inc(x_67); +lean_inc(x_66); lean_inc(x_65); -lean_inc(x_64); -lean_inc(x_63); -lean_inc(x_62); lean_dec(x_4); -x_66 = lean_array_size(x_65); -x_67 = 0; -x_68 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go_spec__0(x_1, x_2, x_3, x_66, x_67, x_65); -x_69 = lean_alloc_ctor(9, 4, 0); -lean_ctor_set(x_69, 0, x_62); -lean_ctor_set(x_69, 1, x_63); -lean_ctor_set(x_69, 2, x_64); -lean_ctor_set(x_69, 3, x_68); -return x_69; +x_69 = lean_array_size(x_68); +x_70 = 0; +x_71 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go_spec__0(x_1, x_2, x_3, x_69, x_70, x_68); +x_72 = lean_alloc_ctor(9, 4, 0); +lean_ctor_set(x_72, 0, x_65); +lean_ctor_set(x_72, 1, x_66); +lean_ctor_set(x_72, 2, x_67); +lean_ctor_set(x_72, 3, x_71); +return x_72; } } default: @@ -522,7 +541,7 @@ goto block_11; case 4: { lean_object* x_17; -x_17 = lean_ctor_get(x_12, 3); +x_17 = lean_ctor_get(x_12, 4); lean_inc(x_17); x_5 = x_12; x_6 = x_17; @@ -531,7 +550,7 @@ goto block_11; case 5: { lean_object* x_18; -x_18 = lean_ctor_get(x_12, 5); +x_18 = lean_ctor_get(x_12, 6); lean_inc(x_18); x_5 = x_12; x_6 = x_18; @@ -2420,7 +2439,7 @@ goto block_199; case 4: { lean_object* x_203; -x_203 = lean_ctor_get(x_3, 3); +x_203 = lean_ctor_get(x_3, 4); lean_inc(x_203); x_99 = x_203; goto block_199; @@ -2428,7 +2447,7 @@ goto block_199; case 5: { lean_object* x_204; -x_204 = lean_ctor_get(x_3, 5); +x_204 = lean_ctor_get(x_3, 6); lean_inc(x_204); x_99 = x_204; goto block_199; @@ -3119,7 +3138,7 @@ x_5 = l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_conta return x_5; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___00Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_IR_ResetReuse_R_spec__1_spec__2_spec__4_spec__6___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___00Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_IR_ResetReuse_R_spec__1_spec__2_spec__4_spec__7___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -3226,7 +3245,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNode___at___0 { lean_object* x_4; lean_object* x_5; x_4 = lean_unsigned_to_nat(0u); -x_5 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___00Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_IR_ResetReuse_R_spec__1_spec__2_spec__4_spec__6___redArg(x_1, x_4, x_2, x_3); +x_5 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___00Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_IR_ResetReuse_R_spec__1_spec__2_spec__4_spec__7___redArg(x_1, x_4, x_2, x_3); return x_5; } } @@ -3555,11 +3574,11 @@ x_5 = l_Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHas return x_5; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___00Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_IR_ResetReuse_R_spec__1_spec__2_spec__4_spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___00Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_IR_ResetReuse_R_spec__1_spec__2_spec__4_spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; -x_6 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___00Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_IR_ResetReuse_R_spec__1_spec__2_spec__4_spec__6___redArg(x_2, x_3, x_4, x_5); +x_6 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___00Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_IR_ResetReuse_R_spec__1_spec__2_spec__4_spec__7___redArg(x_2, x_3, x_4, x_5); return x_6; } } @@ -3579,435 +3598,732 @@ x_8 = l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAu return x_8; } } -LEAN_EXPORT lean_object* l_Lean_IR_ResetReuse_R(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ResetReuse_R_spec__2(lean_object* x_1, uint8_t x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_4; -switch (lean_obj_tag(x_1)) { -case 9: +uint8_t x_8; +x_8 = lean_usize_dec_lt(x_4, x_3); +if (x_8 == 0) { -uint8_t x_18; -x_18 = !lean_is_exclusive(x_1); -if (x_18 == 0) +lean_object* x_9; +lean_dec_ref(x_6); +lean_dec(x_1); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_5); +lean_ctor_set(x_9, 1, x_7); +return x_9; +} +else { -uint8_t x_19; -x_19 = !lean_is_exclusive(x_2); -if (x_19 == 0) +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_10 = lean_array_uget(x_5, x_4); +x_11 = lean_unsigned_to_nat(0u); +x_12 = lean_array_uset(x_5, x_4, x_11); +if (lean_obj_tag(x_10) == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t 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_20 = lean_ctor_get(x_1, 1); -x_21 = lean_ctor_get(x_1, 3); -x_22 = lean_ctor_get(x_2, 2); -lean_inc_ref(x_22); -x_23 = l_Lean_PersistentHashMap_contains___at___00Lean_IR_ResetReuse_R_spec__0___redArg(x_22, x_20); -x_24 = lean_box(0); -lean_inc(x_20); -x_25 = l_Lean_PersistentHashMap_insert___at___00Lean_IR_ResetReuse_R_spec__1___redArg(x_22, x_20, x_24); -lean_ctor_set(x_2, 2, x_25); -x_26 = lean_array_size(x_21); -x_27 = 0; -lean_inc(x_20); -x_28 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ResetReuse_R_spec__2(x_20, x_23, x_26, x_27, x_21, x_2, x_3); -x_29 = !lean_is_exclusive(x_28); -if (x_29 == 0) +uint8_t x_20; +x_20 = !lean_is_exclusive(x_10); +if (x_20 == 0) { -lean_object* x_30; -x_30 = lean_ctor_get(x_28, 0); -lean_ctor_set(x_1, 3, x_30); -lean_ctor_set(x_28, 0, x_1); -return x_28; +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; uint8_t x_35; +x_21 = lean_ctor_get(x_10, 0); +x_22 = lean_ctor_get(x_10, 1); +lean_inc_ref(x_6); +x_23 = l_Lean_IR_ResetReuse_R(x_22, 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_ref(x_23); +lean_inc(x_24); +lean_inc_ref(x_21); +lean_ctor_set(x_10, 1, x_24); +x_35 = l_Lean_IR_CtorInfo_isScalar(x_21); +if (x_35 == 0) +{ +x_26 = x_2; +goto block_34; +} +else +{ +x_26 = x_35; +goto block_34; +} +block_34: +{ +if (x_26 == 0) +{ +lean_object* x_27; uint8_t x_28; +lean_dec_ref(x_10); +lean_inc_ref(x_6); +lean_inc(x_1); +x_27 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_D(x_1, x_21, x_24, x_6, x_25); +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_27, 0); +x_30 = lean_ctor_get(x_27, 1); +lean_ctor_set(x_27, 1, x_29); +lean_ctor_set(x_27, 0, x_21); +x_13 = x_27; +x_14 = x_30; +goto block_19; } else { lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_28, 0); -x_32 = lean_ctor_get(x_28, 1); +x_31 = lean_ctor_get(x_27, 0); +x_32 = lean_ctor_get(x_27, 1); lean_inc(x_32); lean_inc(x_31); -lean_dec(x_28); -lean_ctor_set(x_1, 3, x_31); +lean_dec(x_27); x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_1); -lean_ctor_set(x_33, 1, x_32); -return x_33; +lean_ctor_set(x_33, 0, x_21); +lean_ctor_set(x_33, 1, x_31); +x_13 = x_33; +x_14 = x_32; +goto block_19; } } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; uint8_t x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; size_t x_44; size_t x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_34 = lean_ctor_get(x_1, 1); -x_35 = lean_ctor_get(x_1, 3); -x_36 = lean_ctor_get(x_2, 0); -x_37 = lean_ctor_get(x_2, 1); -x_38 = lean_ctor_get(x_2, 2); -x_39 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); -lean_inc(x_38); -lean_inc(x_37); -lean_inc(x_36); -lean_dec(x_2); -lean_inc_ref(x_38); -x_40 = l_Lean_PersistentHashMap_contains___at___00Lean_IR_ResetReuse_R_spec__0___redArg(x_38, x_34); -x_41 = lean_box(0); -lean_inc(x_34); -x_42 = l_Lean_PersistentHashMap_insert___at___00Lean_IR_ResetReuse_R_spec__1___redArg(x_38, x_34, x_41); -x_43 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_43, 0, x_36); -lean_ctor_set(x_43, 1, x_37); -lean_ctor_set(x_43, 2, x_42); -lean_ctor_set_uint8(x_43, sizeof(void*)*3, x_39); -x_44 = lean_array_size(x_35); -x_45 = 0; -lean_inc(x_34); -x_46 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ResetReuse_R_spec__2(x_34, x_40, x_44, x_45, x_35, x_43, x_3); -x_47 = lean_ctor_get(x_46, 0); -lean_inc(x_47); -x_48 = lean_ctor_get(x_46, 1); -lean_inc(x_48); -if (lean_is_exclusive(x_46)) { - lean_ctor_release(x_46, 0); - lean_ctor_release(x_46, 1); - x_49 = x_46; -} else { - lean_dec_ref(x_46); - x_49 = lean_box(0); +lean_dec(x_24); +lean_dec_ref(x_21); +x_13 = x_10; +x_14 = x_25; +goto block_19; } -lean_ctor_set(x_1, 3, x_47); -if (lean_is_scalar(x_49)) { - x_50 = lean_alloc_ctor(0, 2, 0); -} else { - x_50 = x_49; } -lean_ctor_set(x_50, 0, x_1); -lean_ctor_set(x_50, 1, x_48); -return x_50; } +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; uint8_t x_49; +x_36 = lean_ctor_get(x_10, 0); +x_37 = lean_ctor_get(x_10, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_10); +lean_inc_ref(x_6); +x_38 = l_Lean_IR_ResetReuse_R(x_37, x_6, x_7); +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +lean_dec_ref(x_38); +lean_inc(x_39); +lean_inc_ref(x_36); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_36); +lean_ctor_set(x_41, 1, x_39); +x_49 = l_Lean_IR_CtorInfo_isScalar(x_36); +if (x_49 == 0) +{ +x_42 = x_2; +goto block_48; } else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; uint8_t x_58; lean_object* x_59; uint8_t x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; size_t x_64; size_t x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_51 = lean_ctor_get(x_1, 0); -x_52 = lean_ctor_get(x_1, 1); -x_53 = lean_ctor_get(x_1, 2); -x_54 = lean_ctor_get(x_1, 3); -lean_inc(x_54); -lean_inc(x_53); -lean_inc(x_52); -lean_inc(x_51); -lean_dec(x_1); -x_55 = lean_ctor_get(x_2, 0); -lean_inc_ref(x_55); -x_56 = lean_ctor_get(x_2, 1); -lean_inc(x_56); -x_57 = lean_ctor_get(x_2, 2); -lean_inc_ref(x_57); -x_58 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); -if (lean_is_exclusive(x_2)) { - lean_ctor_release(x_2, 0); - lean_ctor_release(x_2, 1); - lean_ctor_release(x_2, 2); - x_59 = x_2; -} else { - lean_dec_ref(x_2); - x_59 = lean_box(0); +x_42 = x_49; +goto block_48; } -lean_inc_ref(x_57); -x_60 = l_Lean_PersistentHashMap_contains___at___00Lean_IR_ResetReuse_R_spec__0___redArg(x_57, x_52); -x_61 = lean_box(0); -lean_inc(x_52); -x_62 = l_Lean_PersistentHashMap_insert___at___00Lean_IR_ResetReuse_R_spec__1___redArg(x_57, x_52, x_61); -if (lean_is_scalar(x_59)) { - x_63 = lean_alloc_ctor(0, 3, 1); -} else { - x_63 = x_59; -} -lean_ctor_set(x_63, 0, x_55); -lean_ctor_set(x_63, 1, x_56); -lean_ctor_set(x_63, 2, x_62); -lean_ctor_set_uint8(x_63, sizeof(void*)*3, x_58); -x_64 = lean_array_size(x_54); -x_65 = 0; -lean_inc(x_52); -x_66 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ResetReuse_R_spec__2(x_52, x_60, x_64, x_65, x_54, x_63, x_3); -x_67 = lean_ctor_get(x_66, 0); -lean_inc(x_67); -x_68 = lean_ctor_get(x_66, 1); -lean_inc(x_68); -if (lean_is_exclusive(x_66)) { - lean_ctor_release(x_66, 0); - lean_ctor_release(x_66, 1); - x_69 = x_66; -} else { - lean_dec_ref(x_66); - x_69 = lean_box(0); -} -x_70 = lean_alloc_ctor(9, 4, 0); -lean_ctor_set(x_70, 0, x_51); -lean_ctor_set(x_70, 1, x_52); -lean_ctor_set(x_70, 2, x_53); -lean_ctor_set(x_70, 3, x_67); -if (lean_is_scalar(x_69)) { - x_71 = lean_alloc_ctor(0, 2, 0); +block_48: +{ +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +lean_dec_ref(x_41); +lean_inc_ref(x_6); +lean_inc(x_1); +x_43 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_D(x_1, x_36, x_39, x_6, x_40); +x_44 = lean_ctor_get(x_43, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_43, 1); +lean_inc(x_45); +if (lean_is_exclusive(x_43)) { + lean_ctor_release(x_43, 0); + lean_ctor_release(x_43, 1); + x_46 = x_43; +} else { + lean_dec_ref(x_43); + x_46 = lean_box(0); +} +if (lean_is_scalar(x_46)) { + x_47 = lean_alloc_ctor(0, 2, 0); +} else { + x_47 = x_46; +} +lean_ctor_set(x_47, 0, x_36); +lean_ctor_set(x_47, 1, x_44); +x_13 = x_47; +x_14 = x_45; +goto block_19; +} +else +{ +lean_dec(x_39); +lean_dec_ref(x_36); +x_13 = x_41; +x_14 = x_40; +goto block_19; +} +} +} +} +else +{ +uint8_t x_50; +x_50 = !lean_is_exclusive(x_10); +if (x_50 == 0) +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_51 = lean_ctor_get(x_10, 0); +lean_inc_ref(x_6); +x_52 = l_Lean_IR_ResetReuse_R(x_51, x_6, x_7); +x_53 = lean_ctor_get(x_52, 0); +lean_inc(x_53); +x_54 = lean_ctor_get(x_52, 1); +lean_inc(x_54); +lean_dec_ref(x_52); +lean_ctor_set(x_10, 0, x_53); +x_13 = x_10; +x_14 = x_54; +goto block_19; +} +else +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_55 = lean_ctor_get(x_10, 0); +lean_inc(x_55); +lean_dec(x_10); +lean_inc_ref(x_6); +x_56 = l_Lean_IR_ResetReuse_R(x_55, x_6, x_7); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_56, 1); +lean_inc(x_58); +lean_dec_ref(x_56); +x_59 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_59, 0, x_57); +x_13 = x_59; +x_14 = x_58; +goto block_19; +} +} +block_19: +{ +size_t x_15; size_t x_16; lean_object* x_17; +x_15 = 1; +x_16 = lean_usize_add(x_4, x_15); +x_17 = lean_array_uset(x_12, x_4, x_13); +x_4 = x_16; +x_5 = x_17; +x_7 = x_14; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_ResetReuse_R(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +switch (lean_obj_tag(x_1)) { +case 9: +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_1); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_19 = lean_ctor_get(x_1, 1); +x_20 = lean_ctor_get(x_1, 2); +x_21 = lean_ctor_get(x_1, 3); +x_22 = l_Lean_IR_IRType_isStruct(x_20); +if (x_22 == 0) +{ +uint8_t x_23; +x_23 = !lean_is_exclusive(x_2); +if (x_23 == 0) +{ +lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; size_t x_28; size_t x_29; lean_object* x_30; uint8_t x_31; +x_24 = lean_ctor_get(x_2, 2); +lean_inc_ref(x_24); +x_25 = l_Lean_PersistentHashMap_contains___at___00Lean_IR_ResetReuse_R_spec__0___redArg(x_24, x_19); +x_26 = lean_box(0); +lean_inc(x_19); +x_27 = l_Lean_PersistentHashMap_insert___at___00Lean_IR_ResetReuse_R_spec__1___redArg(x_24, x_19, x_26); +lean_ctor_set(x_2, 2, x_27); +x_28 = lean_array_size(x_21); +x_29 = 0; +lean_inc(x_19); +x_30 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ResetReuse_R_spec__2(x_19, x_25, x_28, x_29, x_21, x_2, x_3); +x_31 = !lean_is_exclusive(x_30); +if (x_31 == 0) +{ +lean_object* x_32; +x_32 = lean_ctor_get(x_30, 0); +lean_ctor_set(x_1, 3, x_32); +lean_ctor_set(x_30, 0, x_1); +return x_30; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_30, 0); +x_34 = lean_ctor_get(x_30, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_30); +lean_ctor_set(x_1, 3, x_33); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_1); +lean_ctor_set(x_35, 1, x_34); +return x_35; +} +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; uint8_t x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; size_t x_44; size_t x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_36 = lean_ctor_get(x_2, 0); +x_37 = lean_ctor_get(x_2, 1); +x_38 = lean_ctor_get(x_2, 2); +x_39 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); +lean_inc(x_38); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_2); +lean_inc_ref(x_38); +x_40 = l_Lean_PersistentHashMap_contains___at___00Lean_IR_ResetReuse_R_spec__0___redArg(x_38, x_19); +x_41 = lean_box(0); +lean_inc(x_19); +x_42 = l_Lean_PersistentHashMap_insert___at___00Lean_IR_ResetReuse_R_spec__1___redArg(x_38, x_19, x_41); +x_43 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_43, 0, x_36); +lean_ctor_set(x_43, 1, x_37); +lean_ctor_set(x_43, 2, x_42); +lean_ctor_set_uint8(x_43, sizeof(void*)*3, x_39); +x_44 = lean_array_size(x_21); +x_45 = 0; +lean_inc(x_19); +x_46 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ResetReuse_R_spec__2(x_19, x_40, x_44, x_45, x_21, x_43, x_3); +x_47 = lean_ctor_get(x_46, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_46, 1); +lean_inc(x_48); +if (lean_is_exclusive(x_46)) { + lean_ctor_release(x_46, 0); + lean_ctor_release(x_46, 1); + x_49 = x_46; +} else { + lean_dec_ref(x_46); + x_49 = lean_box(0); +} +lean_ctor_set(x_1, 3, x_47); +if (lean_is_scalar(x_49)) { + x_50 = lean_alloc_ctor(0, 2, 0); +} else { + x_50 = x_49; +} +lean_ctor_set(x_50, 0, x_1); +lean_ctor_set(x_50, 1, x_48); +return x_50; +} +} +else +{ +size_t x_51; size_t x_52; lean_object* x_53; uint8_t x_54; +x_51 = lean_array_size(x_21); +x_52 = 0; +x_53 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ResetReuse_R_spec__3(x_51, x_52, x_21, x_2, x_3); +x_54 = !lean_is_exclusive(x_53); +if (x_54 == 0) +{ +lean_object* x_55; +x_55 = lean_ctor_get(x_53, 0); +lean_ctor_set(x_1, 3, x_55); +lean_ctor_set(x_53, 0, x_1); +return x_53; +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_53, 0); +x_57 = lean_ctor_get(x_53, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_53); +lean_ctor_set(x_1, 3, x_56); +x_58 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_58, 0, x_1); +lean_ctor_set(x_58, 1, x_57); +return x_58; +} +} +} +else +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; +x_59 = lean_ctor_get(x_1, 0); +x_60 = lean_ctor_get(x_1, 1); +x_61 = lean_ctor_get(x_1, 2); +x_62 = lean_ctor_get(x_1, 3); +lean_inc(x_62); +lean_inc(x_61); +lean_inc(x_60); +lean_inc(x_59); +lean_dec(x_1); +x_63 = l_Lean_IR_IRType_isStruct(x_61); +if (x_63 == 0) +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; lean_object* x_68; uint8_t x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; size_t x_73; size_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_64 = lean_ctor_get(x_2, 0); +lean_inc_ref(x_64); +x_65 = lean_ctor_get(x_2, 1); +lean_inc(x_65); +x_66 = lean_ctor_get(x_2, 2); +lean_inc_ref(x_66); +x_67 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); +if (lean_is_exclusive(x_2)) { + lean_ctor_release(x_2, 0); + lean_ctor_release(x_2, 1); + lean_ctor_release(x_2, 2); + x_68 = x_2; +} else { + lean_dec_ref(x_2); + x_68 = lean_box(0); +} +lean_inc_ref(x_66); +x_69 = l_Lean_PersistentHashMap_contains___at___00Lean_IR_ResetReuse_R_spec__0___redArg(x_66, x_60); +x_70 = lean_box(0); +lean_inc(x_60); +x_71 = l_Lean_PersistentHashMap_insert___at___00Lean_IR_ResetReuse_R_spec__1___redArg(x_66, x_60, x_70); +if (lean_is_scalar(x_68)) { + x_72 = lean_alloc_ctor(0, 3, 1); +} else { + x_72 = x_68; +} +lean_ctor_set(x_72, 0, x_64); +lean_ctor_set(x_72, 1, x_65); +lean_ctor_set(x_72, 2, x_71); +lean_ctor_set_uint8(x_72, sizeof(void*)*3, x_67); +x_73 = lean_array_size(x_62); +x_74 = 0; +lean_inc(x_60); +x_75 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ResetReuse_R_spec__2(x_60, x_69, x_73, x_74, x_62, x_72, x_3); +x_76 = lean_ctor_get(x_75, 0); +lean_inc(x_76); +x_77 = lean_ctor_get(x_75, 1); +lean_inc(x_77); +if (lean_is_exclusive(x_75)) { + lean_ctor_release(x_75, 0); + lean_ctor_release(x_75, 1); + x_78 = x_75; +} else { + lean_dec_ref(x_75); + x_78 = lean_box(0); +} +x_79 = lean_alloc_ctor(9, 4, 0); +lean_ctor_set(x_79, 0, x_59); +lean_ctor_set(x_79, 1, x_60); +lean_ctor_set(x_79, 2, x_61); +lean_ctor_set(x_79, 3, x_76); +if (lean_is_scalar(x_78)) { + x_80 = lean_alloc_ctor(0, 2, 0); } else { - x_71 = x_69; + x_80 = x_78; +} +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_77); +return x_80; +} +else +{ +size_t x_81; size_t x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_81 = lean_array_size(x_62); +x_82 = 0; +x_83 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ResetReuse_R_spec__3(x_81, x_82, x_62, x_2, x_3); +x_84 = lean_ctor_get(x_83, 0); +lean_inc(x_84); +x_85 = lean_ctor_get(x_83, 1); +lean_inc(x_85); +if (lean_is_exclusive(x_83)) { + lean_ctor_release(x_83, 0); + lean_ctor_release(x_83, 1); + x_86 = x_83; +} else { + lean_dec_ref(x_83); + x_86 = lean_box(0); +} +x_87 = lean_alloc_ctor(9, 4, 0); +lean_ctor_set(x_87, 0, x_59); +lean_ctor_set(x_87, 1, x_60); +lean_ctor_set(x_87, 2, x_61); +lean_ctor_set(x_87, 3, x_84); +if (lean_is_scalar(x_86)) { + x_88 = lean_alloc_ctor(0, 2, 0); +} else { + x_88 = x_86; +} +lean_ctor_set(x_88, 0, x_87); +lean_ctor_set(x_88, 1, x_85); +return x_88; } -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_68); -return x_71; } } case 1: { -uint8_t x_72; -x_72 = !lean_is_exclusive(x_1); -if (x_72 == 0) +uint8_t x_89; +x_89 = !lean_is_exclusive(x_1); +if (x_89 == 0) { -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; uint8_t x_80; -x_73 = lean_ctor_get(x_1, 0); -x_74 = lean_ctor_get(x_1, 1); -x_75 = lean_ctor_get(x_1, 2); -x_76 = lean_ctor_get(x_1, 3); +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; uint8_t x_97; +x_90 = lean_ctor_get(x_1, 0); +x_91 = lean_ctor_get(x_1, 1); +x_92 = lean_ctor_get(x_1, 2); +x_93 = lean_ctor_get(x_1, 3); lean_inc_ref(x_2); -x_77 = l_Lean_IR_ResetReuse_R(x_75, x_2, x_3); -x_78 = lean_ctor_get(x_77, 0); -lean_inc(x_78); -x_79 = lean_ctor_get(x_77, 1); -lean_inc(x_79); -lean_dec_ref(x_77); -x_80 = !lean_is_exclusive(x_2); -if (x_80 == 0) -{ -lean_object* x_81; lean_object* x_82; lean_object* x_83; uint8_t x_84; -x_81 = lean_ctor_get(x_2, 1); -lean_inc(x_78); -lean_inc_ref(x_74); -lean_inc(x_73); -x_82 = l_Lean_IR_LocalContext_addJP(x_81, x_73, x_74, x_78); -lean_ctor_set(x_2, 1, x_82); -x_83 = l_Lean_IR_ResetReuse_R(x_76, x_2, x_79); -x_84 = !lean_is_exclusive(x_83); -if (x_84 == 0) +x_94 = l_Lean_IR_ResetReuse_R(x_92, x_2, x_3); +x_95 = lean_ctor_get(x_94, 0); +lean_inc(x_95); +x_96 = lean_ctor_get(x_94, 1); +lean_inc(x_96); +lean_dec_ref(x_94); +x_97 = !lean_is_exclusive(x_2); +if (x_97 == 0) { -lean_object* x_85; -x_85 = lean_ctor_get(x_83, 0); -lean_ctor_set(x_1, 3, x_85); -lean_ctor_set(x_1, 2, x_78); -lean_ctor_set(x_83, 0, x_1); -return x_83; +lean_object* x_98; lean_object* x_99; lean_object* x_100; uint8_t x_101; +x_98 = lean_ctor_get(x_2, 1); +lean_inc(x_95); +lean_inc_ref(x_91); +lean_inc(x_90); +x_99 = l_Lean_IR_LocalContext_addJP(x_98, x_90, x_91, x_95); +lean_ctor_set(x_2, 1, x_99); +x_100 = l_Lean_IR_ResetReuse_R(x_93, x_2, x_96); +x_101 = !lean_is_exclusive(x_100); +if (x_101 == 0) +{ +lean_object* x_102; +x_102 = lean_ctor_get(x_100, 0); +lean_ctor_set(x_1, 3, x_102); +lean_ctor_set(x_1, 2, x_95); +lean_ctor_set(x_100, 0, x_1); +return x_100; } else { -lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_86 = lean_ctor_get(x_83, 0); -x_87 = lean_ctor_get(x_83, 1); -lean_inc(x_87); -lean_inc(x_86); -lean_dec(x_83); -lean_ctor_set(x_1, 3, x_86); -lean_ctor_set(x_1, 2, x_78); -x_88 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_88, 0, x_1); -lean_ctor_set(x_88, 1, x_87); -return x_88; +lean_object* x_103; lean_object* x_104; lean_object* x_105; +x_103 = lean_ctor_get(x_100, 0); +x_104 = lean_ctor_get(x_100, 1); +lean_inc(x_104); +lean_inc(x_103); +lean_dec(x_100); +lean_ctor_set(x_1, 3, x_103); +lean_ctor_set(x_1, 2, x_95); +x_105 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_105, 0, x_1); +lean_ctor_set(x_105, 1, x_104); +return x_105; } } else { -lean_object* x_89; lean_object* x_90; lean_object* x_91; uint8_t x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_89 = lean_ctor_get(x_2, 0); -x_90 = lean_ctor_get(x_2, 1); -x_91 = lean_ctor_get(x_2, 2); -x_92 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); -lean_inc(x_91); -lean_inc(x_90); -lean_inc(x_89); +lean_object* x_106; lean_object* x_107; lean_object* x_108; uint8_t x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_106 = lean_ctor_get(x_2, 0); +x_107 = lean_ctor_get(x_2, 1); +x_108 = lean_ctor_get(x_2, 2); +x_109 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); +lean_inc(x_108); +lean_inc(x_107); +lean_inc(x_106); lean_dec(x_2); -lean_inc(x_78); -lean_inc_ref(x_74); -lean_inc(x_73); -x_93 = l_Lean_IR_LocalContext_addJP(x_90, x_73, x_74, x_78); -x_94 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_94, 0, x_89); -lean_ctor_set(x_94, 1, x_93); -lean_ctor_set(x_94, 2, x_91); -lean_ctor_set_uint8(x_94, sizeof(void*)*3, x_92); -x_95 = l_Lean_IR_ResetReuse_R(x_76, x_94, x_79); -x_96 = lean_ctor_get(x_95, 0); -lean_inc(x_96); -x_97 = lean_ctor_get(x_95, 1); -lean_inc(x_97); -if (lean_is_exclusive(x_95)) { - lean_ctor_release(x_95, 0); - lean_ctor_release(x_95, 1); - x_98 = x_95; +lean_inc(x_95); +lean_inc_ref(x_91); +lean_inc(x_90); +x_110 = l_Lean_IR_LocalContext_addJP(x_107, x_90, x_91, x_95); +x_111 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_111, 0, x_106); +lean_ctor_set(x_111, 1, x_110); +lean_ctor_set(x_111, 2, x_108); +lean_ctor_set_uint8(x_111, sizeof(void*)*3, x_109); +x_112 = l_Lean_IR_ResetReuse_R(x_93, x_111, x_96); +x_113 = lean_ctor_get(x_112, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_112, 1); +lean_inc(x_114); +if (lean_is_exclusive(x_112)) { + lean_ctor_release(x_112, 0); + lean_ctor_release(x_112, 1); + x_115 = x_112; } else { - lean_dec_ref(x_95); - x_98 = lean_box(0); + lean_dec_ref(x_112); + x_115 = lean_box(0); } -lean_ctor_set(x_1, 3, x_96); -lean_ctor_set(x_1, 2, x_78); -if (lean_is_scalar(x_98)) { - x_99 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1, 3, x_113); +lean_ctor_set(x_1, 2, x_95); +if (lean_is_scalar(x_115)) { + x_116 = lean_alloc_ctor(0, 2, 0); } else { - x_99 = x_98; + x_116 = x_115; } -lean_ctor_set(x_99, 0, x_1); -lean_ctor_set(x_99, 1, x_97); -return x_99; +lean_ctor_set(x_116, 0, x_1); +lean_ctor_set(x_116, 1, x_114); +return x_116; } } else { -lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; uint8_t x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; -x_100 = lean_ctor_get(x_1, 0); -x_101 = lean_ctor_get(x_1, 1); -x_102 = lean_ctor_get(x_1, 2); -x_103 = lean_ctor_get(x_1, 3); -lean_inc(x_103); -lean_inc(x_102); -lean_inc(x_101); -lean_inc(x_100); +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; uint8_t x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; +x_117 = lean_ctor_get(x_1, 0); +x_118 = lean_ctor_get(x_1, 1); +x_119 = lean_ctor_get(x_1, 2); +x_120 = lean_ctor_get(x_1, 3); +lean_inc(x_120); +lean_inc(x_119); +lean_inc(x_118); +lean_inc(x_117); lean_dec(x_1); lean_inc_ref(x_2); -x_104 = l_Lean_IR_ResetReuse_R(x_102, x_2, x_3); -x_105 = lean_ctor_get(x_104, 0); -lean_inc(x_105); -x_106 = lean_ctor_get(x_104, 1); -lean_inc(x_106); -lean_dec_ref(x_104); -x_107 = lean_ctor_get(x_2, 0); -lean_inc_ref(x_107); -x_108 = lean_ctor_get(x_2, 1); -lean_inc(x_108); -x_109 = lean_ctor_get(x_2, 2); -lean_inc_ref(x_109); -x_110 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); +x_121 = l_Lean_IR_ResetReuse_R(x_119, x_2, x_3); +x_122 = lean_ctor_get(x_121, 0); +lean_inc(x_122); +x_123 = lean_ctor_get(x_121, 1); +lean_inc(x_123); +lean_dec_ref(x_121); +x_124 = lean_ctor_get(x_2, 0); +lean_inc_ref(x_124); +x_125 = lean_ctor_get(x_2, 1); +lean_inc(x_125); +x_126 = lean_ctor_get(x_2, 2); +lean_inc_ref(x_126); +x_127 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); if (lean_is_exclusive(x_2)) { lean_ctor_release(x_2, 0); lean_ctor_release(x_2, 1); lean_ctor_release(x_2, 2); - x_111 = x_2; + x_128 = x_2; } else { lean_dec_ref(x_2); - x_111 = lean_box(0); + x_128 = lean_box(0); } -lean_inc(x_105); -lean_inc_ref(x_101); -lean_inc(x_100); -x_112 = l_Lean_IR_LocalContext_addJP(x_108, x_100, x_101, x_105); -if (lean_is_scalar(x_111)) { - x_113 = lean_alloc_ctor(0, 3, 1); +lean_inc(x_122); +lean_inc_ref(x_118); +lean_inc(x_117); +x_129 = l_Lean_IR_LocalContext_addJP(x_125, x_117, x_118, x_122); +if (lean_is_scalar(x_128)) { + x_130 = lean_alloc_ctor(0, 3, 1); } else { - x_113 = x_111; -} -lean_ctor_set(x_113, 0, x_107); -lean_ctor_set(x_113, 1, x_112); -lean_ctor_set(x_113, 2, x_109); -lean_ctor_set_uint8(x_113, sizeof(void*)*3, x_110); -x_114 = l_Lean_IR_ResetReuse_R(x_103, x_113, x_106); -x_115 = lean_ctor_get(x_114, 0); -lean_inc(x_115); -x_116 = lean_ctor_get(x_114, 1); -lean_inc(x_116); -if (lean_is_exclusive(x_114)) { - lean_ctor_release(x_114, 0); - lean_ctor_release(x_114, 1); - x_117 = x_114; + x_130 = x_128; +} +lean_ctor_set(x_130, 0, x_124); +lean_ctor_set(x_130, 1, x_129); +lean_ctor_set(x_130, 2, x_126); +lean_ctor_set_uint8(x_130, sizeof(void*)*3, x_127); +x_131 = l_Lean_IR_ResetReuse_R(x_120, x_130, x_123); +x_132 = lean_ctor_get(x_131, 0); +lean_inc(x_132); +x_133 = lean_ctor_get(x_131, 1); +lean_inc(x_133); +if (lean_is_exclusive(x_131)) { + lean_ctor_release(x_131, 0); + lean_ctor_release(x_131, 1); + x_134 = x_131; } else { - lean_dec_ref(x_114); - x_117 = lean_box(0); -} -x_118 = lean_alloc_ctor(1, 4, 0); -lean_ctor_set(x_118, 0, x_100); -lean_ctor_set(x_118, 1, x_101); -lean_ctor_set(x_118, 2, x_105); -lean_ctor_set(x_118, 3, x_115); -if (lean_is_scalar(x_117)) { - x_119 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_131); + x_134 = lean_box(0); +} +x_135 = lean_alloc_ctor(1, 4, 0); +lean_ctor_set(x_135, 0, x_117); +lean_ctor_set(x_135, 1, x_118); +lean_ctor_set(x_135, 2, x_122); +lean_ctor_set(x_135, 3, x_132); +if (lean_is_scalar(x_134)) { + x_136 = lean_alloc_ctor(0, 2, 0); } else { - x_119 = x_117; + x_136 = x_134; } -lean_ctor_set(x_119, 0, x_118); -lean_ctor_set(x_119, 1, x_116); -return x_119; +lean_ctor_set(x_136, 0, x_135); +lean_ctor_set(x_136, 1, x_133); +return x_136; } } default: { -uint8_t x_120; -x_120 = l_Lean_IR_FnBody_isTerminal(x_1); -if (x_120 == 0) +uint8_t x_137; +x_137 = l_Lean_IR_FnBody_isTerminal(x_1); +if (x_137 == 0) { switch (lean_obj_tag(x_1)) { case 0: { -lean_object* x_121; -x_121 = lean_ctor_get(x_1, 3); -lean_inc(x_121); -x_4 = x_121; +lean_object* x_138; +x_138 = lean_ctor_get(x_1, 3); +lean_inc(x_138); +x_4 = x_138; goto block_17; } case 1: { -lean_object* x_122; -x_122 = lean_ctor_get(x_1, 3); -lean_inc(x_122); -x_4 = x_122; +lean_object* x_139; +x_139 = lean_ctor_get(x_1, 3); +lean_inc(x_139); +x_4 = x_139; goto block_17; } case 2: { -lean_object* x_123; -x_123 = lean_ctor_get(x_1, 3); -lean_inc(x_123); -x_4 = x_123; +lean_object* x_140; +x_140 = lean_ctor_get(x_1, 3); +lean_inc(x_140); +x_4 = x_140; goto block_17; } case 4: { -lean_object* x_124; -x_124 = lean_ctor_get(x_1, 3); -lean_inc(x_124); -x_4 = x_124; +lean_object* x_141; +x_141 = lean_ctor_get(x_1, 4); +lean_inc(x_141); +x_4 = x_141; goto block_17; } case 5: { -lean_object* x_125; -x_125 = lean_ctor_get(x_1, 5); -lean_inc(x_125); -x_4 = x_125; +lean_object* x_142; +x_142 = lean_ctor_get(x_1, 6); +lean_inc(x_142); +x_4 = x_142; goto block_17; } case 3: { -lean_object* x_126; -x_126 = lean_ctor_get(x_1, 2); -lean_inc(x_126); -x_4 = x_126; +lean_object* x_143; +x_143 = lean_ctor_get(x_1, 2); +lean_inc(x_143); +x_4 = x_143; goto block_17; } case 6: { -lean_object* x_127; -x_127 = lean_ctor_get(x_1, 2); -lean_inc(x_127); -x_4 = x_127; +lean_object* x_144; +x_144 = lean_ctor_get(x_1, 2); +lean_inc(x_144); +x_4 = x_144; goto block_17; } case 7: { -lean_object* x_128; -x_128 = lean_ctor_get(x_1, 2); -lean_inc(x_128); -x_4 = x_128; +lean_object* x_145; +x_145 = lean_ctor_get(x_1, 2); +lean_inc(x_145); +x_4 = x_145; goto block_17; } case 8: { -lean_object* x_129; -x_129 = lean_ctor_get(x_1, 1); -lean_inc(x_129); -x_4 = x_129; +lean_object* x_146; +x_146 = lean_ctor_get(x_1, 1); +lean_inc(x_146); +x_4 = x_146; goto block_17; } default: @@ -4020,12 +4336,12 @@ goto block_17; } else { -lean_object* x_130; +lean_object* x_147; lean_dec_ref(x_2); -x_130 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_130, 0, x_1); -lean_ctor_set(x_130, 1, x_3); -return x_130; +x_147 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_147, 0, x_1); +lean_ctor_set(x_147, 1, x_3); +return x_147; } } } @@ -4063,227 +4379,118 @@ return x_16; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ResetReuse_R_spec__2(lean_object* x_1, uint8_t x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ResetReuse_R_spec__3(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -uint8_t x_8; -x_8 = lean_usize_dec_lt(x_4, x_3); -if (x_8 == 0) +uint8_t x_6; +x_6 = lean_usize_dec_lt(x_2, x_1); +if (x_6 == 0) { -lean_object* x_9; -lean_dec_ref(x_6); -lean_dec(x_1); -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_5); -lean_ctor_set(x_9, 1, x_7); -return x_9; +lean_object* x_7; +lean_dec_ref(x_4); +x_7 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_7, 0, x_3); +lean_ctor_set(x_7, 1, x_5); +return x_7; } else { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_10 = lean_array_uget(x_5, x_4); -x_11 = lean_unsigned_to_nat(0u); -x_12 = lean_array_uset(x_5, x_4, x_11); -if (lean_obj_tag(x_10) == 0) -{ -uint8_t x_20; -x_20 = !lean_is_exclusive(x_10); -if (x_20 == 0) +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_8 = lean_array_uget(x_3, x_2); +x_9 = lean_unsigned_to_nat(0u); +x_10 = lean_array_uset(x_3, x_2, x_9); +if (lean_obj_tag(x_8) == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; uint8_t x_35; -x_21 = lean_ctor_get(x_10, 0); -x_22 = lean_ctor_get(x_10, 1); -lean_inc_ref(x_6); -x_23 = l_Lean_IR_ResetReuse_R(x_22, 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_ref(x_23); -lean_inc(x_24); -lean_inc_ref(x_21); -lean_ctor_set(x_10, 1, x_24); -x_35 = l_Lean_IR_CtorInfo_isScalar(x_21); -if (x_35 == 0) +uint8_t x_18; +x_18 = !lean_is_exclusive(x_8); +if (x_18 == 0) { -x_26 = x_2; -goto block_34; +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_19 = lean_ctor_get(x_8, 1); +lean_inc_ref(x_4); +x_20 = l_Lean_IR_ResetReuse_R(x_19, x_4, x_5); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec_ref(x_20); +lean_ctor_set(x_8, 1, x_21); +x_11 = x_8; +x_12 = x_22; +goto block_17; } else { -x_26 = x_35; -goto block_34; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_23 = lean_ctor_get(x_8, 0); +x_24 = lean_ctor_get(x_8, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_8); +lean_inc_ref(x_4); +x_25 = l_Lean_IR_ResetReuse_R(x_24, x_4, x_5); +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +lean_dec_ref(x_25); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_23); +lean_ctor_set(x_28, 1, x_26); +x_11 = x_28; +x_12 = x_27; +goto block_17; } -block_34: -{ -if (x_26 == 0) -{ -lean_object* x_27; uint8_t x_28; -lean_dec_ref(x_10); -lean_inc_ref(x_6); -lean_inc(x_1); -x_27 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_D(x_1, x_21, x_24, x_6, x_25); -x_28 = !lean_is_exclusive(x_27); -if (x_28 == 0) -{ -lean_object* x_29; lean_object* x_30; -x_29 = lean_ctor_get(x_27, 0); -x_30 = lean_ctor_get(x_27, 1); -lean_ctor_set(x_27, 1, x_29); -lean_ctor_set(x_27, 0, x_21); -x_13 = x_27; -x_14 = x_30; -goto block_19; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_27, 0); -x_32 = lean_ctor_get(x_27, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_27); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_21); -lean_ctor_set(x_33, 1, x_31); -x_13 = x_33; -x_14 = x_32; -goto block_19; -} -} -else +uint8_t x_29; +x_29 = !lean_is_exclusive(x_8); +if (x_29 == 0) { -lean_dec(x_24); -lean_dec_ref(x_21); -x_13 = x_10; -x_14 = x_25; -goto block_19; -} -} +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_8, 0); +lean_inc_ref(x_4); +x_31 = l_Lean_IR_ResetReuse_R(x_30, x_4, x_5); +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec_ref(x_31); +lean_ctor_set(x_8, 0, x_32); +x_11 = x_8; +x_12 = x_33; +goto block_17; } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; uint8_t x_49; -x_36 = lean_ctor_get(x_10, 0); -x_37 = lean_ctor_get(x_10, 1); -lean_inc(x_37); +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_34 = lean_ctor_get(x_8, 0); +lean_inc(x_34); +lean_dec(x_8); +lean_inc_ref(x_4); +x_35 = l_Lean_IR_ResetReuse_R(x_34, x_4, x_5); +x_36 = lean_ctor_get(x_35, 0); lean_inc(x_36); -lean_dec(x_10); -lean_inc_ref(x_6); -x_38 = l_Lean_IR_ResetReuse_R(x_37, x_6, x_7); -x_39 = lean_ctor_get(x_38, 0); -lean_inc(x_39); -x_40 = lean_ctor_get(x_38, 1); -lean_inc(x_40); -lean_dec_ref(x_38); -lean_inc(x_39); -lean_inc_ref(x_36); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_36); -lean_ctor_set(x_41, 1, x_39); -x_49 = l_Lean_IR_CtorInfo_isScalar(x_36); -if (x_49 == 0) -{ -x_42 = x_2; -goto block_48; -} -else -{ -x_42 = x_49; -goto block_48; -} -block_48: -{ -if (x_42 == 0) -{ -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -lean_dec_ref(x_41); -lean_inc_ref(x_6); -lean_inc(x_1); -x_43 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_D(x_1, x_36, x_39, x_6, x_40); -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_43, 1); -lean_inc(x_45); -if (lean_is_exclusive(x_43)) { - lean_ctor_release(x_43, 0); - lean_ctor_release(x_43, 1); - x_46 = x_43; -} else { - lean_dec_ref(x_43); - x_46 = lean_box(0); -} -if (lean_is_scalar(x_46)) { - x_47 = lean_alloc_ctor(0, 2, 0); -} else { - x_47 = x_46; -} -lean_ctor_set(x_47, 0, x_36); -lean_ctor_set(x_47, 1, x_44); -x_13 = x_47; -x_14 = x_45; -goto block_19; -} -else -{ -lean_dec(x_39); -lean_dec_ref(x_36); -x_13 = x_41; -x_14 = x_40; -goto block_19; -} -} -} -} -else -{ -uint8_t x_50; -x_50 = !lean_is_exclusive(x_10); -if (x_50 == 0) -{ -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_51 = lean_ctor_get(x_10, 0); -lean_inc_ref(x_6); -x_52 = l_Lean_IR_ResetReuse_R(x_51, x_6, x_7); -x_53 = lean_ctor_get(x_52, 0); -lean_inc(x_53); -x_54 = lean_ctor_get(x_52, 1); -lean_inc(x_54); -lean_dec_ref(x_52); -lean_ctor_set(x_10, 0, x_53); -x_13 = x_10; -x_14 = x_54; -goto block_19; -} -else -{ -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_55 = lean_ctor_get(x_10, 0); -lean_inc(x_55); -lean_dec(x_10); -lean_inc_ref(x_6); -x_56 = l_Lean_IR_ResetReuse_R(x_55, x_6, x_7); -x_57 = lean_ctor_get(x_56, 0); -lean_inc(x_57); -x_58 = lean_ctor_get(x_56, 1); -lean_inc(x_58); -lean_dec_ref(x_56); -x_59 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_59, 0, x_57); -x_13 = x_59; -x_14 = x_58; -goto block_19; +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec_ref(x_35); +x_38 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_38, 0, x_36); +x_11 = x_38; +x_12 = x_37; +goto block_17; } } -block_19: +block_17: { -size_t x_15; size_t x_16; lean_object* x_17; -x_15 = 1; -x_16 = lean_usize_add(x_4, x_15); -x_17 = lean_array_uset(x_12, x_4, x_13); -x_4 = x_16; -x_5 = x_17; -x_7 = x_14; +size_t x_13; size_t x_14; lean_object* x_15; +x_13 = 1; +x_14 = lean_usize_add(x_2, x_13); +x_15 = lean_array_uset(x_10, x_2, x_11); +x_2 = x_14; +x_3 = x_15; +x_5 = x_12; goto _start; } } @@ -4392,6 +4599,18 @@ x_6 = lean_box(x_5); return x_6; } } +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ResetReuse_R_spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_7 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_8 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ResetReuse_R_spec__3(x_6, x_7, x_3, x_4, x_5); +return x_8; +} +} LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ResetReuse_R_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: { diff --git a/stage0/stdlib/Lean/Compiler/IR/StructRC.c b/stage0/stdlib/Lean/Compiler/IR/StructRC.c new file mode 100644 index 000000000000..2a264a327f2e --- /dev/null +++ b/stage0/stdlib/Lean/Compiler/IR/StructRC.c @@ -0,0 +1,12620 @@ +// Lean compiler output +// Module: Lean.Compiler.IR.StructRC +// Imports: public import Lean.Compiler.IR.FreeVars public import Lean.Compiler.IR.Format import Std.Data.TreeMap.AdditionalOperations +#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 +extern lean_object* l_Lean_IR_instInhabitedArg_default; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_StructRC_Context_useVar_spec__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___00Array_Array_repr___at___00instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1_spec__3_spec__5___lam__0(lean_object*); +static lean_object* l_Lean_IR_StructRC_Context_renameVar_x21___closed__0; +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_ProjEntry_var_elim(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_StructRC_instReprProjEntry_repr___closed__4; +lean_object* lean_mk_empty_array_with_capacity(lean_object*); +static lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__24_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_remove___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_IR_StructRC_instReprEntry_repr___closed__5; +uint8_t l_Lean_IR_IRType_isDefiniteRef(lean_object*); +static lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__8_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +static lean_object* l_Lean_IR_StructRC_instReprEntry_repr___closed__3; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_StructRC_Context_useVar_spec__1___redArg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_StructRC_Context_addCtorKnowledge_spec__0(lean_object*); +static lean_object* l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__9; +static lean_object* l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__4; +lean_object* l_Lean_IR_Decl_maxIndex(lean_object*); +static lean_object* l_Lean_IR_StructRC_visitExpr___closed__0; +static lean_object* l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__9; +LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_StructRC_atConstructorIndex_spec__0(lean_object*); +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_Const_get_x3f___at___00Lean_IR_StructRC_Context_renameVar_x21_spec__0___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_StructRC_Context_resetRC___closed__0; +lean_object* l_Std_Format_fill(lean_object*); +static lean_object* l_Lean_IR_StructRC_instReprEntry_repr___closed__0; +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_ProjEntry_ctorElim(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__2; +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_atConstructorIndex___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_StructRC_visitFnBody_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_Const_getD___at___00Lean_IR_StructRC_Context_renameVar_spec__0___redArg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_StructRC_Context_useVar_spec__0___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_StructRC_instReprProjEntry_repr___closed__7; +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_instReprProjEntry_repr___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_accumulateRCDiff___lam__0(lean_object*, lean_object*, lean_object*); +lean_object* lean_array_push(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00Lean_IR_StructRC_visitExpr_spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_StructRC_instReprProjEntry_repr___closed__1; +uint8_t lean_usize_dec_eq(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_atConstructorIndex(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_instReprEntry_repr(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Vector_Basic_0__Vector_mapM_go___at___00Lean_IR_StructRC_Context_accumulateRCDiff_spec__0(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_Vector_Basic_0__Vector_mapM_go___at___00Lean_IR_StructRC_Context_accumulateRCDiff_spec__0___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__8; +lean_object* lean_mk_array(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00Lean_IR_StructRC_visitExpr_spec__1___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_StructRC_instReprEntry_repr___closed__6; +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_instInhabitedEntry; +lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__3; +static lean_object* l_Lean_IR_StructRC_Context_addVarBasic___closed__0; +static lean_object* l_Lean_IR_StructRC_Entry_ofStruct___closed__2; +static lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__30_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +static lean_object* l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__7; +static lean_object* l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__4; +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_ctorIdx___boxed(lean_object*); +static lean_object* l_Lean_IR_StructRC_Context_addProjInfo___closed__1; +lean_object* lean_array_fget_borrowed(lean_object*, lean_object*); +static lean_object* l_Lean_IR_StructRC_Context_addProjInfo___closed__0; +static lean_object* l_Lean_IR_StructRC_instReprEntry_repr___closed__8; +static lean_object* l_Lean_IR_StructRC_instReprProjEntry_repr___closed__0; +static lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__3_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +static lean_object* l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__3; +static lean_object* l_Lean_IR_StructRC_Context_accumulateRCDiff___closed__0; +static lean_object* l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__12; +static lean_object* l_Lean_IR_StructRC_instReprEntry_repr___closed__1; +lean_object* l_panic___redArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_addParams(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_struct_elim(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_ProjEntry_erased_elim(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_ProjEntry_ctorIdx(lean_object*); +static lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__5_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_StructRC_Context_useVar_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_StructRC_Context_renameArgs_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_StructRC_instReprEntry_repr___closed__2; +lean_object* lean_string_length(lean_object*); +static lean_object* l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__13; +uint8_t lean_int_dec_le(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_renameVar___boxed(lean_object*, lean_object*); +lean_object* l_Id_instMonad___lam__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_resetRC(lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_renameArg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_ProjEntry_ctorIdx___boxed(lean_object*); +lean_object* l_Array_ofFn___redArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_contains___at___00Lean_IR_StructRC_Context_insertIfNew_spec__0___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_insertRename(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_ofType(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___00Array_Array_repr___at___00instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1_spec__3_spec__5(lean_object*, lean_object*); +static lean_object* l_Lean_IR_StructRC_Context_renameVar_x21___closed__1; +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_accumulateRCDiff___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Nat_reprFast(lean_object*); +lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_addCtorKnowledge(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_Const_get_x3f___at___00Lean_IR_StructRC_Context_renameVar_x21_spec__0___redArg(lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__2_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +size_t lean_usize_of_nat(lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_finish(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_erase___at___00Lean_IR_StructRC_Context_remove_spec__0___redArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_StructRC_visitExpr_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_StructRC_Entry_ofStruct___closed__1; +static lean_object* l_Lean_IR_StructRC_instReprProjEntry_repr___closed__8; +static lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__15_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +static lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__13_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +static lean_object* l_Lean_IR_StructRC_Context_addProjInfo___closed__7; +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_ctorElim___redArg(lean_object*, lean_object*); +static lean_object* l_Lean_IR_StructRC_instReprProjEntry_repr___closed__9; +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_Const_get_x3f___at___00Lean_IR_StructRC_Context_renameVar_x21_spec__0(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_foldl___at___00Std_Format_joinSep___at___00Array_Array_repr___at___00instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1_spec__3_spec__5_spec__7(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_IR_StructRC_needsRC(lean_object*); +static lean_object* l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__1; +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_map___at___00Lean_IR_StructRC_Context_resetRC_spec__0(lean_object*); +static lean_object* l_Lean_IR_StructRC_Context_addProjInfo___closed__6; +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_renameVar_x21(lean_object*, lean_object*); +lean_object* l_Lean_IR_instReprVarId_repr___redArg(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00Lean_IR_StructRC_visitExpr_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_addInstr(lean_object*, lean_object*); +static lean_object* l_Lean_IR_StructRC_Context_addProjInfo___closed__4; +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_addVar(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_nat_to_int(lean_object*); +static lean_object* l_Lean_IR_StructRC_Context_addProjInfo___closed__3; +static lean_object* l_Lean_IR_StructRC_instReprEntry_repr___closed__7; +static lean_object* l_Lean_IR_StructRC_instReprEntry_repr___closed__4; +lean_object* l_Array_empty(lean_object*); +lean_object* l_Lean_IR_instReprIRType_repr(lean_object*, lean_object*); +static lean_object* l_Lean_IR_StructRC_atConstructorIndex___closed__1; +static lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__27_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_ProjEntry_ctorElim___redArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Nat_cast___at___00Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0_spec__1(lean_object*); +static lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__17_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_ProjEntry_unbound_elim___redArg(lean_object*, lean_object*); +lean_object* l_Id_instMonad___lam__2___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___00Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0_spec__0___lam__0(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_StructRC_visitFnBody_spec__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Lean_IR_StructRC_Context_renameVar_x21___closed__3; +static lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +lean_object* l_instInhabitedOfMonad___redArg(lean_object*, lean_object*); +static lean_object* l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__15; +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_ctorElim___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_useArg(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__7_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +lean_object* lean_array_to_list(lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_visitDecl(lean_object*); +LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_StructRC_Context_renameVar_x21_spec__1(lean_object*); +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_contains___at___00Lean_IR_StructRC_Context_insertIfNew_spec__0___redArg___boxed(lean_object*, lean_object*); +lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__14_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +static lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__16_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +static lean_object* l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__0; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_StructRC_visitExpr_spec__0(lean_object*, size_t, size_t, lean_object*, lean_object*); +lean_object* l_Int_repr(lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_instInhabitedEntry_default; +static lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +static lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__29_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_instInhabitedProjEntry_default; +static lean_object* l_Lean_IR_StructRC_instReprProjEntry_repr___closed__6; +lean_object* l_Std_DTreeMap_Internal_Impl_minView___redArg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_instReprEntry_repr___boxed(lean_object*, lean_object*); +lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +lean_object* l_Id_instMonad___lam__3(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_StructRC_instReprProjEntry_repr___closed__2; +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_insert(lean_object*, lean_object*, lean_object*); +static lean_object* l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__10; +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00Lean_IR_StructRC_visitExpr_spec__1___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_visitExpr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__5; +static lean_object* l_Lean_IR_StructRC_Entry_ofStruct___closed__0; +static lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__18_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_insert___at___00Lean_IR_StructRC_Context_insertRename_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_visitExpr___lam__0(lean_object*, lean_object*); +static lean_object* l_Lean_IR_StructRC_Context_addProjInfo___closed__2; +static lean_object* l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__7; +static lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__4_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +LEAN_EXPORT lean_object* l_List_foldl___at___00List_foldl___at___00Std_Format_joinSep___at___00Array_Array_repr___at___00instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1_spec__3_spec__5_spec__7_spec__8(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_visitExpr___lam__0___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_emitRCDiff(lean_object*, uint8_t, lean_object*, lean_object*); +static lean_object* l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__6; +static lean_object* l_Lean_IR_StructRC_visitExpr___lam__0___closed__0; +static lean_object* l_Lean_IR_StructRC_Context_addProjInfo___closed__8; +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_ref_elim(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_StructRC_instReprEntry___closed__0; +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_Const_getD___at___00Lean_IR_StructRC_Context_renameVar_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__16; +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_remove(lean_object*, lean_object*); +lean_object* lean_array_get_borrowed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_addProjInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_unknownUnion_elim(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Vector_Basic_0__Vector_mapM_go___at___00Lean_IR_StructRC_Context_accumulateRCDiff_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__20_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___00Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0_spec__0(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_ProjEntry_var_elim___redArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +static lean_object* l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__0; +LEAN_EXPORT lean_object* l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_StructRC_needsRC_spec__0___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__12_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_emitRCDiff___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__14; +lean_object* l_outOfBounds___redArg(lean_object*); +static lean_object* l_Lean_IR_StructRC_instReprProjEntry_repr___closed__5; +static lean_object* l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__5; +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_useVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0(lean_object*); +static lean_object* l_Lean_IR_StructRC_instReprProjEntry_repr___closed__10; +lean_object* lean_array_fget(lean_object*, lean_object*); +lean_object* lean_nat_abs(lean_object*); +LEAN_EXPORT lean_object* l_Array_Array_repr___at___00instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1_spec__3(lean_object*); +static lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__21_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_ctorIdx(lean_object*); +lean_object* l_Id_instMonad___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_StructRC_Context_addParams_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_mkFreshVar(lean_object*); +static lean_object* l_Lean_IR_StructRC_instReprEntry_repr___closed__9; +static lean_object* l_Lean_IR_StructRC_instReprProjEntry_repr___closed__3; +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__26_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +static lean_object* l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__2; +uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +static lean_object* l_Lean_IR_StructRC_Context_addProjInfo___closed__5; +static lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__10_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +static lean_object* l_panic___at___00Lean_IR_StructRC_atConstructorIndex_spec__0___closed__0; +static lean_object* l_Lean_IR_StructRC_Context_addProjInfo___closed__10; +static lean_object* l_Lean_IR_StructRC_instReprProjEntry___closed__0; +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_Const_get_x3f___at___00Lean_IR_StructRC_Context_renameVar_x21_spec__0___redArg___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_IR_StructRC_Context_addProjInfo___closed__11; +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_useVar(lean_object*, lean_object*, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_ofStruct___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_renameArgs___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_IR_StructRC_visitDecl___closed__0; +LEAN_EXPORT uint8_t l_Std_DTreeMap_Internal_Impl_contains___at___00Lean_IR_StructRC_Context_insertIfNew_spec__0___redArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_foldl___at___00List_foldl___at___00Std_Format_joinSep___at___00Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0_spec__0_spec__1_spec__4(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_renameVar(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_accumulateRCDiff(lean_object*, uint8_t, lean_object*, lean_object*); +lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_StructRC_Context_useVar_spec__1___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_addParams___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_ctorElim(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Repr_addAppParen(lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__19_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +lean_object* lean_panic_fn(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_persistent_elim___redArg(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_DTreeMap_Internal_Impl_contains___at___00Lean_IR_StructRC_Context_insertIfNew_spec__0(lean_object*, lean_object*, lean_object*); +uint8_t lean_int_dec_lt(lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__22_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +static lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__28_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +lean_object* lean_nat_sub(lean_object*, lean_object*); +lean_object* l_Bool_repr___redArg(uint8_t); +lean_object* lean_nat_mul(lean_object*, lean_object*); +lean_object* l_Std_DTreeMap_Internal_Impl_maxView___redArg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__11; +lean_object* l_Id_instMonad___lam__6(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__10; +static lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__9_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +static lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__23_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_needsRC___boxed(lean_object*); +static lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__25_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +uint8_t l_Lean_IR_IRType_isStruct(lean_object*); +static lean_object* l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__6; +LEAN_EXPORT uint8_t l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_StructRC_Context_useVar_spec__0(lean_object*, size_t, size_t); +static lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__6_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +static lean_object* l_Lean_IR_StructRC_visitExpr___closed__1; +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_erase___at___00Lean_IR_StructRC_Context_remove_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__1; +size_t lean_usize_add(size_t, size_t); +lean_object* l_mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_StructRC_needsRC_spec__0(lean_object*, size_t, size_t); +lean_object* lean_array_uget(lean_object*, size_t); +size_t lean_array_size(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_StructRC_Context_addParams_spec__0(lean_object*, size_t, size_t, lean_object*); +lean_object* lean_int_add(lean_object*, lean_object*); +static lean_object* l_Lean_IR_StructRC_Context_addProjInfo___closed__9; +static lean_object* l_Lean_IR_StructRC_atConstructorIndex___closed__2; +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_renameArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_Const_getD___at___00Lean_IR_StructRC_Context_renameVar_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_renameVar_x21___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__11_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_instInhabitedProjEntry; +uint8_t lean_int_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_StructRC_visitFnBody_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_ofStruct(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_instReprProjEntry_repr(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_foldl___at___00Std_Format_joinSep___at___00Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0_spec__0_spec__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_StructRC_visitFnBody___closed__0; +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_ProjEntry_unbound_elim(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_persistent_elim(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_renameArgs(lean_object*, lean_object*); +lean_object* lean_array_get_size(lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_ref_elim___redArg(lean_object*, lean_object*); +lean_object* l_Id_instMonad___lam__0(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_insertIfNew(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_ProjEntry_ctorElim___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_int_neg(lean_object*); +lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_IR_reshape(lean_object*, lean_object*); +static lean_object* l_Lean_IR_StructRC_Context_renameVar_x21___closed__2; +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___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_StructRC_Context_renameArgs_spec__0(lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_struct_elim___redArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_visitFnBody(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_ProjEntry_erased_elim___redArg(lean_object*, lean_object*); +lean_object* lean_nat_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_instReprEntry; +static lean_object* l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__8; +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_erase___at___00Lean_IR_StructRC_Context_remove_spec__0(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_Const_getD___at___00Lean_IR_StructRC_Context_renameVar_spec__0(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2____boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_StructRC_visitFnBody_spec__0(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); +static lean_object* l_Lean_IR_StructRC_instReprEntry_repr___closed__10; +LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_StructRC_visitExpr_spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_forInStep___at___00Lean_IR_StructRC_Context_finish_spec__0(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_insert___at___00Lean_IR_StructRC_Context_insertRename_spec__0___redArg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_StructRC_atConstructorIndex___closed__0; +LEAN_EXPORT lean_object* l___private_Init_Data_Vector_Basic_0__Vector_mapM_go___at___00Lean_IR_StructRC_Context_accumulateRCDiff_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +lean_object* l_Id_instMonad___lam__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_addVarBasic(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_unknownUnion_elim___redArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_erase___at___00Lean_IR_StructRC_Context_remove_spec__0___redArg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_instReprProjEntry; +LEAN_EXPORT lean_object* l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_ProjEntry_ctorIdx(lean_object* x_1) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 0: +{ +lean_object* x_2; +x_2 = lean_unsigned_to_nat(0u); +return x_2; +} +case 1: +{ +lean_object* x_3; +x_3 = lean_unsigned_to_nat(1u); +return x_3; +} +default: +{ +lean_object* x_4; +x_4 = lean_unsigned_to_nat(2u); +return x_4; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_ProjEntry_ctorIdx___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_IR_StructRC_ProjEntry_ctorIdx(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_ProjEntry_ctorElim___redArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +return x_2; +} +else +{ +lean_object* x_3; lean_object* x_4; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +lean_dec(x_1); +x_4 = lean_apply_1(x_2, x_3); +return x_4; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_ProjEntry_ctorElim(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_IR_StructRC_ProjEntry_ctorElim___redArg(x_3, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_ProjEntry_ctorElim___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_IR_StructRC_ProjEntry_ctorElim(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_2); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_ProjEntry_erased_elim(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_IR_StructRC_ProjEntry_ctorElim___redArg(x_2, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_ProjEntry_erased_elim___redArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_IR_StructRC_ProjEntry_ctorElim___redArg(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_ProjEntry_var_elim(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_IR_StructRC_ProjEntry_ctorElim___redArg(x_2, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_ProjEntry_var_elim___redArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_IR_StructRC_ProjEntry_ctorElim___redArg(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_ProjEntry_unbound_elim(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_IR_StructRC_ProjEntry_ctorElim___redArg(x_2, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_ProjEntry_unbound_elim___redArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_IR_StructRC_ProjEntry_ctorElim___redArg(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_IR_StructRC_instInhabitedProjEntry_default() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_StructRC_instInhabitedProjEntry() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_StructRC_instReprProjEntry_repr___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.IR.StructRC.ProjEntry.erased", 33, 33); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_StructRC_instReprProjEntry_repr___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__0; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_IR_StructRC_instReprProjEntry_repr___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(2u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_IR_StructRC_instReprProjEntry_repr___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_IR_StructRC_instReprProjEntry_repr___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.IR.StructRC.ProjEntry.var", 30, 30); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_StructRC_instReprProjEntry_repr___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__4; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_IR_StructRC_instReprProjEntry_repr___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(1); +x_2 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__5; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_IR_StructRC_instReprProjEntry_repr___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.IR.StructRC.ProjEntry.unbound", 34, 34); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_StructRC_instReprProjEntry_repr___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__7; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_IR_StructRC_instReprProjEntry_repr___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(1); +x_2 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__8; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_IR_StructRC_instReprProjEntry_repr___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_instReprProjEntry_repr(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_12; +switch (lean_obj_tag(x_1)) { +case 0: +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_unsigned_to_nat(1024u); +x_20 = lean_nat_dec_le(x_19, x_2); +if (x_20 == 0) +{ +lean_object* x_21; +x_21 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__2; +x_12 = x_21; +goto block_18; +} +else +{ +lean_object* x_22; +x_22 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__3; +x_12 = x_22; +goto block_18; +} +} +case 1: +{ +lean_object* x_23; lean_object* x_24; lean_object* x_33; uint8_t x_34; +x_23 = lean_ctor_get(x_1, 0); +lean_inc(x_23); +lean_dec_ref(x_1); +x_33 = lean_unsigned_to_nat(1024u); +x_34 = lean_nat_dec_le(x_33, x_2); +if (x_34 == 0) +{ +lean_object* x_35; +x_35 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__2; +x_24 = x_35; +goto block_32; +} +else +{ +lean_object* x_36; +x_36 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__3; +x_24 = x_36; +goto block_32; +} +block_32: +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; lean_object* x_31; +x_25 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__6; +x_26 = l_Lean_IR_instReprVarId_repr___redArg(x_23); +x_27 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_28, 0, x_24); +lean_ctor_set(x_28, 1, x_27); +x_29 = 0; +x_30 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set_uint8(x_30, sizeof(void*)*1, x_29); +x_31 = l_Repr_addAppParen(x_30, x_2); +return x_31; +} +} +default: +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_50; uint8_t x_51; +x_37 = lean_ctor_get(x_1, 0); +lean_inc(x_37); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + x_38 = x_1; +} else { + lean_dec_ref(x_1); + x_38 = lean_box(0); +} +x_50 = lean_unsigned_to_nat(1024u); +x_51 = lean_nat_dec_le(x_50, x_2); +if (x_51 == 0) +{ +lean_object* x_52; +x_52 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__2; +x_39 = x_52; +goto block_49; +} +else +{ +lean_object* x_53; +x_53 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__3; +x_39 = x_53; +goto block_49; +} +block_49: +{ +lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_40 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__9; +x_41 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__10; +x_42 = lean_int_dec_lt(x_37, x_41); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; +x_43 = l_Int_repr(x_37); +lean_dec(x_37); +if (lean_is_scalar(x_38)) { + x_44 = lean_alloc_ctor(3, 1, 0); +} else { + x_44 = x_38; + lean_ctor_set_tag(x_44, 3); +} +lean_ctor_set(x_44, 0, x_43); +x_3 = x_40; +x_4 = x_39; +x_5 = x_44; +goto block_11; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_45 = lean_unsigned_to_nat(1024u); +x_46 = l_Int_repr(x_37); +lean_dec(x_37); +if (lean_is_scalar(x_38)) { + x_47 = lean_alloc_ctor(3, 1, 0); +} else { + x_47 = x_38; + lean_ctor_set_tag(x_47, 3); +} +lean_ctor_set(x_47, 0, x_46); +x_48 = l_Repr_addAppParen(x_47, x_45); +x_3 = x_40; +x_4 = x_39; +x_5 = x_48; +goto block_11; +} +} +} +} +block_11: +{ +lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; +x_6 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_6, 0, x_3); +lean_ctor_set(x_6, 1, x_5); +x_7 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_7, 0, x_4); +lean_ctor_set(x_7, 1, x_6); +x_8 = 0; +x_9 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_9, 0, x_7); +lean_ctor_set_uint8(x_9, sizeof(void*)*1, x_8); +x_10 = l_Repr_addAppParen(x_9, x_2); +return x_10; +} +block_18: +{ +lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; +x_13 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__1; +x_14 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +x_15 = 0; +x_16 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set_uint8(x_16, sizeof(void*)*1, x_15); +x_17 = l_Repr_addAppParen(x_16, x_2); +return x_17; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_instReprProjEntry_repr___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_IR_StructRC_instReprProjEntry_repr(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_IR_StructRC_instReprProjEntry___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_IR_StructRC_instReprProjEntry_repr___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_StructRC_instReprProjEntry() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_IR_StructRC_instReprProjEntry___closed__0; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_ctorIdx(lean_object* x_1) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 0: +{ +lean_object* x_2; +x_2 = lean_unsigned_to_nat(0u); +return x_2; +} +case 1: +{ +lean_object* x_3; +x_3 = lean_unsigned_to_nat(1u); +return x_3; +} +case 2: +{ +lean_object* x_4; +x_4 = lean_unsigned_to_nat(2u); +return x_4; +} +default: +{ +lean_object* x_5; +x_5 = lean_unsigned_to_nat(3u); +return x_5; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_ctorIdx___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_IR_StructRC_Entry_ctorIdx(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_ctorElim___redArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 0: +{ +return x_2; +} +case 1: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_ctor_get(x_1, 0); +lean_inc_ref(x_3); +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +lean_dec_ref(x_1); +x_5 = lean_apply_2(x_2, x_3, x_4); +return x_5; +} +case 2: +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_6 = lean_ctor_get(x_1, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_7); +x_8 = lean_ctor_get(x_1, 2); +lean_inc_ref(x_8); +lean_dec_ref(x_1); +x_9 = lean_apply_3(x_2, x_6, x_7, x_8); +return x_9; +} +default: +{ +uint8_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_ctor_get_uint8(x_1, sizeof(void*)*1); +x_11 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +lean_dec_ref(x_1); +x_12 = lean_box(x_10); +x_13 = lean_apply_2(x_2, x_12, x_11); +return x_13; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_ctorElim(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_IR_StructRC_Entry_ctorElim___redArg(x_3, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_ctorElim___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_IR_StructRC_Entry_ctorElim(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_2); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_persistent_elim(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_IR_StructRC_Entry_ctorElim___redArg(x_2, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_persistent_elim___redArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_IR_StructRC_Entry_ctorElim___redArg(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_unknownUnion_elim(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_IR_StructRC_Entry_ctorElim___redArg(x_2, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_unknownUnion_elim___redArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_IR_StructRC_Entry_ctorElim___redArg(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_struct_elim(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_IR_StructRC_Entry_ctorElim___redArg(x_2, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_struct_elim___redArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_IR_StructRC_Entry_ctorElim___redArg(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_ref_elim(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_IR_StructRC_Entry_ctorElim___redArg(x_2, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_ref_elim___redArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_IR_StructRC_Entry_ctorElim___redArg(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_IR_StructRC_instInhabitedEntry_default() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_StructRC_instInhabitedEntry() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Nat_cast___at___00Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0_spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" := ", 4, 4); +return x_1; +} +} +static lean_object* _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__4; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("toArray", 7, 7); +return x_1; +} +} +static lean_object* _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__2; +x_2 = lean_box(0); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__5; +x_2 = l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__3; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(11u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(",", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(1); +x_2 = l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__2; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___00Array_Array_repr___at___00instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1_spec__3_spec__5___lam__0(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_unsigned_to_nat(0u); +x_3 = l_Lean_IR_StructRC_instReprProjEntry_repr(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_List_foldl___at___00List_foldl___at___00Std_Format_joinSep___at___00Array_Array_repr___at___00instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1_spec__3_spec__5_spec__7_spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_dec(x_1); +return x_2; +} +else +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_3, 1); +lean_inc(x_1); +lean_ctor_set_tag(x_3, 5); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 0, x_2); +x_7 = lean_unsigned_to_nat(0u); +x_8 = l_Lean_IR_StructRC_instReprProjEntry_repr(x_5, x_7); +x_9 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_9, 0, x_3); +lean_ctor_set(x_9, 1, x_8); +x_2 = x_9; +x_3 = x_6; +goto _start; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_11 = lean_ctor_get(x_3, 0); +x_12 = lean_ctor_get(x_3, 1); +lean_inc(x_12); +lean_inc(x_11); +lean_dec(x_3); +lean_inc(x_1); +x_13 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_13, 0, x_2); +lean_ctor_set(x_13, 1, x_1); +x_14 = lean_unsigned_to_nat(0u); +x_15 = l_Lean_IR_StructRC_instReprProjEntry_repr(x_11, x_14); +x_16 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_16, 0, x_13); +lean_ctor_set(x_16, 1, x_15); +x_2 = x_16; +x_3 = x_12; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_List_foldl___at___00Std_Format_joinSep___at___00Array_Array_repr___at___00instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1_spec__3_spec__5_spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_dec(x_1); +return x_2; +} +else +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_3, 1); +lean_inc(x_1); +lean_ctor_set_tag(x_3, 5); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 0, x_2); +x_7 = lean_unsigned_to_nat(0u); +x_8 = l_Lean_IR_StructRC_instReprProjEntry_repr(x_5, x_7); +x_9 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_9, 0, x_3); +lean_ctor_set(x_9, 1, x_8); +x_10 = l_List_foldl___at___00List_foldl___at___00Std_Format_joinSep___at___00Array_Array_repr___at___00instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1_spec__3_spec__5_spec__7_spec__8(x_1, x_9, x_6); +return x_10; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_11 = lean_ctor_get(x_3, 0); +x_12 = lean_ctor_get(x_3, 1); +lean_inc(x_12); +lean_inc(x_11); +lean_dec(x_3); +lean_inc(x_1); +x_13 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_13, 0, x_2); +lean_ctor_set(x_13, 1, x_1); +x_14 = lean_unsigned_to_nat(0u); +x_15 = l_Lean_IR_StructRC_instReprProjEntry_repr(x_11, x_14); +x_16 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_16, 0, x_13); +lean_ctor_set(x_16, 1, x_15); +x_17 = l_List_foldl___at___00List_foldl___at___00Std_Format_joinSep___at___00Array_Array_repr___at___00instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1_spec__3_spec__5_spec__7_spec__8(x_1, x_16, x_12); +return x_17; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___00Array_Array_repr___at___00instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1_spec__3_spec__5(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; +lean_dec(x_2); +x_3 = lean_box(0); +return x_3; +} +else +{ +lean_object* x_4; +x_4 = lean_ctor_get(x_1, 1); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; +lean_dec(x_2); +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +lean_dec_ref(x_1); +x_6 = l_Std_Format_joinSep___at___00Array_Array_repr___at___00instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1_spec__3_spec__5___lam__0(x_5); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +lean_inc(x_4); +x_7 = lean_ctor_get(x_1, 0); +lean_inc(x_7); +lean_dec_ref(x_1); +x_8 = l_Std_Format_joinSep___at___00Array_Array_repr___at___00instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1_spec__3_spec__5___lam__0(x_7); +x_9 = l_List_foldl___at___00Std_Format_joinSep___at___00Array_Array_repr___at___00instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1_spec__3_spec__5_spec__7(x_2, x_8, x_4); +return x_9; +} +} +} +} +static lean_object* _init_l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("#[", 2, 2); +return x_1; +} +} +static lean_object* _init_l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__0; +x_2 = lean_string_length(x_1); +return x_2; +} +} +static lean_object* _init_l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__5; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__0; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("]", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__4; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("#[]", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__9; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_Array_repr___at___00instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1_spec__3(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = lean_array_get_size(x_1); +x_3 = lean_unsigned_to_nat(0u); +x_4 = lean_nat_dec_eq(x_2, x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_5 = lean_array_to_list(x_1); +x_6 = l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__3; +x_7 = l_Std_Format_joinSep___at___00Array_Array_repr___at___00instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1_spec__3_spec__5(x_5, x_6); +x_8 = l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__6; +x_9 = l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__7; +x_10 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_7); +x_11 = l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__8; +x_12 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set(x_12, 1, x_11); +x_13 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_13, 0, x_8); +lean_ctor_set(x_13, 1, x_12); +x_14 = l_Std_Format_fill(x_13); +return x_14; +} +else +{ +lean_object* x_15; +lean_dec_ref(x_1); +x_15 = l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__10; +return x_15; +} +} +} +static lean_object* _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("size_toArray", 12, 12); +return x_1; +} +} +static lean_object* _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__8; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_", 1, 1); +return x_1; +} +} +static lean_object* _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__10; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("{ ", 2, 2); +return x_1; +} +} +static lean_object* _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__0; +x_2 = lean_string_length(x_1); +return x_2; +} +} +static lean_object* _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__13; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__0; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" }", 2, 2); +return x_1; +} +} +static lean_object* _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__12; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg(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; uint8_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_2 = l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__5; +x_3 = l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__6; +x_4 = l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__7; +x_5 = l_Array_Array_repr___at___00instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1_spec__3(x_1); +x_6 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_6, 0, x_4); +lean_ctor_set(x_6, 1, x_5); +x_7 = 0; +x_8 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_8, 0, x_6); +lean_ctor_set_uint8(x_8, sizeof(void*)*1, x_7); +x_9 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_9, 0, x_3); +lean_ctor_set(x_9, 1, x_8); +x_10 = l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__2; +x_11 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_11, 0, x_9); +lean_ctor_set(x_11, 1, x_10); +x_12 = lean_box(1); +x_13 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_13, 0, x_11); +lean_ctor_set(x_13, 1, x_12); +x_14 = l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__9; +x_15 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_14); +x_16 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_2); +x_17 = l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__11; +x_18 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +x_19 = l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__14; +x_20 = l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__15; +x_21 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_18); +x_22 = l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__16; +x_23 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +x_24 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_24, 0, x_19); +lean_ctor_set(x_24, 1, x_23); +x_25 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set_uint8(x_25, sizeof(void*)*1, x_7); +return x_25; +} +} +LEAN_EXPORT lean_object* l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___00Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0_spec__0___lam__0(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_unsigned_to_nat(0u); +x_3 = l_Lean_IR_instReprIRType_repr(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_List_foldl___at___00List_foldl___at___00Std_Format_joinSep___at___00Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0_spec__0_spec__1_spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_dec(x_1); +return x_2; +} +else +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_3, 1); +lean_inc(x_1); +lean_ctor_set_tag(x_3, 5); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 0, x_2); +x_7 = lean_unsigned_to_nat(0u); +x_8 = l_Lean_IR_instReprIRType_repr(x_5, x_7); +x_9 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_9, 0, x_3); +lean_ctor_set(x_9, 1, x_8); +x_2 = x_9; +x_3 = x_6; +goto _start; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_11 = lean_ctor_get(x_3, 0); +x_12 = lean_ctor_get(x_3, 1); +lean_inc(x_12); +lean_inc(x_11); +lean_dec(x_3); +lean_inc(x_1); +x_13 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_13, 0, x_2); +lean_ctor_set(x_13, 1, x_1); +x_14 = lean_unsigned_to_nat(0u); +x_15 = l_Lean_IR_instReprIRType_repr(x_11, x_14); +x_16 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_16, 0, x_13); +lean_ctor_set(x_16, 1, x_15); +x_2 = x_16; +x_3 = x_12; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_List_foldl___at___00Std_Format_joinSep___at___00Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0_spec__0_spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_dec(x_1); +return x_2; +} +else +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_3, 1); +lean_inc(x_1); +lean_ctor_set_tag(x_3, 5); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 0, x_2); +x_7 = lean_unsigned_to_nat(0u); +x_8 = l_Lean_IR_instReprIRType_repr(x_5, x_7); +x_9 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_9, 0, x_3); +lean_ctor_set(x_9, 1, x_8); +x_10 = l_List_foldl___at___00List_foldl___at___00Std_Format_joinSep___at___00Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0_spec__0_spec__1_spec__4(x_1, x_9, x_6); +return x_10; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_11 = lean_ctor_get(x_3, 0); +x_12 = lean_ctor_get(x_3, 1); +lean_inc(x_12); +lean_inc(x_11); +lean_dec(x_3); +lean_inc(x_1); +x_13 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_13, 0, x_2); +lean_ctor_set(x_13, 1, x_1); +x_14 = lean_unsigned_to_nat(0u); +x_15 = l_Lean_IR_instReprIRType_repr(x_11, x_14); +x_16 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_16, 0, x_13); +lean_ctor_set(x_16, 1, x_15); +x_17 = l_List_foldl___at___00List_foldl___at___00Std_Format_joinSep___at___00Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0_spec__0_spec__1_spec__4(x_1, x_16, x_12); +return x_17; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___00Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0_spec__0(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; +lean_dec(x_2); +x_3 = lean_box(0); +return x_3; +} +else +{ +lean_object* x_4; +x_4 = lean_ctor_get(x_1, 1); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; +lean_dec(x_2); +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +lean_dec_ref(x_1); +x_6 = l_Std_Format_joinSep___at___00Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0_spec__0___lam__0(x_5); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +lean_inc(x_4); +x_7 = lean_ctor_get(x_1, 0); +lean_inc(x_7); +lean_dec_ref(x_1); +x_8 = l_Std_Format_joinSep___at___00Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0_spec__0___lam__0(x_7); +x_9 = l_List_foldl___at___00Std_Format_joinSep___at___00Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0_spec__0_spec__1(x_2, x_8, x_4); +return x_9; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = lean_array_get_size(x_1); +x_3 = lean_unsigned_to_nat(0u); +x_4 = lean_nat_dec_eq(x_2, x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_5 = lean_array_to_list(x_1); +x_6 = l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__3; +x_7 = l_Std_Format_joinSep___at___00Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0_spec__0(x_5, x_6); +x_8 = l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__6; +x_9 = l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__7; +x_10 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_7); +x_11 = l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__8; +x_12 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set(x_12, 1, x_11); +x_13 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_13, 0, x_8); +lean_ctor_set(x_13, 1, x_12); +x_14 = l_Std_Format_fill(x_13); +return x_14; +} +else +{ +lean_object* x_15; +lean_dec_ref(x_1); +x_15 = l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__10; +return x_15; +} +} +} +static lean_object* _init_l_Lean_IR_StructRC_instReprEntry_repr___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.IR.StructRC.Entry.persistent", 33, 33); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_StructRC_instReprEntry_repr___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_IR_StructRC_instReprEntry_repr___closed__0; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_IR_StructRC_instReprEntry_repr___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.IR.StructRC.Entry.unknownUnion", 35, 35); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_StructRC_instReprEntry_repr___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_IR_StructRC_instReprEntry_repr___closed__2; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_IR_StructRC_instReprEntry_repr___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(1); +x_2 = l_Lean_IR_StructRC_instReprEntry_repr___closed__3; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_IR_StructRC_instReprEntry_repr___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.IR.StructRC.Entry.struct", 29, 29); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_StructRC_instReprEntry_repr___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_IR_StructRC_instReprEntry_repr___closed__5; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_IR_StructRC_instReprEntry_repr___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(1); +x_2 = l_Lean_IR_StructRC_instReprEntry_repr___closed__6; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_IR_StructRC_instReprEntry_repr___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.IR.StructRC.Entry.ref", 26, 26); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_StructRC_instReprEntry_repr___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_IR_StructRC_instReprEntry_repr___closed__8; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_IR_StructRC_instReprEntry_repr___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(1); +x_2 = l_Lean_IR_StructRC_instReprEntry_repr___closed__9; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_instReprEntry_repr(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_21; +switch (lean_obj_tag(x_1)) { +case 0: +{ +lean_object* x_28; uint8_t x_29; +x_28 = lean_unsigned_to_nat(1024u); +x_29 = lean_nat_dec_le(x_28, x_2); +if (x_29 == 0) +{ +lean_object* x_30; +x_30 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__2; +x_21 = x_30; +goto block_27; +} +else +{ +lean_object* x_31; +x_31 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__3; +x_21 = x_31; +goto block_27; +} +} +case 1: +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_50; uint8_t x_51; +x_32 = lean_ctor_get(x_1, 0); +lean_inc_ref(x_32); +x_33 = lean_ctor_get(x_1, 1); +lean_inc(x_33); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + x_34 = x_1; +} else { + lean_dec_ref(x_1); + x_34 = lean_box(0); +} +x_50 = lean_unsigned_to_nat(1024u); +x_51 = lean_nat_dec_le(x_50, x_2); +if (x_51 == 0) +{ +lean_object* x_52; +x_52 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__2; +x_35 = x_52; +goto block_49; +} +else +{ +lean_object* x_53; +x_53 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__3; +x_35 = x_53; +goto block_49; +} +block_49: +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_36 = lean_box(1); +x_37 = l_Lean_IR_StructRC_instReprEntry_repr___closed__4; +x_38 = l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0(x_32); +if (lean_is_scalar(x_34)) { + x_39 = lean_alloc_ctor(5, 2, 0); +} else { + x_39 = x_34; + lean_ctor_set_tag(x_39, 5); +} +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +x_40 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_36); +x_41 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__10; +x_42 = lean_int_dec_lt(x_33, x_41); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; +x_43 = l_Int_repr(x_33); +lean_dec(x_33); +x_44 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_44, 0, x_43); +x_12 = x_35; +x_13 = x_40; +x_14 = x_44; +goto block_20; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_45 = lean_unsigned_to_nat(1024u); +x_46 = l_Int_repr(x_33); +lean_dec(x_33); +x_47 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_47, 0, x_46); +x_48 = l_Repr_addAppParen(x_47, x_45); +x_12 = x_35; +x_13 = x_40; +x_14 = x_48; +goto block_20; +} +} +} +case 2: +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_74; uint8_t x_75; +x_54 = lean_ctor_get(x_1, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_55); +x_56 = lean_ctor_get(x_1, 2); +lean_inc_ref(x_56); +lean_dec_ref(x_1); +x_74 = lean_unsigned_to_nat(1024u); +x_75 = lean_nat_dec_le(x_74, x_2); +if (x_75 == 0) +{ +lean_object* x_76; +x_76 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__2; +x_57 = x_76; +goto block_73; +} +else +{ +lean_object* x_77; +x_77 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__3; +x_57 = x_77; +goto block_73; +} +block_73: +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; uint8_t x_70; lean_object* x_71; lean_object* x_72; +x_58 = lean_box(1); +x_59 = l_Lean_IR_StructRC_instReprEntry_repr___closed__7; +x_60 = l_Nat_reprFast(x_54); +x_61 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_61, 0, x_60); +x_62 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_62, 0, x_59); +lean_ctor_set(x_62, 1, x_61); +x_63 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_63, 0, x_62); +lean_ctor_set(x_63, 1, x_58); +x_64 = l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0(x_55); +x_65 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_64); +x_66 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_66, 0, x_65); +lean_ctor_set(x_66, 1, x_58); +x_67 = l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg(x_56); +x_68 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_68, 0, x_66); +lean_ctor_set(x_68, 1, x_67); +x_69 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_69, 0, x_57); +lean_ctor_set(x_69, 1, x_68); +x_70 = 0; +x_71 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_71, 0, x_69); +lean_ctor_set_uint8(x_71, sizeof(void*)*1, x_70); +x_72 = l_Repr_addAppParen(x_71, x_2); +return x_72; +} +} +default: +{ +uint8_t x_78; lean_object* x_79; lean_object* x_80; lean_object* x_95; uint8_t x_96; +x_78 = lean_ctor_get_uint8(x_1, sizeof(void*)*1); +x_79 = lean_ctor_get(x_1, 0); +lean_inc(x_79); +lean_dec_ref(x_1); +x_95 = lean_unsigned_to_nat(1024u); +x_96 = lean_nat_dec_le(x_95, x_2); +if (x_96 == 0) +{ +lean_object* x_97; +x_97 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__2; +x_80 = x_97; +goto block_94; +} +else +{ +lean_object* x_98; +x_98 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__3; +x_80 = x_98; +goto block_94; +} +block_94: +{ +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; uint8_t x_87; +x_81 = lean_box(1); +x_82 = l_Lean_IR_StructRC_instReprEntry_repr___closed__10; +x_83 = l_Bool_repr___redArg(x_78); +x_84 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_84, 0, x_82); +lean_ctor_set(x_84, 1, x_83); +x_85 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_85, 0, x_84); +lean_ctor_set(x_85, 1, x_81); +x_86 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__10; +x_87 = lean_int_dec_lt(x_79, x_86); +if (x_87 == 0) +{ +lean_object* x_88; lean_object* x_89; +x_88 = l_Int_repr(x_79); +lean_dec(x_79); +x_89 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_89, 0, x_88); +x_3 = x_85; +x_4 = x_80; +x_5 = x_89; +goto block_11; +} +else +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_90 = lean_unsigned_to_nat(1024u); +x_91 = l_Int_repr(x_79); +lean_dec(x_79); +x_92 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_92, 0, x_91); +x_93 = l_Repr_addAppParen(x_92, x_90); +x_3 = x_85; +x_4 = x_80; +x_5 = x_93; +goto block_11; +} +} +} +} +block_11: +{ +lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; +x_6 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_6, 0, x_3); +lean_ctor_set(x_6, 1, x_5); +x_7 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_7, 0, x_4); +lean_ctor_set(x_7, 1, x_6); +x_8 = 0; +x_9 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_9, 0, x_7); +lean_ctor_set_uint8(x_9, sizeof(void*)*1, x_8); +x_10 = l_Repr_addAppParen(x_9, x_2); +return x_10; +} +block_20: +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; +x_15 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_14); +x_16 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_16, 0, x_12); +lean_ctor_set(x_16, 1, x_15); +x_17 = 0; +x_18 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set_uint8(x_18, sizeof(void*)*1, x_17); +x_19 = l_Repr_addAppParen(x_18, x_2); +return x_19; +} +block_27: +{ +lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; +x_22 = l_Lean_IR_StructRC_instReprEntry_repr___closed__1; +x_23 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +x_24 = 0; +x_25 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set_uint8(x_25, sizeof(void*)*1, x_24); +x_26 = l_Repr_addAppParen(x_25, x_2); +return x_26; +} +} +} +LEAN_EXPORT lean_object* l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_instReprEntry_repr___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_IR_StructRC_instReprEntry_repr(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_IR_StructRC_instReprEntry___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_IR_StructRC_instReprEntry_repr___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_StructRC_instReprEntry() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_IR_StructRC_instReprEntry___closed__0; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_Const_getD___at___00Lean_IR_StructRC_Context_renameVar_spec__0___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_4 = lean_ctor_get(x_1, 1); +x_5 = lean_ctor_get(x_1, 2); +x_6 = lean_ctor_get(x_1, 3); +x_7 = lean_ctor_get(x_1, 4); +x_8 = lean_nat_dec_lt(x_2, x_4); +if (x_8 == 0) +{ +uint8_t x_9; +x_9 = lean_nat_dec_eq(x_2, x_4); +if (x_9 == 0) +{ +x_1 = x_7; +goto _start; +} +else +{ +lean_inc(x_5); +return x_5; +} +} +else +{ +x_1 = x_6; +goto _start; +} +} +else +{ +lean_inc(x_3); +return x_3; +} +} +} +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_Const_getD___at___00Lean_IR_StructRC_Context_renameVar_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Std_DTreeMap_Internal_Impl_Const_getD___at___00Lean_IR_StructRC_Context_renameVar_spec__0___redArg(x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_renameVar(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_ctor_get(x_1, 2); +lean_inc(x_2); +x_4 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_4, 0, x_2); +x_5 = l_Std_DTreeMap_Internal_Impl_Const_getD___at___00Lean_IR_StructRC_Context_renameVar_spec__0___redArg(x_3, x_2, x_4); +lean_dec_ref(x_4); +lean_dec(x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_Const_getD___at___00Lean_IR_StructRC_Context_renameVar_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Std_DTreeMap_Internal_Impl_Const_getD___at___00Lean_IR_StructRC_Context_renameVar_spec__0(x_1, x_2, x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_renameVar___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_IR_StructRC_Context_renameVar(x_1, x_2); +lean_dec_ref(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_Const_getD___at___00Lean_IR_StructRC_Context_renameVar_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_DTreeMap_Internal_Impl_Const_getD___at___00Lean_IR_StructRC_Context_renameVar_spec__0___redArg(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_Const_get_x3f___at___00Lean_IR_StructRC_Context_renameVar_x21_spec__0___redArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_3 = lean_ctor_get(x_1, 1); +x_4 = lean_ctor_get(x_1, 2); +x_5 = lean_ctor_get(x_1, 3); +x_6 = lean_ctor_get(x_1, 4); +x_7 = lean_nat_dec_lt(x_2, x_3); +if (x_7 == 0) +{ +uint8_t x_8; +x_8 = lean_nat_dec_eq(x_2, x_3); +if (x_8 == 0) +{ +x_1 = x_6; +goto _start; +} +else +{ +lean_object* x_10; +lean_inc(x_4); +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_4); +return x_10; +} +} +else +{ +x_1 = x_5; +goto _start; +} +} +else +{ +lean_object* x_12; +x_12 = lean_box(0); +return x_12; +} +} +} +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_Const_get_x3f___at___00Lean_IR_StructRC_Context_renameVar_x21_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_DTreeMap_Internal_Impl_Const_get_x3f___at___00Lean_IR_StructRC_Context_renameVar_x21_spec__0___redArg(x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_StructRC_Context_renameVar_x21_spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_unsigned_to_nat(0u); +x_3 = lean_panic_fn(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_IR_StructRC_Context_renameVar_x21___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Compiler.IR.StructRC", 25, 25); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_StructRC_Context_renameVar_x21___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.IR.StructRC.Context.renameVar!", 35, 35); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_StructRC_Context_renameVar_x21___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("malformed IR", 12, 12); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_StructRC_Context_renameVar_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_Lean_IR_StructRC_Context_renameVar_x21___closed__2; +x_2 = lean_unsigned_to_nat(20u); +x_3 = lean_unsigned_to_nat(80u); +x_4 = l_Lean_IR_StructRC_Context_renameVar_x21___closed__1; +x_5 = l_Lean_IR_StructRC_Context_renameVar_x21___closed__0; +x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_renameVar_x21(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = lean_ctor_get(x_1, 2); +x_4 = l_Std_DTreeMap_Internal_Impl_Const_get_x3f___at___00Lean_IR_StructRC_Context_renameVar_x21_spec__0___redArg(x_3, x_2); +if (lean_obj_tag(x_4) == 0) +{ +lean_inc(x_2); +return x_2; +} +else +{ +lean_object* x_5; +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +lean_dec_ref(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); +lean_dec_ref(x_5); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; +x_7 = l_Lean_IR_StructRC_Context_renameVar_x21___closed__3; +x_8 = l_panic___at___00Lean_IR_StructRC_Context_renameVar_x21_spec__1(x_7); +return x_8; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_Const_get_x3f___at___00Lean_IR_StructRC_Context_renameVar_x21_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_DTreeMap_Internal_Impl_Const_get_x3f___at___00Lean_IR_StructRC_Context_renameVar_x21_spec__0(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_Const_get_x3f___at___00Lean_IR_StructRC_Context_renameVar_x21_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_DTreeMap_Internal_Impl_Const_get_x3f___at___00Lean_IR_StructRC_Context_renameVar_x21_spec__0___redArg(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_renameVar_x21___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_IR_StructRC_Context_renameVar_x21(x_1, x_2); +lean_dec(x_2); +lean_dec_ref(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_renameArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_ctor_get(x_2, 0); +x_4 = lean_ctor_get(x_1, 2); +x_5 = l_Std_DTreeMap_Internal_Impl_Const_getD___at___00Lean_IR_StructRC_Context_renameVar_spec__0___redArg(x_4, x_3, x_2); +return x_5; +} +else +{ +return x_2; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_renameArg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_IR_StructRC_Context_renameArg(x_1, x_2); +lean_dec(x_2); +lean_dec_ref(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_StructRC_Context_renameArgs_spec__0(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_usize_dec_lt(x_3, x_2); +if (x_5 == 0) +{ +return x_4; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; size_t x_10; size_t x_11; lean_object* x_12; +x_6 = lean_array_uget(x_4, x_3); +x_7 = lean_unsigned_to_nat(0u); +x_8 = lean_array_uset(x_4, x_3, x_7); +x_9 = l_Lean_IR_StructRC_Context_renameArg(x_1, x_6); +lean_dec(x_6); +x_10 = 1; +x_11 = lean_usize_add(x_3, x_10); +x_12 = lean_array_uset(x_8, x_3, x_9); +x_3 = x_11; +x_4 = x_12; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_renameArgs(lean_object* x_1, lean_object* x_2) { +_start: +{ +size_t x_3; size_t x_4; lean_object* x_5; +x_3 = lean_array_size(x_2); +x_4 = 0; +x_5 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_StructRC_Context_renameArgs_spec__0(x_1, x_3, x_4, x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_renameArgs___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_IR_StructRC_Context_renameArgs(x_1, x_2); +lean_dec_ref(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_StructRC_Context_renameArgs_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; lean_object* x_7; +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_StructRC_Context_renameArgs_spec__0(x_1, x_5, x_6, x_4); +lean_dec_ref(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_insert___at___00Lean_IR_StructRC_Context_insertRename_spec__0___redArg(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_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_3, 1); +lean_inc(x_5); +x_6 = lean_ctor_get(x_3, 2); +lean_inc(x_6); +x_7 = lean_ctor_get(x_3, 3); +lean_inc(x_7); +x_8 = lean_ctor_get(x_3, 4); +lean_inc(x_8); +if (lean_is_exclusive(x_3)) { + lean_ctor_release(x_3, 0); + lean_ctor_release(x_3, 1); + lean_ctor_release(x_3, 2); + lean_ctor_release(x_3, 3); + lean_ctor_release(x_3, 4); + x_9 = x_3; +} else { + lean_dec_ref(x_3); + x_9 = lean_box(0); +} +x_10 = lean_nat_dec_lt(x_1, x_5); +if (x_10 == 0) +{ +uint8_t x_11; +x_11 = lean_nat_dec_eq(x_1, x_5); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_4); +x_12 = l_Std_DTreeMap_Internal_Impl_insert___at___00Lean_IR_StructRC_Context_insertRename_spec__0___redArg(x_1, x_2, x_8); +x_13 = lean_unsigned_to_nat(1u); +if (lean_obj_tag(x_7) == 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; uint8_t x_22; +x_14 = lean_ctor_get(x_7, 0); +x_15 = lean_ctor_get(x_12, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_12, 1); +lean_inc(x_16); +x_17 = lean_ctor_get(x_12, 2); +lean_inc(x_17); +x_18 = lean_ctor_get(x_12, 3); +lean_inc(x_18); +x_19 = lean_ctor_get(x_12, 4); +lean_inc(x_19); +x_20 = lean_unsigned_to_nat(3u); +x_21 = lean_nat_mul(x_20, x_14); +x_22 = lean_nat_dec_lt(x_21, x_15); +lean_dec(x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +x_23 = lean_nat_add(x_13, x_14); +x_24 = lean_nat_add(x_23, x_15); +lean_dec(x_15); +lean_dec(x_23); +if (lean_is_scalar(x_9)) { + x_25 = lean_alloc_ctor(0, 5, 0); +} else { + x_25 = x_9; +} +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_5); +lean_ctor_set(x_25, 2, x_6); +lean_ctor_set(x_25, 3, x_7); +lean_ctor_set(x_25, 4, x_12); +return x_25; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +if (lean_is_exclusive(x_12)) { + lean_ctor_release(x_12, 0); + lean_ctor_release(x_12, 1); + lean_ctor_release(x_12, 2); + lean_ctor_release(x_12, 3); + lean_ctor_release(x_12, 4); + x_26 = x_12; +} else { + lean_dec_ref(x_12); + x_26 = lean_box(0); +} +x_27 = lean_ctor_get(x_18, 0); +x_28 = lean_ctor_get(x_18, 1); +x_29 = lean_ctor_get(x_18, 2); +x_30 = lean_ctor_get(x_18, 3); +x_31 = lean_ctor_get(x_18, 4); +x_32 = lean_ctor_get(x_19, 0); +x_33 = lean_unsigned_to_nat(2u); +x_34 = lean_nat_mul(x_33, x_32); +x_35 = lean_nat_dec_lt(x_27, x_34); +lean_dec(x_34); +if (x_35 == 0) +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_46; +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +if (lean_is_exclusive(x_18)) { + lean_ctor_release(x_18, 0); + lean_ctor_release(x_18, 1); + lean_ctor_release(x_18, 2); + lean_ctor_release(x_18, 3); + lean_ctor_release(x_18, 4); + x_36 = x_18; +} else { + lean_dec_ref(x_18); + x_36 = lean_box(0); +} +x_37 = lean_nat_add(x_13, x_14); +x_38 = lean_nat_add(x_37, x_15); +lean_dec(x_15); +if (lean_obj_tag(x_30) == 0) +{ +lean_object* x_53; +x_53 = lean_ctor_get(x_30, 0); +lean_inc(x_53); +x_46 = x_53; +goto block_52; +} +else +{ +lean_object* x_54; +x_54 = lean_unsigned_to_nat(0u); +x_46 = x_54; +goto block_52; +} +block_45: +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_nat_add(x_40, x_41); +lean_dec(x_41); +lean_dec(x_40); +if (lean_is_scalar(x_36)) { + x_43 = lean_alloc_ctor(0, 5, 0); +} else { + x_43 = x_36; +} +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_16); +lean_ctor_set(x_43, 2, x_17); +lean_ctor_set(x_43, 3, x_31); +lean_ctor_set(x_43, 4, x_19); +if (lean_is_scalar(x_26)) { + x_44 = lean_alloc_ctor(0, 5, 0); +} else { + x_44 = x_26; +} +lean_ctor_set(x_44, 0, x_38); +lean_ctor_set(x_44, 1, x_28); +lean_ctor_set(x_44, 2, x_29); +lean_ctor_set(x_44, 3, x_39); +lean_ctor_set(x_44, 4, x_43); +return x_44; +} +block_52: +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_nat_add(x_37, x_46); +lean_dec(x_46); +lean_dec(x_37); +if (lean_is_scalar(x_9)) { + x_48 = lean_alloc_ctor(0, 5, 0); +} else { + x_48 = x_9; +} +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_5); +lean_ctor_set(x_48, 2, x_6); +lean_ctor_set(x_48, 3, x_7); +lean_ctor_set(x_48, 4, x_30); +x_49 = lean_nat_add(x_13, x_32); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_50; +x_50 = lean_ctor_get(x_31, 0); +lean_inc(x_50); +x_39 = x_48; +x_40 = x_49; +x_41 = x_50; +goto block_45; +} +else +{ +lean_object* x_51; +x_51 = lean_unsigned_to_nat(0u); +x_39 = x_48; +x_40 = x_49; +x_41 = x_51; +goto block_45; +} +} +} +else +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; +lean_dec(x_9); +x_55 = lean_nat_add(x_13, x_14); +x_56 = lean_nat_add(x_55, x_15); +lean_dec(x_15); +x_57 = lean_nat_add(x_55, x_27); +lean_dec(x_55); +lean_inc_ref(x_7); +if (lean_is_scalar(x_26)) { + x_58 = lean_alloc_ctor(0, 5, 0); +} else { + x_58 = x_26; +} +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_5); +lean_ctor_set(x_58, 2, x_6); +lean_ctor_set(x_58, 3, x_7); +lean_ctor_set(x_58, 4, x_18); +x_59 = !lean_is_exclusive(x_7); +if (x_59 == 0) +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_60 = lean_ctor_get(x_7, 4); +lean_dec(x_60); +x_61 = lean_ctor_get(x_7, 3); +lean_dec(x_61); +x_62 = lean_ctor_get(x_7, 2); +lean_dec(x_62); +x_63 = lean_ctor_get(x_7, 1); +lean_dec(x_63); +x_64 = lean_ctor_get(x_7, 0); +lean_dec(x_64); +lean_ctor_set(x_7, 4, x_19); +lean_ctor_set(x_7, 3, x_58); +lean_ctor_set(x_7, 2, x_17); +lean_ctor_set(x_7, 1, x_16); +lean_ctor_set(x_7, 0, x_56); +return x_7; +} +else +{ +lean_object* x_65; +lean_dec(x_7); +x_65 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_65, 0, x_56); +lean_ctor_set(x_65, 1, x_16); +lean_ctor_set(x_65, 2, x_17); +lean_ctor_set(x_65, 3, x_58); +lean_ctor_set(x_65, 4, x_19); +return x_65; +} +} +} +} +else +{ +lean_object* x_66; +x_66 = lean_ctor_get(x_12, 3); +lean_inc(x_66); +if (lean_obj_tag(x_66) == 0) +{ +uint8_t x_67; +x_67 = !lean_is_exclusive(x_12); +if (x_67 == 0) +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; +x_68 = lean_ctor_get(x_12, 4); +x_69 = lean_ctor_get(x_12, 3); +lean_dec(x_69); +x_70 = lean_ctor_get(x_12, 0); +lean_dec(x_70); +x_71 = !lean_is_exclusive(x_66); +if (x_71 == 0) +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_72 = lean_ctor_get(x_66, 1); +x_73 = lean_ctor_get(x_66, 2); +x_74 = lean_ctor_get(x_66, 4); +lean_dec(x_74); +x_75 = lean_ctor_get(x_66, 3); +lean_dec(x_75); +x_76 = lean_ctor_get(x_66, 0); +lean_dec(x_76); +x_77 = lean_unsigned_to_nat(3u); +lean_inc_n(x_68, 2); +lean_ctor_set(x_66, 4, x_68); +lean_ctor_set(x_66, 3, x_68); +lean_ctor_set(x_66, 2, x_6); +lean_ctor_set(x_66, 1, x_5); +lean_ctor_set(x_66, 0, x_13); +lean_inc(x_68); +lean_ctor_set(x_12, 3, x_68); +lean_ctor_set(x_12, 0, x_13); +if (lean_is_scalar(x_9)) { + x_78 = lean_alloc_ctor(0, 5, 0); +} else { + x_78 = x_9; +} +lean_ctor_set(x_78, 0, x_77); +lean_ctor_set(x_78, 1, x_72); +lean_ctor_set(x_78, 2, x_73); +lean_ctor_set(x_78, 3, x_66); +lean_ctor_set(x_78, 4, x_12); +return x_78; +} +else +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_79 = lean_ctor_get(x_66, 1); +x_80 = lean_ctor_get(x_66, 2); +lean_inc(x_80); +lean_inc(x_79); +lean_dec(x_66); +x_81 = lean_unsigned_to_nat(3u); +lean_inc_n(x_68, 2); +x_82 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_82, 0, x_13); +lean_ctor_set(x_82, 1, x_5); +lean_ctor_set(x_82, 2, x_6); +lean_ctor_set(x_82, 3, x_68); +lean_ctor_set(x_82, 4, x_68); +lean_inc(x_68); +lean_ctor_set(x_12, 3, x_68); +lean_ctor_set(x_12, 0, x_13); +if (lean_is_scalar(x_9)) { + x_83 = lean_alloc_ctor(0, 5, 0); +} else { + x_83 = x_9; +} +lean_ctor_set(x_83, 0, x_81); +lean_ctor_set(x_83, 1, x_79); +lean_ctor_set(x_83, 2, x_80); +lean_ctor_set(x_83, 3, x_82); +lean_ctor_set(x_83, 4, x_12); +return x_83; +} +} +else +{ +lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_84 = lean_ctor_get(x_12, 4); +x_85 = lean_ctor_get(x_12, 1); +x_86 = lean_ctor_get(x_12, 2); +lean_inc(x_84); +lean_inc(x_86); +lean_inc(x_85); +lean_dec(x_12); +x_87 = lean_ctor_get(x_66, 1); +lean_inc(x_87); +x_88 = lean_ctor_get(x_66, 2); +lean_inc(x_88); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + lean_ctor_release(x_66, 2); + lean_ctor_release(x_66, 3); + lean_ctor_release(x_66, 4); + x_89 = x_66; +} else { + lean_dec_ref(x_66); + x_89 = lean_box(0); +} +x_90 = lean_unsigned_to_nat(3u); +lean_inc_n(x_84, 2); +if (lean_is_scalar(x_89)) { + x_91 = lean_alloc_ctor(0, 5, 0); +} else { + x_91 = x_89; +} +lean_ctor_set(x_91, 0, x_13); +lean_ctor_set(x_91, 1, x_5); +lean_ctor_set(x_91, 2, x_6); +lean_ctor_set(x_91, 3, x_84); +lean_ctor_set(x_91, 4, x_84); +lean_inc(x_84); +x_92 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_92, 0, x_13); +lean_ctor_set(x_92, 1, x_85); +lean_ctor_set(x_92, 2, x_86); +lean_ctor_set(x_92, 3, x_84); +lean_ctor_set(x_92, 4, x_84); +if (lean_is_scalar(x_9)) { + x_93 = lean_alloc_ctor(0, 5, 0); +} else { + x_93 = x_9; +} +lean_ctor_set(x_93, 0, x_90); +lean_ctor_set(x_93, 1, x_87); +lean_ctor_set(x_93, 2, x_88); +lean_ctor_set(x_93, 3, x_91); +lean_ctor_set(x_93, 4, x_92); +return x_93; +} +} +else +{ +lean_object* x_94; +x_94 = lean_ctor_get(x_12, 4); +lean_inc(x_94); +if (lean_obj_tag(x_94) == 0) +{ +uint8_t x_95; +x_95 = !lean_is_exclusive(x_12); +if (x_95 == 0) +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_96 = lean_ctor_get(x_12, 1); +x_97 = lean_ctor_get(x_12, 2); +x_98 = lean_ctor_get(x_12, 4); +lean_dec(x_98); +x_99 = lean_ctor_get(x_12, 3); +lean_dec(x_99); +x_100 = lean_ctor_get(x_12, 0); +lean_dec(x_100); +x_101 = lean_unsigned_to_nat(3u); +lean_ctor_set(x_12, 4, x_66); +lean_ctor_set(x_12, 2, x_6); +lean_ctor_set(x_12, 1, x_5); +lean_ctor_set(x_12, 0, x_13); +if (lean_is_scalar(x_9)) { + x_102 = lean_alloc_ctor(0, 5, 0); +} else { + x_102 = x_9; +} +lean_ctor_set(x_102, 0, x_101); +lean_ctor_set(x_102, 1, x_96); +lean_ctor_set(x_102, 2, x_97); +lean_ctor_set(x_102, 3, x_12); +lean_ctor_set(x_102, 4, x_94); +return x_102; +} +else +{ +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_103 = lean_ctor_get(x_12, 1); +x_104 = lean_ctor_get(x_12, 2); +lean_inc(x_104); +lean_inc(x_103); +lean_dec(x_12); +x_105 = lean_unsigned_to_nat(3u); +x_106 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_106, 0, x_13); +lean_ctor_set(x_106, 1, x_5); +lean_ctor_set(x_106, 2, x_6); +lean_ctor_set(x_106, 3, x_66); +lean_ctor_set(x_106, 4, x_66); +if (lean_is_scalar(x_9)) { + x_107 = lean_alloc_ctor(0, 5, 0); +} else { + x_107 = x_9; +} +lean_ctor_set(x_107, 0, x_105); +lean_ctor_set(x_107, 1, x_103); +lean_ctor_set(x_107, 2, x_104); +lean_ctor_set(x_107, 3, x_106); +lean_ctor_set(x_107, 4, x_94); +return x_107; +} +} +else +{ +lean_object* x_108; lean_object* x_109; +x_108 = lean_unsigned_to_nat(2u); +if (lean_is_scalar(x_9)) { + x_109 = lean_alloc_ctor(0, 5, 0); +} else { + x_109 = x_9; +} +lean_ctor_set(x_109, 0, x_108); +lean_ctor_set(x_109, 1, x_5); +lean_ctor_set(x_109, 2, x_6); +lean_ctor_set(x_109, 3, x_94); +lean_ctor_set(x_109, 4, x_12); +return x_109; +} +} +} +} +else +{ +lean_object* x_110; +lean_dec(x_6); +lean_dec(x_5); +if (lean_is_scalar(x_9)) { + x_110 = lean_alloc_ctor(0, 5, 0); +} else { + x_110 = x_9; +} +lean_ctor_set(x_110, 0, x_4); +lean_ctor_set(x_110, 1, x_1); +lean_ctor_set(x_110, 2, x_2); +lean_ctor_set(x_110, 3, x_7); +lean_ctor_set(x_110, 4, x_8); +return x_110; +} +} +else +{ +lean_object* x_111; lean_object* x_112; +lean_dec(x_4); +x_111 = l_Std_DTreeMap_Internal_Impl_insert___at___00Lean_IR_StructRC_Context_insertRename_spec__0___redArg(x_1, x_2, x_7); +x_112 = lean_unsigned_to_nat(1u); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; uint8_t x_121; +x_113 = lean_ctor_get(x_8, 0); +x_114 = lean_ctor_get(x_111, 0); +lean_inc(x_114); +x_115 = lean_ctor_get(x_111, 1); +lean_inc(x_115); +x_116 = lean_ctor_get(x_111, 2); +lean_inc(x_116); +x_117 = lean_ctor_get(x_111, 3); +lean_inc(x_117); +x_118 = lean_ctor_get(x_111, 4); +lean_inc(x_118); +x_119 = lean_unsigned_to_nat(3u); +x_120 = lean_nat_mul(x_119, x_113); +x_121 = lean_nat_dec_lt(x_120, x_114); +lean_dec(x_120); +if (x_121 == 0) +{ +lean_object* x_122; lean_object* x_123; lean_object* x_124; +lean_dec(x_118); +lean_dec(x_117); +lean_dec(x_116); +lean_dec(x_115); +x_122 = lean_nat_add(x_112, x_114); +lean_dec(x_114); +x_123 = lean_nat_add(x_122, x_113); +lean_dec(x_122); +if (lean_is_scalar(x_9)) { + x_124 = lean_alloc_ctor(0, 5, 0); +} else { + x_124 = x_9; +} +lean_ctor_set(x_124, 0, x_123); +lean_ctor_set(x_124, 1, x_5); +lean_ctor_set(x_124, 2, x_6); +lean_ctor_set(x_124, 3, x_111); +lean_ctor_set(x_124, 4, x_8); +return x_124; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; uint8_t x_134; +if (lean_is_exclusive(x_111)) { + lean_ctor_release(x_111, 0); + lean_ctor_release(x_111, 1); + lean_ctor_release(x_111, 2); + lean_ctor_release(x_111, 3); + lean_ctor_release(x_111, 4); + x_125 = x_111; +} else { + lean_dec_ref(x_111); + x_125 = lean_box(0); +} +x_126 = lean_ctor_get(x_117, 0); +x_127 = lean_ctor_get(x_118, 0); +x_128 = lean_ctor_get(x_118, 1); +x_129 = lean_ctor_get(x_118, 2); +x_130 = lean_ctor_get(x_118, 3); +x_131 = lean_ctor_get(x_118, 4); +x_132 = lean_unsigned_to_nat(2u); +x_133 = lean_nat_mul(x_132, x_126); +x_134 = lean_nat_dec_lt(x_127, x_133); +lean_dec(x_133); +if (x_134 == 0) +{ +lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_145; lean_object* x_146; +lean_inc(x_131); +lean_inc(x_130); +lean_inc(x_129); +lean_inc(x_128); +if (lean_is_exclusive(x_118)) { + lean_ctor_release(x_118, 0); + lean_ctor_release(x_118, 1); + lean_ctor_release(x_118, 2); + lean_ctor_release(x_118, 3); + lean_ctor_release(x_118, 4); + x_135 = x_118; +} else { + lean_dec_ref(x_118); + x_135 = lean_box(0); +} +x_136 = lean_nat_add(x_112, x_114); +lean_dec(x_114); +x_137 = lean_nat_add(x_136, x_113); +lean_dec(x_136); +x_145 = lean_nat_add(x_112, x_126); +if (lean_obj_tag(x_130) == 0) +{ +lean_object* x_153; +x_153 = lean_ctor_get(x_130, 0); +lean_inc(x_153); +x_146 = x_153; +goto block_152; +} +else +{ +lean_object* x_154; +x_154 = lean_unsigned_to_nat(0u); +x_146 = x_154; +goto block_152; +} +block_144: +{ +lean_object* x_141; lean_object* x_142; lean_object* x_143; +x_141 = lean_nat_add(x_139, x_140); +lean_dec(x_140); +lean_dec(x_139); +if (lean_is_scalar(x_135)) { + x_142 = lean_alloc_ctor(0, 5, 0); +} else { + x_142 = x_135; +} +lean_ctor_set(x_142, 0, x_141); +lean_ctor_set(x_142, 1, x_5); +lean_ctor_set(x_142, 2, x_6); +lean_ctor_set(x_142, 3, x_131); +lean_ctor_set(x_142, 4, x_8); +if (lean_is_scalar(x_125)) { + x_143 = lean_alloc_ctor(0, 5, 0); +} else { + x_143 = x_125; +} +lean_ctor_set(x_143, 0, x_137); +lean_ctor_set(x_143, 1, x_128); +lean_ctor_set(x_143, 2, x_129); +lean_ctor_set(x_143, 3, x_138); +lean_ctor_set(x_143, 4, x_142); +return x_143; +} +block_152: +{ +lean_object* x_147; lean_object* x_148; lean_object* x_149; +x_147 = lean_nat_add(x_145, x_146); +lean_dec(x_146); +lean_dec(x_145); +if (lean_is_scalar(x_9)) { + x_148 = lean_alloc_ctor(0, 5, 0); +} else { + x_148 = x_9; +} +lean_ctor_set(x_148, 0, x_147); +lean_ctor_set(x_148, 1, x_115); +lean_ctor_set(x_148, 2, x_116); +lean_ctor_set(x_148, 3, x_117); +lean_ctor_set(x_148, 4, x_130); +x_149 = lean_nat_add(x_112, x_113); +if (lean_obj_tag(x_131) == 0) +{ +lean_object* x_150; +x_150 = lean_ctor_get(x_131, 0); +lean_inc(x_150); +x_138 = x_148; +x_139 = x_149; +x_140 = x_150; +goto block_144; +} +else +{ +lean_object* x_151; +x_151 = lean_unsigned_to_nat(0u); +x_138 = x_148; +x_139 = x_149; +x_140 = x_151; +goto block_144; +} +} +} +else +{ +lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; uint8_t x_160; +lean_dec(x_9); +x_155 = lean_nat_add(x_112, x_114); +lean_dec(x_114); +x_156 = lean_nat_add(x_155, x_113); +lean_dec(x_155); +x_157 = lean_nat_add(x_112, x_113); +x_158 = lean_nat_add(x_157, x_127); +lean_dec(x_157); +lean_inc_ref(x_8); +if (lean_is_scalar(x_125)) { + x_159 = lean_alloc_ctor(0, 5, 0); +} else { + x_159 = x_125; +} +lean_ctor_set(x_159, 0, x_158); +lean_ctor_set(x_159, 1, x_5); +lean_ctor_set(x_159, 2, x_6); +lean_ctor_set(x_159, 3, x_118); +lean_ctor_set(x_159, 4, x_8); +x_160 = !lean_is_exclusive(x_8); +if (x_160 == 0) +{ +lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; +x_161 = lean_ctor_get(x_8, 4); +lean_dec(x_161); +x_162 = lean_ctor_get(x_8, 3); +lean_dec(x_162); +x_163 = lean_ctor_get(x_8, 2); +lean_dec(x_163); +x_164 = lean_ctor_get(x_8, 1); +lean_dec(x_164); +x_165 = lean_ctor_get(x_8, 0); +lean_dec(x_165); +lean_ctor_set(x_8, 4, x_159); +lean_ctor_set(x_8, 3, x_117); +lean_ctor_set(x_8, 2, x_116); +lean_ctor_set(x_8, 1, x_115); +lean_ctor_set(x_8, 0, x_156); +return x_8; +} +else +{ +lean_object* x_166; +lean_dec(x_8); +x_166 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_166, 0, x_156); +lean_ctor_set(x_166, 1, x_115); +lean_ctor_set(x_166, 2, x_116); +lean_ctor_set(x_166, 3, x_117); +lean_ctor_set(x_166, 4, x_159); +return x_166; +} +} +} +} +else +{ +lean_object* x_167; +x_167 = lean_ctor_get(x_111, 3); +lean_inc(x_167); +if (lean_obj_tag(x_167) == 0) +{ +uint8_t x_168; +x_168 = !lean_is_exclusive(x_111); +if (x_168 == 0) +{ +lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; +x_169 = lean_ctor_get(x_111, 4); +x_170 = lean_ctor_get(x_111, 1); +x_171 = lean_ctor_get(x_111, 2); +x_172 = lean_ctor_get(x_111, 3); +lean_dec(x_172); +x_173 = lean_ctor_get(x_111, 0); +lean_dec(x_173); +x_174 = lean_unsigned_to_nat(3u); +lean_inc(x_169); +lean_ctor_set(x_111, 3, x_169); +lean_ctor_set(x_111, 2, x_6); +lean_ctor_set(x_111, 1, x_5); +lean_ctor_set(x_111, 0, x_112); +if (lean_is_scalar(x_9)) { + x_175 = lean_alloc_ctor(0, 5, 0); +} else { + x_175 = x_9; +} +lean_ctor_set(x_175, 0, x_174); +lean_ctor_set(x_175, 1, x_170); +lean_ctor_set(x_175, 2, x_171); +lean_ctor_set(x_175, 3, x_167); +lean_ctor_set(x_175, 4, x_111); +return x_175; +} +else +{ +lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; +x_176 = lean_ctor_get(x_111, 4); +x_177 = lean_ctor_get(x_111, 1); +x_178 = lean_ctor_get(x_111, 2); +lean_inc(x_176); +lean_inc(x_178); +lean_inc(x_177); +lean_dec(x_111); +x_179 = lean_unsigned_to_nat(3u); +lean_inc(x_176); +x_180 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_180, 0, x_112); +lean_ctor_set(x_180, 1, x_5); +lean_ctor_set(x_180, 2, x_6); +lean_ctor_set(x_180, 3, x_176); +lean_ctor_set(x_180, 4, x_176); +if (lean_is_scalar(x_9)) { + x_181 = lean_alloc_ctor(0, 5, 0); +} else { + x_181 = x_9; +} +lean_ctor_set(x_181, 0, x_179); +lean_ctor_set(x_181, 1, x_177); +lean_ctor_set(x_181, 2, x_178); +lean_ctor_set(x_181, 3, x_167); +lean_ctor_set(x_181, 4, x_180); +return x_181; +} +} +else +{ +lean_object* x_182; +x_182 = lean_ctor_get(x_111, 4); +lean_inc(x_182); +if (lean_obj_tag(x_182) == 0) +{ +uint8_t x_183; +x_183 = !lean_is_exclusive(x_111); +if (x_183 == 0) +{ +lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; uint8_t x_189; +x_184 = lean_ctor_get(x_111, 1); +x_185 = lean_ctor_get(x_111, 2); +x_186 = lean_ctor_get(x_111, 4); +lean_dec(x_186); +x_187 = lean_ctor_get(x_111, 3); +lean_dec(x_187); +x_188 = lean_ctor_get(x_111, 0); +lean_dec(x_188); +x_189 = !lean_is_exclusive(x_182); +if (x_189 == 0) +{ +lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; +x_190 = lean_ctor_get(x_182, 1); +x_191 = lean_ctor_get(x_182, 2); +x_192 = lean_ctor_get(x_182, 4); +lean_dec(x_192); +x_193 = lean_ctor_get(x_182, 3); +lean_dec(x_193); +x_194 = lean_ctor_get(x_182, 0); +lean_dec(x_194); +x_195 = lean_unsigned_to_nat(3u); +lean_ctor_set(x_182, 4, x_167); +lean_ctor_set(x_182, 3, x_167); +lean_ctor_set(x_182, 2, x_185); +lean_ctor_set(x_182, 1, x_184); +lean_ctor_set(x_182, 0, x_112); +lean_ctor_set(x_111, 4, x_167); +lean_ctor_set(x_111, 2, x_6); +lean_ctor_set(x_111, 1, x_5); +lean_ctor_set(x_111, 0, x_112); +if (lean_is_scalar(x_9)) { + x_196 = lean_alloc_ctor(0, 5, 0); +} else { + x_196 = x_9; +} +lean_ctor_set(x_196, 0, x_195); +lean_ctor_set(x_196, 1, x_190); +lean_ctor_set(x_196, 2, x_191); +lean_ctor_set(x_196, 3, x_182); +lean_ctor_set(x_196, 4, x_111); +return x_196; +} +else +{ +lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; +x_197 = lean_ctor_get(x_182, 1); +x_198 = lean_ctor_get(x_182, 2); +lean_inc(x_198); +lean_inc(x_197); +lean_dec(x_182); +x_199 = lean_unsigned_to_nat(3u); +x_200 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_200, 0, x_112); +lean_ctor_set(x_200, 1, x_184); +lean_ctor_set(x_200, 2, x_185); +lean_ctor_set(x_200, 3, x_167); +lean_ctor_set(x_200, 4, x_167); +lean_ctor_set(x_111, 4, x_167); +lean_ctor_set(x_111, 2, x_6); +lean_ctor_set(x_111, 1, x_5); +lean_ctor_set(x_111, 0, x_112); +if (lean_is_scalar(x_9)) { + x_201 = lean_alloc_ctor(0, 5, 0); +} else { + x_201 = x_9; +} +lean_ctor_set(x_201, 0, x_199); +lean_ctor_set(x_201, 1, x_197); +lean_ctor_set(x_201, 2, x_198); +lean_ctor_set(x_201, 3, x_200); +lean_ctor_set(x_201, 4, x_111); +return x_201; +} +} +else +{ +lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; +x_202 = lean_ctor_get(x_111, 1); +x_203 = lean_ctor_get(x_111, 2); +lean_inc(x_203); +lean_inc(x_202); +lean_dec(x_111); +x_204 = lean_ctor_get(x_182, 1); +lean_inc(x_204); +x_205 = lean_ctor_get(x_182, 2); +lean_inc(x_205); +if (lean_is_exclusive(x_182)) { + lean_ctor_release(x_182, 0); + lean_ctor_release(x_182, 1); + lean_ctor_release(x_182, 2); + lean_ctor_release(x_182, 3); + lean_ctor_release(x_182, 4); + x_206 = x_182; +} else { + lean_dec_ref(x_182); + x_206 = lean_box(0); +} +x_207 = lean_unsigned_to_nat(3u); +if (lean_is_scalar(x_206)) { + x_208 = lean_alloc_ctor(0, 5, 0); +} else { + x_208 = x_206; +} +lean_ctor_set(x_208, 0, x_112); +lean_ctor_set(x_208, 1, x_202); +lean_ctor_set(x_208, 2, x_203); +lean_ctor_set(x_208, 3, x_167); +lean_ctor_set(x_208, 4, x_167); +x_209 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_209, 0, x_112); +lean_ctor_set(x_209, 1, x_5); +lean_ctor_set(x_209, 2, x_6); +lean_ctor_set(x_209, 3, x_167); +lean_ctor_set(x_209, 4, x_167); +if (lean_is_scalar(x_9)) { + x_210 = lean_alloc_ctor(0, 5, 0); +} else { + x_210 = x_9; +} +lean_ctor_set(x_210, 0, x_207); +lean_ctor_set(x_210, 1, x_204); +lean_ctor_set(x_210, 2, x_205); +lean_ctor_set(x_210, 3, x_208); +lean_ctor_set(x_210, 4, x_209); +return x_210; +} +} +else +{ +lean_object* x_211; lean_object* x_212; +x_211 = lean_unsigned_to_nat(2u); +if (lean_is_scalar(x_9)) { + x_212 = lean_alloc_ctor(0, 5, 0); +} else { + x_212 = x_9; +} +lean_ctor_set(x_212, 0, x_211); +lean_ctor_set(x_212, 1, x_5); +lean_ctor_set(x_212, 2, x_6); +lean_ctor_set(x_212, 3, x_111); +lean_ctor_set(x_212, 4, x_182); +return x_212; +} +} +} +} +} +else +{ +lean_object* x_213; lean_object* x_214; +x_213 = lean_unsigned_to_nat(1u); +x_214 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_214, 0, x_213); +lean_ctor_set(x_214, 1, x_1); +lean_ctor_set(x_214, 2, x_2); +lean_ctor_set(x_214, 3, x_3); +lean_ctor_set(x_214, 4, x_3); +return x_214; +} +} +} +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_insert___at___00Lean_IR_StructRC_Context_insertRename_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = l_Std_DTreeMap_Internal_Impl_insert___at___00Lean_IR_StructRC_Context_insertRename_spec__0___redArg(x_2, x_3, x_4); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_insertRename(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_1, 2); +x_6 = l_Std_DTreeMap_Internal_Impl_insert___at___00Lean_IR_StructRC_Context_insertRename_spec__0___redArg(x_2, x_3, x_5); +lean_ctor_set(x_1, 2, x_6); +return x_1; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_7 = lean_ctor_get(x_1, 0); +x_8 = lean_ctor_get(x_1, 1); +x_9 = lean_ctor_get(x_1, 2); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_dec(x_1); +x_10 = l_Std_DTreeMap_Internal_Impl_insert___at___00Lean_IR_StructRC_Context_insertRename_spec__0___redArg(x_2, x_3, x_9); +x_11 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_11, 0, x_7); +lean_ctor_set(x_11, 1, x_8); +lean_ctor_set(x_11, 2, x_10); +return x_11; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_insert(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_3, 0); +x_6 = l_Std_DTreeMap_Internal_Impl_insert___at___00Lean_IR_StructRC_Context_insertRename_spec__0___redArg(x_1, x_2, x_5); +lean_ctor_set(x_3, 0, x_6); +return x_3; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_7 = lean_ctor_get(x_3, 0); +x_8 = lean_ctor_get(x_3, 1); +x_9 = lean_ctor_get(x_3, 2); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_dec(x_3); +x_10 = l_Std_DTreeMap_Internal_Impl_insert___at___00Lean_IR_StructRC_Context_insertRename_spec__0___redArg(x_1, x_2, x_7); +x_11 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_8); +lean_ctor_set(x_11, 2, x_9); +return x_11; +} +} +} +LEAN_EXPORT uint8_t l_Std_DTreeMap_Internal_Impl_contains___at___00Lean_IR_StructRC_Context_insertIfNew_spec__0___redArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_3 = lean_ctor_get(x_2, 1); +x_4 = lean_ctor_get(x_2, 3); +x_5 = lean_ctor_get(x_2, 4); +x_6 = lean_nat_dec_lt(x_1, x_3); +if (x_6 == 0) +{ +uint8_t x_7; +x_7 = lean_nat_dec_eq(x_1, x_3); +if (x_7 == 0) +{ +x_2 = x_5; +goto _start; +} +else +{ +return x_7; +} +} +else +{ +x_2 = x_4; +goto _start; +} +} +else +{ +uint8_t x_10; +x_10 = 0; +return x_10; +} +} +} +LEAN_EXPORT uint8_t l_Std_DTreeMap_Internal_Impl_contains___at___00Lean_IR_StructRC_Context_insertIfNew_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = l_Std_DTreeMap_Internal_Impl_contains___at___00Lean_IR_StructRC_Context_insertIfNew_spec__0___redArg(x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_insertIfNew(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; uint8_t x_7; +x_4 = lean_ctor_get(x_3, 0); +x_5 = lean_ctor_get(x_3, 1); +x_6 = lean_ctor_get(x_3, 2); +x_7 = l_Std_DTreeMap_Internal_Impl_contains___at___00Lean_IR_StructRC_Context_insertIfNew_spec__0___redArg(x_1, x_4); +if (x_7 == 0) +{ +uint8_t x_8; +lean_inc(x_6); +lean_inc_ref(x_5); +lean_inc(x_4); +x_8 = !lean_is_exclusive(x_3); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_3, 2); +lean_dec(x_9); +x_10 = lean_ctor_get(x_3, 1); +lean_dec(x_10); +x_11 = lean_ctor_get(x_3, 0); +lean_dec(x_11); +x_12 = l_Std_DTreeMap_Internal_Impl_insert___at___00Lean_IR_StructRC_Context_insertRename_spec__0___redArg(x_1, x_2, x_4); +lean_ctor_set(x_3, 0, x_12); +return x_3; +} +else +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_3); +x_13 = l_Std_DTreeMap_Internal_Impl_insert___at___00Lean_IR_StructRC_Context_insertRename_spec__0___redArg(x_1, x_2, x_4); +x_14 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_5); +lean_ctor_set(x_14, 2, x_6); +return x_14; +} +} +else +{ +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +} +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_contains___at___00Lean_IR_StructRC_Context_insertIfNew_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = l_Std_DTreeMap_Internal_Impl_contains___at___00Lean_IR_StructRC_Context_insertIfNew_spec__0(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +x_5 = lean_box(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_contains___at___00Lean_IR_StructRC_Context_insertIfNew_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_DTreeMap_Internal_Impl_contains___at___00Lean_IR_StructRC_Context_insertIfNew_spec__0___redArg(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_DTreeMap_Internal_Impl_erase___at___00Lean_IR_StructRC_Context_remove_spec__0___redArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_3 = lean_ctor_get(x_2, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_2, 2); +lean_inc(x_4); +x_5 = lean_ctor_get(x_2, 3); +lean_inc(x_5); +x_6 = lean_ctor_get(x_2, 4); +lean_inc(x_6); +if (lean_is_exclusive(x_2)) { + lean_ctor_release(x_2, 0); + lean_ctor_release(x_2, 1); + lean_ctor_release(x_2, 2); + lean_ctor_release(x_2, 3); + lean_ctor_release(x_2, 4); + x_7 = x_2; +} else { + lean_dec_ref(x_2); + x_7 = lean_box(0); +} +x_8 = lean_nat_dec_lt(x_1, x_3); +if (x_8 == 0) +{ +uint8_t x_9; +x_9 = lean_nat_dec_eq(x_1, x_3); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; +x_10 = l_Std_DTreeMap_Internal_Impl_erase___at___00Lean_IR_StructRC_Context_remove_spec__0___redArg(x_1, x_6); +x_11 = lean_unsigned_to_nat(1u); +if (lean_obj_tag(x_10) == 0) +{ +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_12 = lean_ctor_get(x_10, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_5, 0); +x_14 = lean_ctor_get(x_5, 1); +x_15 = lean_ctor_get(x_5, 2); +x_16 = lean_ctor_get(x_5, 3); +x_17 = lean_ctor_get(x_5, 4); +lean_inc(x_17); +x_18 = lean_unsigned_to_nat(3u); +x_19 = lean_nat_mul(x_18, x_12); +x_20 = lean_nat_dec_lt(x_19, x_13); +lean_dec(x_19); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +lean_dec(x_17); +x_21 = lean_nat_add(x_11, x_13); +x_22 = lean_nat_add(x_21, x_12); +lean_dec(x_12); +lean_dec(x_21); +if (lean_is_scalar(x_7)) { + x_23 = lean_alloc_ctor(0, 5, 0); +} else { + x_23 = x_7; +} +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_3); +lean_ctor_set(x_23, 2, x_4); +lean_ctor_set(x_23, 3, x_5); +lean_ctor_set(x_23, 4, x_10); +return x_23; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +if (lean_is_exclusive(x_5)) { + lean_ctor_release(x_5, 0); + lean_ctor_release(x_5, 1); + lean_ctor_release(x_5, 2); + lean_ctor_release(x_5, 3); + lean_ctor_release(x_5, 4); + x_24 = x_5; +} else { + lean_dec_ref(x_5); + x_24 = lean_box(0); +} +x_25 = lean_ctor_get(x_16, 0); +x_26 = lean_ctor_get(x_17, 0); +x_27 = lean_ctor_get(x_17, 1); +x_28 = lean_ctor_get(x_17, 2); +x_29 = lean_ctor_get(x_17, 3); +x_30 = lean_ctor_get(x_17, 4); +x_31 = lean_unsigned_to_nat(2u); +x_32 = lean_nat_mul(x_31, x_25); +x_33 = lean_nat_dec_lt(x_26, x_32); +lean_dec(x_32); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_44; lean_object* x_45; +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_27); +if (lean_is_exclusive(x_17)) { + lean_ctor_release(x_17, 0); + lean_ctor_release(x_17, 1); + lean_ctor_release(x_17, 2); + lean_ctor_release(x_17, 3); + lean_ctor_release(x_17, 4); + x_34 = x_17; +} else { + lean_dec_ref(x_17); + x_34 = lean_box(0); +} +x_35 = lean_nat_add(x_11, x_13); +lean_dec(x_13); +x_36 = lean_nat_add(x_35, x_12); +lean_dec(x_35); +x_44 = lean_nat_add(x_11, x_25); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_52; +x_52 = lean_ctor_get(x_29, 0); +lean_inc(x_52); +x_45 = x_52; +goto block_51; +} +else +{ +lean_object* x_53; +x_53 = lean_unsigned_to_nat(0u); +x_45 = x_53; +goto block_51; +} +block_43: +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_nat_add(x_37, x_39); +lean_dec(x_39); +lean_dec(x_37); +if (lean_is_scalar(x_34)) { + x_41 = lean_alloc_ctor(0, 5, 0); +} else { + x_41 = x_34; +} +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_3); +lean_ctor_set(x_41, 2, x_4); +lean_ctor_set(x_41, 3, x_30); +lean_ctor_set(x_41, 4, x_10); +if (lean_is_scalar(x_24)) { + x_42 = lean_alloc_ctor(0, 5, 0); +} else { + x_42 = x_24; +} +lean_ctor_set(x_42, 0, x_36); +lean_ctor_set(x_42, 1, x_27); +lean_ctor_set(x_42, 2, x_28); +lean_ctor_set(x_42, 3, x_38); +lean_ctor_set(x_42, 4, x_41); +return x_42; +} +block_51: +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_nat_add(x_44, x_45); +lean_dec(x_45); +lean_dec(x_44); +if (lean_is_scalar(x_7)) { + x_47 = lean_alloc_ctor(0, 5, 0); +} else { + x_47 = x_7; +} +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_14); +lean_ctor_set(x_47, 2, x_15); +lean_ctor_set(x_47, 3, x_16); +lean_ctor_set(x_47, 4, x_29); +x_48 = lean_nat_add(x_11, x_12); +lean_dec(x_12); +if (lean_obj_tag(x_30) == 0) +{ +lean_object* x_49; +x_49 = lean_ctor_get(x_30, 0); +lean_inc(x_49); +x_37 = x_48; +x_38 = x_47; +x_39 = x_49; +goto block_43; +} +else +{ +lean_object* x_50; +x_50 = lean_unsigned_to_nat(0u); +x_37 = x_48; +x_38 = x_47; +x_39 = x_50; +goto block_43; +} +} +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; +lean_dec(x_7); +x_54 = lean_nat_add(x_11, x_13); +lean_dec(x_13); +x_55 = lean_nat_add(x_54, x_12); +lean_dec(x_54); +x_56 = lean_nat_add(x_11, x_12); +lean_dec(x_12); +x_57 = lean_nat_add(x_56, x_26); +lean_dec(x_56); +lean_inc_ref(x_10); +if (lean_is_scalar(x_24)) { + x_58 = lean_alloc_ctor(0, 5, 0); +} else { + x_58 = x_24; +} +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_3); +lean_ctor_set(x_58, 2, x_4); +lean_ctor_set(x_58, 3, x_17); +lean_ctor_set(x_58, 4, x_10); +x_59 = !lean_is_exclusive(x_10); +if (x_59 == 0) +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_60 = lean_ctor_get(x_10, 4); +lean_dec(x_60); +x_61 = lean_ctor_get(x_10, 3); +lean_dec(x_61); +x_62 = lean_ctor_get(x_10, 2); +lean_dec(x_62); +x_63 = lean_ctor_get(x_10, 1); +lean_dec(x_63); +x_64 = lean_ctor_get(x_10, 0); +lean_dec(x_64); +lean_ctor_set(x_10, 4, x_58); +lean_ctor_set(x_10, 3, x_16); +lean_ctor_set(x_10, 2, x_15); +lean_ctor_set(x_10, 1, x_14); +lean_ctor_set(x_10, 0, x_55); +return x_10; +} +else +{ +lean_object* x_65; +lean_dec(x_10); +x_65 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_65, 0, x_55); +lean_ctor_set(x_65, 1, x_14); +lean_ctor_set(x_65, 2, x_15); +lean_ctor_set(x_65, 3, x_16); +lean_ctor_set(x_65, 4, x_58); +return x_65; +} +} +} +} +else +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_66 = lean_ctor_get(x_10, 0); +lean_inc(x_66); +x_67 = lean_nat_add(x_11, x_66); +lean_dec(x_66); +if (lean_is_scalar(x_7)) { + x_68 = lean_alloc_ctor(0, 5, 0); +} else { + x_68 = x_7; +} +lean_ctor_set(x_68, 0, x_67); +lean_ctor_set(x_68, 1, x_3); +lean_ctor_set(x_68, 2, x_4); +lean_ctor_set(x_68, 3, x_5); +lean_ctor_set(x_68, 4, x_10); +return x_68; +} +} +else +{ +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_69; +x_69 = lean_ctor_get(x_5, 3); +if (lean_obj_tag(x_69) == 0) +{ +lean_object* x_70; +lean_inc_ref(x_69); +x_70 = lean_ctor_get(x_5, 4); +lean_inc(x_70); +if (lean_obj_tag(x_70) == 0) +{ +uint8_t x_71; +x_71 = !lean_is_exclusive(x_5); +if (x_71 == 0) +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_72 = lean_ctor_get(x_5, 0); +x_73 = lean_ctor_get(x_5, 1); +x_74 = lean_ctor_get(x_5, 2); +x_75 = lean_ctor_get(x_5, 4); +lean_dec(x_75); +x_76 = lean_ctor_get(x_5, 3); +lean_dec(x_76); +x_77 = lean_ctor_get(x_70, 0); +x_78 = lean_nat_add(x_11, x_72); +lean_dec(x_72); +x_79 = lean_nat_add(x_11, x_77); +lean_ctor_set(x_5, 4, x_10); +lean_ctor_set(x_5, 3, x_70); +lean_ctor_set(x_5, 2, x_4); +lean_ctor_set(x_5, 1, x_3); +lean_ctor_set(x_5, 0, x_79); +if (lean_is_scalar(x_7)) { + x_80 = lean_alloc_ctor(0, 5, 0); +} else { + x_80 = x_7; +} +lean_ctor_set(x_80, 0, x_78); +lean_ctor_set(x_80, 1, x_73); +lean_ctor_set(x_80, 2, x_74); +lean_ctor_set(x_80, 3, x_69); +lean_ctor_set(x_80, 4, x_5); +return x_80; +} +else +{ +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_81 = lean_ctor_get(x_5, 0); +x_82 = lean_ctor_get(x_5, 1); +x_83 = lean_ctor_get(x_5, 2); +lean_inc(x_83); +lean_inc(x_82); +lean_inc(x_81); +lean_dec(x_5); +x_84 = lean_ctor_get(x_70, 0); +x_85 = lean_nat_add(x_11, x_81); +lean_dec(x_81); +x_86 = lean_nat_add(x_11, x_84); +x_87 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_87, 0, x_86); +lean_ctor_set(x_87, 1, x_3); +lean_ctor_set(x_87, 2, x_4); +lean_ctor_set(x_87, 3, x_70); +lean_ctor_set(x_87, 4, x_10); +if (lean_is_scalar(x_7)) { + x_88 = lean_alloc_ctor(0, 5, 0); +} else { + x_88 = x_7; +} +lean_ctor_set(x_88, 0, x_85); +lean_ctor_set(x_88, 1, x_82); +lean_ctor_set(x_88, 2, x_83); +lean_ctor_set(x_88, 3, x_69); +lean_ctor_set(x_88, 4, x_87); +return x_88; +} +} +else +{ +uint8_t x_89; +x_89 = !lean_is_exclusive(x_5); +if (x_89 == 0) +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_90 = lean_ctor_get(x_5, 1); +x_91 = lean_ctor_get(x_5, 2); +x_92 = lean_ctor_get(x_5, 4); +lean_dec(x_92); +x_93 = lean_ctor_get(x_5, 3); +lean_dec(x_93); +x_94 = lean_ctor_get(x_5, 0); +lean_dec(x_94); +x_95 = lean_unsigned_to_nat(3u); +lean_ctor_set(x_5, 3, x_70); +lean_ctor_set(x_5, 2, x_4); +lean_ctor_set(x_5, 1, x_3); +lean_ctor_set(x_5, 0, x_11); +if (lean_is_scalar(x_7)) { + x_96 = lean_alloc_ctor(0, 5, 0); +} else { + x_96 = x_7; +} +lean_ctor_set(x_96, 0, x_95); +lean_ctor_set(x_96, 1, x_90); +lean_ctor_set(x_96, 2, x_91); +lean_ctor_set(x_96, 3, x_69); +lean_ctor_set(x_96, 4, x_5); +return x_96; +} +else +{ +lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_97 = lean_ctor_get(x_5, 1); +x_98 = lean_ctor_get(x_5, 2); +lean_inc(x_98); +lean_inc(x_97); +lean_dec(x_5); +x_99 = lean_unsigned_to_nat(3u); +x_100 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_100, 0, x_11); +lean_ctor_set(x_100, 1, x_3); +lean_ctor_set(x_100, 2, x_4); +lean_ctor_set(x_100, 3, x_70); +lean_ctor_set(x_100, 4, x_70); +if (lean_is_scalar(x_7)) { + x_101 = lean_alloc_ctor(0, 5, 0); +} else { + x_101 = x_7; +} +lean_ctor_set(x_101, 0, x_99); +lean_ctor_set(x_101, 1, x_97); +lean_ctor_set(x_101, 2, x_98); +lean_ctor_set(x_101, 3, x_69); +lean_ctor_set(x_101, 4, x_100); +return x_101; +} +} +} +else +{ +lean_object* x_102; +x_102 = lean_ctor_get(x_5, 4); +lean_inc(x_102); +if (lean_obj_tag(x_102) == 0) +{ +uint8_t x_103; +lean_inc(x_69); +x_103 = !lean_is_exclusive(x_5); +if (x_103 == 0) +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; uint8_t x_109; +x_104 = lean_ctor_get(x_5, 1); +x_105 = lean_ctor_get(x_5, 2); +x_106 = lean_ctor_get(x_5, 4); +lean_dec(x_106); +x_107 = lean_ctor_get(x_5, 3); +lean_dec(x_107); +x_108 = lean_ctor_get(x_5, 0); +lean_dec(x_108); +x_109 = !lean_is_exclusive(x_102); +if (x_109 == 0) +{ +lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_110 = lean_ctor_get(x_102, 1); +x_111 = lean_ctor_get(x_102, 2); +x_112 = lean_ctor_get(x_102, 4); +lean_dec(x_112); +x_113 = lean_ctor_get(x_102, 3); +lean_dec(x_113); +x_114 = lean_ctor_get(x_102, 0); +lean_dec(x_114); +x_115 = lean_unsigned_to_nat(3u); +lean_ctor_set(x_102, 4, x_69); +lean_ctor_set(x_102, 3, x_69); +lean_ctor_set(x_102, 2, x_105); +lean_ctor_set(x_102, 1, x_104); +lean_ctor_set(x_102, 0, x_11); +lean_ctor_set(x_5, 4, x_69); +lean_ctor_set(x_5, 2, x_4); +lean_ctor_set(x_5, 1, x_3); +lean_ctor_set(x_5, 0, x_11); +if (lean_is_scalar(x_7)) { + x_116 = lean_alloc_ctor(0, 5, 0); +} else { + x_116 = x_7; +} +lean_ctor_set(x_116, 0, x_115); +lean_ctor_set(x_116, 1, x_110); +lean_ctor_set(x_116, 2, x_111); +lean_ctor_set(x_116, 3, x_102); +lean_ctor_set(x_116, 4, x_5); +return x_116; +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; +x_117 = lean_ctor_get(x_102, 1); +x_118 = lean_ctor_get(x_102, 2); +lean_inc(x_118); +lean_inc(x_117); +lean_dec(x_102); +x_119 = lean_unsigned_to_nat(3u); +x_120 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_120, 0, x_11); +lean_ctor_set(x_120, 1, x_104); +lean_ctor_set(x_120, 2, x_105); +lean_ctor_set(x_120, 3, x_69); +lean_ctor_set(x_120, 4, x_69); +lean_ctor_set(x_5, 4, x_69); +lean_ctor_set(x_5, 2, x_4); +lean_ctor_set(x_5, 1, x_3); +lean_ctor_set(x_5, 0, x_11); +if (lean_is_scalar(x_7)) { + x_121 = lean_alloc_ctor(0, 5, 0); +} else { + x_121 = x_7; +} +lean_ctor_set(x_121, 0, x_119); +lean_ctor_set(x_121, 1, x_117); +lean_ctor_set(x_121, 2, x_118); +lean_ctor_set(x_121, 3, x_120); +lean_ctor_set(x_121, 4, x_5); +return x_121; +} +} +else +{ +lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; +x_122 = lean_ctor_get(x_5, 1); +x_123 = lean_ctor_get(x_5, 2); +lean_inc(x_123); +lean_inc(x_122); +lean_dec(x_5); +x_124 = lean_ctor_get(x_102, 1); +lean_inc(x_124); +x_125 = lean_ctor_get(x_102, 2); +lean_inc(x_125); +if (lean_is_exclusive(x_102)) { + lean_ctor_release(x_102, 0); + lean_ctor_release(x_102, 1); + lean_ctor_release(x_102, 2); + lean_ctor_release(x_102, 3); + lean_ctor_release(x_102, 4); + x_126 = x_102; +} else { + lean_dec_ref(x_102); + x_126 = lean_box(0); +} +x_127 = lean_unsigned_to_nat(3u); +if (lean_is_scalar(x_126)) { + x_128 = lean_alloc_ctor(0, 5, 0); +} else { + x_128 = x_126; +} +lean_ctor_set(x_128, 0, x_11); +lean_ctor_set(x_128, 1, x_122); +lean_ctor_set(x_128, 2, x_123); +lean_ctor_set(x_128, 3, x_69); +lean_ctor_set(x_128, 4, x_69); +x_129 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_129, 0, x_11); +lean_ctor_set(x_129, 1, x_3); +lean_ctor_set(x_129, 2, x_4); +lean_ctor_set(x_129, 3, x_69); +lean_ctor_set(x_129, 4, x_69); +if (lean_is_scalar(x_7)) { + x_130 = lean_alloc_ctor(0, 5, 0); +} else { + x_130 = x_7; +} +lean_ctor_set(x_130, 0, x_127); +lean_ctor_set(x_130, 1, x_124); +lean_ctor_set(x_130, 2, x_125); +lean_ctor_set(x_130, 3, x_128); +lean_ctor_set(x_130, 4, x_129); +return x_130; +} +} +else +{ +lean_object* x_131; lean_object* x_132; +x_131 = lean_unsigned_to_nat(2u); +if (lean_is_scalar(x_7)) { + x_132 = lean_alloc_ctor(0, 5, 0); +} else { + x_132 = x_7; +} +lean_ctor_set(x_132, 0, x_131); +lean_ctor_set(x_132, 1, x_3); +lean_ctor_set(x_132, 2, x_4); +lean_ctor_set(x_132, 3, x_5); +lean_ctor_set(x_132, 4, x_102); +return x_132; +} +} +} +else +{ +lean_object* x_133; +if (lean_is_scalar(x_7)) { + x_133 = lean_alloc_ctor(0, 5, 0); +} else { + x_133 = x_7; +} +lean_ctor_set(x_133, 0, x_11); +lean_ctor_set(x_133, 1, x_3); +lean_ctor_set(x_133, 2, x_4); +lean_ctor_set(x_133, 3, x_5); +lean_ctor_set(x_133, 4, x_5); +return x_133; +} +} +} +else +{ +lean_dec(x_7); +lean_dec(x_4); +lean_dec(x_3); +if (lean_obj_tag(x_5) == 0) +{ +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; uint8_t x_145; +x_134 = lean_ctor_get(x_5, 0); +x_135 = lean_ctor_get(x_5, 1); +x_136 = lean_ctor_get(x_5, 2); +x_137 = lean_ctor_get(x_5, 3); +x_138 = lean_ctor_get(x_5, 4); +lean_inc(x_138); +x_139 = lean_ctor_get(x_6, 0); +x_140 = lean_ctor_get(x_6, 1); +x_141 = lean_ctor_get(x_6, 2); +x_142 = lean_ctor_get(x_6, 3); +lean_inc(x_142); +x_143 = lean_ctor_get(x_6, 4); +x_144 = lean_unsigned_to_nat(1u); +x_145 = lean_nat_dec_lt(x_134, x_139); +if (x_145 == 0) +{ +lean_object* x_146; lean_object* x_147; lean_object* x_148; +lean_inc(x_137); +lean_inc(x_136); +lean_inc(x_135); +if (lean_is_exclusive(x_5)) { + lean_ctor_release(x_5, 0); + lean_ctor_release(x_5, 1); + lean_ctor_release(x_5, 2); + lean_ctor_release(x_5, 3); + lean_ctor_release(x_5, 4); + x_146 = x_5; +} else { + lean_dec_ref(x_5); + x_146 = lean_box(0); +} +x_147 = l_Std_DTreeMap_Internal_Impl_maxView___redArg(x_135, x_136, x_137, x_138); +x_148 = lean_ctor_get(x_147, 2); +lean_inc(x_148); +if (lean_obj_tag(x_148) == 0) +{ +lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; uint8_t x_154; +x_149 = lean_ctor_get(x_147, 0); +lean_inc(x_149); +x_150 = lean_ctor_get(x_147, 1); +lean_inc(x_150); +lean_dec_ref(x_147); +x_151 = lean_ctor_get(x_148, 0); +x_152 = lean_unsigned_to_nat(3u); +x_153 = lean_nat_mul(x_152, x_151); +x_154 = lean_nat_dec_lt(x_153, x_139); +lean_dec(x_153); +if (x_154 == 0) +{ +lean_object* x_155; lean_object* x_156; lean_object* x_157; +lean_dec(x_142); +x_155 = lean_nat_add(x_144, x_151); +x_156 = lean_nat_add(x_155, x_139); +lean_dec(x_155); +if (lean_is_scalar(x_146)) { + x_157 = lean_alloc_ctor(0, 5, 0); +} else { + x_157 = x_146; +} +lean_ctor_set(x_157, 0, x_156); +lean_ctor_set(x_157, 1, x_149); +lean_ctor_set(x_157, 2, x_150); +lean_ctor_set(x_157, 3, x_148); +lean_ctor_set(x_157, 4, x_6); +return x_157; +} +else +{ +lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; uint8_t x_167; +lean_inc(x_143); +lean_inc(x_141); +lean_inc(x_140); +lean_inc(x_139); +if (lean_is_exclusive(x_6)) { + lean_ctor_release(x_6, 0); + lean_ctor_release(x_6, 1); + lean_ctor_release(x_6, 2); + lean_ctor_release(x_6, 3); + lean_ctor_release(x_6, 4); + x_158 = x_6; +} else { + lean_dec_ref(x_6); + x_158 = lean_box(0); +} +x_159 = lean_ctor_get(x_142, 0); +x_160 = lean_ctor_get(x_142, 1); +x_161 = lean_ctor_get(x_142, 2); +x_162 = lean_ctor_get(x_142, 3); +x_163 = lean_ctor_get(x_142, 4); +x_164 = lean_ctor_get(x_143, 0); +x_165 = lean_unsigned_to_nat(2u); +x_166 = lean_nat_mul(x_165, x_164); +x_167 = lean_nat_dec_lt(x_159, x_166); +lean_dec(x_166); +if (x_167 == 0) +{ +lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_178; +lean_inc(x_163); +lean_inc(x_162); +lean_inc(x_161); +lean_inc(x_160); +if (lean_is_exclusive(x_142)) { + lean_ctor_release(x_142, 0); + lean_ctor_release(x_142, 1); + lean_ctor_release(x_142, 2); + lean_ctor_release(x_142, 3); + lean_ctor_release(x_142, 4); + x_168 = x_142; +} else { + lean_dec_ref(x_142); + x_168 = lean_box(0); +} +x_169 = lean_nat_add(x_144, x_151); +x_170 = lean_nat_add(x_169, x_139); +lean_dec(x_139); +if (lean_obj_tag(x_162) == 0) +{ +lean_object* x_185; +x_185 = lean_ctor_get(x_162, 0); +lean_inc(x_185); +x_178 = x_185; +goto block_184; +} +else +{ +lean_object* x_186; +x_186 = lean_unsigned_to_nat(0u); +x_178 = x_186; +goto block_184; +} +block_177: +{ +lean_object* x_174; lean_object* x_175; lean_object* x_176; +x_174 = lean_nat_add(x_171, x_173); +lean_dec(x_173); +lean_dec(x_171); +if (lean_is_scalar(x_168)) { + x_175 = lean_alloc_ctor(0, 5, 0); +} else { + x_175 = x_168; +} +lean_ctor_set(x_175, 0, x_174); +lean_ctor_set(x_175, 1, x_140); +lean_ctor_set(x_175, 2, x_141); +lean_ctor_set(x_175, 3, x_163); +lean_ctor_set(x_175, 4, x_143); +if (lean_is_scalar(x_158)) { + x_176 = lean_alloc_ctor(0, 5, 0); +} else { + x_176 = x_158; +} +lean_ctor_set(x_176, 0, x_170); +lean_ctor_set(x_176, 1, x_160); +lean_ctor_set(x_176, 2, x_161); +lean_ctor_set(x_176, 3, x_172); +lean_ctor_set(x_176, 4, x_175); +return x_176; +} +block_184: +{ +lean_object* x_179; lean_object* x_180; lean_object* x_181; +x_179 = lean_nat_add(x_169, x_178); +lean_dec(x_178); +lean_dec(x_169); +if (lean_is_scalar(x_146)) { + x_180 = lean_alloc_ctor(0, 5, 0); +} else { + x_180 = x_146; +} +lean_ctor_set(x_180, 0, x_179); +lean_ctor_set(x_180, 1, x_149); +lean_ctor_set(x_180, 2, x_150); +lean_ctor_set(x_180, 3, x_148); +lean_ctor_set(x_180, 4, x_162); +x_181 = lean_nat_add(x_144, x_164); +if (lean_obj_tag(x_163) == 0) +{ +lean_object* x_182; +x_182 = lean_ctor_get(x_163, 0); +lean_inc(x_182); +x_171 = x_181; +x_172 = x_180; +x_173 = x_182; +goto block_177; +} +else +{ +lean_object* x_183; +x_183 = lean_unsigned_to_nat(0u); +x_171 = x_181; +x_172 = x_180; +x_173 = x_183; +goto block_177; +} +} +} +else +{ +lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; +x_187 = lean_nat_add(x_144, x_151); +x_188 = lean_nat_add(x_187, x_139); +lean_dec(x_139); +x_189 = lean_nat_add(x_187, x_159); +lean_dec(x_187); +if (lean_is_scalar(x_158)) { + x_190 = lean_alloc_ctor(0, 5, 0); +} else { + x_190 = x_158; +} +lean_ctor_set(x_190, 0, x_189); +lean_ctor_set(x_190, 1, x_149); +lean_ctor_set(x_190, 2, x_150); +lean_ctor_set(x_190, 3, x_148); +lean_ctor_set(x_190, 4, x_142); +if (lean_is_scalar(x_146)) { + x_191 = lean_alloc_ctor(0, 5, 0); +} else { + x_191 = x_146; +} +lean_ctor_set(x_191, 0, x_188); +lean_ctor_set(x_191, 1, x_140); +lean_ctor_set(x_191, 2, x_141); +lean_ctor_set(x_191, 3, x_190); +lean_ctor_set(x_191, 4, x_143); +return x_191; +} +} +} +else +{ +uint8_t x_192; +lean_inc(x_143); +lean_inc(x_141); +lean_inc(x_140); +lean_inc(x_139); +x_192 = !lean_is_exclusive(x_6); +if (x_192 == 0) +{ +lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; +x_193 = lean_ctor_get(x_6, 4); +lean_dec(x_193); +x_194 = lean_ctor_get(x_6, 3); +lean_dec(x_194); +x_195 = lean_ctor_get(x_6, 2); +lean_dec(x_195); +x_196 = lean_ctor_get(x_6, 1); +lean_dec(x_196); +x_197 = lean_ctor_get(x_6, 0); +lean_dec(x_197); +if (lean_obj_tag(x_142) == 0) +{ +if (lean_obj_tag(x_143) == 0) +{ +lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; +x_198 = lean_ctor_get(x_147, 0); +lean_inc(x_198); +x_199 = lean_ctor_get(x_147, 1); +lean_inc(x_199); +lean_dec_ref(x_147); +x_200 = lean_ctor_get(x_142, 0); +x_201 = lean_nat_add(x_144, x_139); +lean_dec(x_139); +x_202 = lean_nat_add(x_144, x_200); +lean_ctor_set(x_6, 4, x_142); +lean_ctor_set(x_6, 3, x_148); +lean_ctor_set(x_6, 2, x_199); +lean_ctor_set(x_6, 1, x_198); +lean_ctor_set(x_6, 0, x_202); +if (lean_is_scalar(x_146)) { + x_203 = lean_alloc_ctor(0, 5, 0); +} else { + x_203 = x_146; +} +lean_ctor_set(x_203, 0, x_201); +lean_ctor_set(x_203, 1, x_140); +lean_ctor_set(x_203, 2, x_141); +lean_ctor_set(x_203, 3, x_6); +lean_ctor_set(x_203, 4, x_143); +return x_203; +} +else +{ +lean_object* x_204; lean_object* x_205; uint8_t x_206; +lean_dec(x_139); +x_204 = lean_ctor_get(x_147, 0); +lean_inc(x_204); +x_205 = lean_ctor_get(x_147, 1); +lean_inc(x_205); +lean_dec_ref(x_147); +x_206 = !lean_is_exclusive(x_142); +if (x_206 == 0) +{ +lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; +x_207 = lean_ctor_get(x_142, 1); +x_208 = lean_ctor_get(x_142, 2); +x_209 = lean_ctor_get(x_142, 4); +lean_dec(x_209); +x_210 = lean_ctor_get(x_142, 3); +lean_dec(x_210); +x_211 = lean_ctor_get(x_142, 0); +lean_dec(x_211); +x_212 = lean_unsigned_to_nat(3u); +lean_ctor_set(x_142, 4, x_143); +lean_ctor_set(x_142, 3, x_143); +lean_ctor_set(x_142, 2, x_205); +lean_ctor_set(x_142, 1, x_204); +lean_ctor_set(x_142, 0, x_144); +lean_ctor_set(x_6, 3, x_143); +lean_ctor_set(x_6, 0, x_144); +if (lean_is_scalar(x_146)) { + x_213 = lean_alloc_ctor(0, 5, 0); +} else { + x_213 = x_146; +} +lean_ctor_set(x_213, 0, x_212); +lean_ctor_set(x_213, 1, x_207); +lean_ctor_set(x_213, 2, x_208); +lean_ctor_set(x_213, 3, x_142); +lean_ctor_set(x_213, 4, x_6); +return x_213; +} +else +{ +lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; +x_214 = lean_ctor_get(x_142, 1); +x_215 = lean_ctor_get(x_142, 2); +lean_inc(x_215); +lean_inc(x_214); +lean_dec(x_142); +x_216 = lean_unsigned_to_nat(3u); +x_217 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_217, 0, x_144); +lean_ctor_set(x_217, 1, x_204); +lean_ctor_set(x_217, 2, x_205); +lean_ctor_set(x_217, 3, x_143); +lean_ctor_set(x_217, 4, x_143); +lean_ctor_set(x_6, 3, x_143); +lean_ctor_set(x_6, 0, x_144); +if (lean_is_scalar(x_146)) { + x_218 = lean_alloc_ctor(0, 5, 0); +} else { + x_218 = x_146; +} +lean_ctor_set(x_218, 0, x_216); +lean_ctor_set(x_218, 1, x_214); +lean_ctor_set(x_218, 2, x_215); +lean_ctor_set(x_218, 3, x_217); +lean_ctor_set(x_218, 4, x_6); +return x_218; +} +} +} +else +{ +if (lean_obj_tag(x_143) == 0) +{ +lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; +lean_dec(x_139); +x_219 = lean_ctor_get(x_147, 0); +lean_inc(x_219); +x_220 = lean_ctor_get(x_147, 1); +lean_inc(x_220); +lean_dec_ref(x_147); +x_221 = lean_unsigned_to_nat(3u); +lean_ctor_set(x_6, 4, x_142); +lean_ctor_set(x_6, 2, x_220); +lean_ctor_set(x_6, 1, x_219); +lean_ctor_set(x_6, 0, x_144); +if (lean_is_scalar(x_146)) { + x_222 = lean_alloc_ctor(0, 5, 0); +} else { + x_222 = x_146; +} +lean_ctor_set(x_222, 0, x_221); +lean_ctor_set(x_222, 1, x_140); +lean_ctor_set(x_222, 2, x_141); +lean_ctor_set(x_222, 3, x_6); +lean_ctor_set(x_222, 4, x_143); +return x_222; +} +else +{ +lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; +x_223 = lean_ctor_get(x_147, 0); +lean_inc(x_223); +x_224 = lean_ctor_get(x_147, 1); +lean_inc(x_224); +lean_dec_ref(x_147); +lean_ctor_set(x_6, 3, x_143); +x_225 = lean_unsigned_to_nat(2u); +if (lean_is_scalar(x_146)) { + x_226 = lean_alloc_ctor(0, 5, 0); +} else { + x_226 = x_146; +} +lean_ctor_set(x_226, 0, x_225); +lean_ctor_set(x_226, 1, x_223); +lean_ctor_set(x_226, 2, x_224); +lean_ctor_set(x_226, 3, x_143); +lean_ctor_set(x_226, 4, x_6); +return x_226; +} +} +} +else +{ +lean_dec(x_6); +if (lean_obj_tag(x_142) == 0) +{ +if (lean_obj_tag(x_143) == 0) +{ +lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; +x_227 = lean_ctor_get(x_147, 0); +lean_inc(x_227); +x_228 = lean_ctor_get(x_147, 1); +lean_inc(x_228); +lean_dec_ref(x_147); +x_229 = lean_ctor_get(x_142, 0); +x_230 = lean_nat_add(x_144, x_139); +lean_dec(x_139); +x_231 = lean_nat_add(x_144, x_229); +x_232 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_232, 0, x_231); +lean_ctor_set(x_232, 1, x_227); +lean_ctor_set(x_232, 2, x_228); +lean_ctor_set(x_232, 3, x_148); +lean_ctor_set(x_232, 4, x_142); +if (lean_is_scalar(x_146)) { + x_233 = lean_alloc_ctor(0, 5, 0); +} else { + x_233 = x_146; +} +lean_ctor_set(x_233, 0, x_230); +lean_ctor_set(x_233, 1, x_140); +lean_ctor_set(x_233, 2, x_141); +lean_ctor_set(x_233, 3, x_232); +lean_ctor_set(x_233, 4, x_143); +return x_233; +} +else +{ +lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; +lean_dec(x_139); +x_234 = lean_ctor_get(x_147, 0); +lean_inc(x_234); +x_235 = lean_ctor_get(x_147, 1); +lean_inc(x_235); +lean_dec_ref(x_147); +x_236 = lean_ctor_get(x_142, 1); +lean_inc(x_236); +x_237 = lean_ctor_get(x_142, 2); +lean_inc(x_237); +if (lean_is_exclusive(x_142)) { + lean_ctor_release(x_142, 0); + lean_ctor_release(x_142, 1); + lean_ctor_release(x_142, 2); + lean_ctor_release(x_142, 3); + lean_ctor_release(x_142, 4); + x_238 = x_142; +} else { + lean_dec_ref(x_142); + x_238 = lean_box(0); +} +x_239 = lean_unsigned_to_nat(3u); +if (lean_is_scalar(x_238)) { + x_240 = lean_alloc_ctor(0, 5, 0); +} else { + x_240 = x_238; +} +lean_ctor_set(x_240, 0, x_144); +lean_ctor_set(x_240, 1, x_234); +lean_ctor_set(x_240, 2, x_235); +lean_ctor_set(x_240, 3, x_143); +lean_ctor_set(x_240, 4, x_143); +x_241 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_241, 0, x_144); +lean_ctor_set(x_241, 1, x_140); +lean_ctor_set(x_241, 2, x_141); +lean_ctor_set(x_241, 3, x_143); +lean_ctor_set(x_241, 4, x_143); +if (lean_is_scalar(x_146)) { + x_242 = lean_alloc_ctor(0, 5, 0); +} else { + x_242 = x_146; +} +lean_ctor_set(x_242, 0, x_239); +lean_ctor_set(x_242, 1, x_236); +lean_ctor_set(x_242, 2, x_237); +lean_ctor_set(x_242, 3, x_240); +lean_ctor_set(x_242, 4, x_241); +return x_242; +} +} +else +{ +if (lean_obj_tag(x_143) == 0) +{ +lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; +lean_dec(x_139); +x_243 = lean_ctor_get(x_147, 0); +lean_inc(x_243); +x_244 = lean_ctor_get(x_147, 1); +lean_inc(x_244); +lean_dec_ref(x_147); +x_245 = lean_unsigned_to_nat(3u); +x_246 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_246, 0, x_144); +lean_ctor_set(x_246, 1, x_243); +lean_ctor_set(x_246, 2, x_244); +lean_ctor_set(x_246, 3, x_142); +lean_ctor_set(x_246, 4, x_142); +if (lean_is_scalar(x_146)) { + x_247 = lean_alloc_ctor(0, 5, 0); +} else { + x_247 = x_146; +} +lean_ctor_set(x_247, 0, x_245); +lean_ctor_set(x_247, 1, x_140); +lean_ctor_set(x_247, 2, x_141); +lean_ctor_set(x_247, 3, x_246); +lean_ctor_set(x_247, 4, x_143); +return x_247; +} +else +{ +lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; +x_248 = lean_ctor_get(x_147, 0); +lean_inc(x_248); +x_249 = lean_ctor_get(x_147, 1); +lean_inc(x_249); +lean_dec_ref(x_147); +x_250 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_250, 0, x_139); +lean_ctor_set(x_250, 1, x_140); +lean_ctor_set(x_250, 2, x_141); +lean_ctor_set(x_250, 3, x_143); +lean_ctor_set(x_250, 4, x_143); +x_251 = lean_unsigned_to_nat(2u); +if (lean_is_scalar(x_146)) { + x_252 = lean_alloc_ctor(0, 5, 0); +} else { + x_252 = x_146; +} +lean_ctor_set(x_252, 0, x_251); +lean_ctor_set(x_252, 1, x_248); +lean_ctor_set(x_252, 2, x_249); +lean_ctor_set(x_252, 3, x_143); +lean_ctor_set(x_252, 4, x_250); +return x_252; +} +} +} +} +} +else +{ +lean_object* x_253; lean_object* x_254; lean_object* x_255; +lean_inc(x_143); +lean_inc(x_141); +lean_inc(x_140); +if (lean_is_exclusive(x_6)) { + lean_ctor_release(x_6, 0); + lean_ctor_release(x_6, 1); + lean_ctor_release(x_6, 2); + lean_ctor_release(x_6, 3); + lean_ctor_release(x_6, 4); + x_253 = x_6; +} else { + lean_dec_ref(x_6); + x_253 = lean_box(0); +} +x_254 = l_Std_DTreeMap_Internal_Impl_minView___redArg(x_140, x_141, x_142, x_143); +x_255 = lean_ctor_get(x_254, 2); +lean_inc(x_255); +if (lean_obj_tag(x_255) == 0) +{ +lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; uint8_t x_261; +x_256 = lean_ctor_get(x_254, 0); +lean_inc(x_256); +x_257 = lean_ctor_get(x_254, 1); +lean_inc(x_257); +lean_dec_ref(x_254); +x_258 = lean_ctor_get(x_255, 0); +x_259 = lean_unsigned_to_nat(3u); +x_260 = lean_nat_mul(x_259, x_258); +x_261 = lean_nat_dec_lt(x_260, x_134); +lean_dec(x_260); +if (x_261 == 0) +{ +lean_object* x_262; lean_object* x_263; lean_object* x_264; +lean_dec(x_138); +x_262 = lean_nat_add(x_144, x_134); +x_263 = lean_nat_add(x_262, x_258); +lean_dec(x_262); +if (lean_is_scalar(x_253)) { + x_264 = lean_alloc_ctor(0, 5, 0); +} else { + x_264 = x_253; +} +lean_ctor_set(x_264, 0, x_263); +lean_ctor_set(x_264, 1, x_256); +lean_ctor_set(x_264, 2, x_257); +lean_ctor_set(x_264, 3, x_5); +lean_ctor_set(x_264, 4, x_255); +return x_264; +} +else +{ +uint8_t x_265; +lean_inc(x_137); +lean_inc(x_136); +lean_inc(x_135); +lean_inc(x_134); +x_265 = !lean_is_exclusive(x_5); +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; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; uint8_t x_279; +x_266 = lean_ctor_get(x_5, 4); +lean_dec(x_266); +x_267 = lean_ctor_get(x_5, 3); +lean_dec(x_267); +x_268 = lean_ctor_get(x_5, 2); +lean_dec(x_268); +x_269 = lean_ctor_get(x_5, 1); +lean_dec(x_269); +x_270 = lean_ctor_get(x_5, 0); +lean_dec(x_270); +x_271 = lean_ctor_get(x_137, 0); +x_272 = lean_ctor_get(x_138, 0); +x_273 = lean_ctor_get(x_138, 1); +x_274 = lean_ctor_get(x_138, 2); +x_275 = lean_ctor_get(x_138, 3); +x_276 = lean_ctor_get(x_138, 4); +x_277 = lean_unsigned_to_nat(2u); +x_278 = lean_nat_mul(x_277, x_271); +x_279 = lean_nat_dec_lt(x_272, x_278); +lean_dec(x_278); +if (x_279 == 0) +{ +lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_296; lean_object* x_297; +lean_inc(x_276); +lean_inc(x_275); +lean_inc(x_274); +lean_inc(x_273); +lean_free_object(x_5); +if (lean_is_exclusive(x_138)) { + lean_ctor_release(x_138, 0); + lean_ctor_release(x_138, 1); + lean_ctor_release(x_138, 2); + lean_ctor_release(x_138, 3); + lean_ctor_release(x_138, 4); + x_280 = x_138; +} else { + lean_dec_ref(x_138); + x_280 = lean_box(0); +} +x_281 = lean_nat_add(x_144, x_134); +lean_dec(x_134); +x_282 = lean_nat_add(x_281, x_258); +lean_dec(x_281); +x_296 = lean_nat_add(x_144, x_271); +if (lean_obj_tag(x_275) == 0) +{ +lean_object* x_304; +x_304 = lean_ctor_get(x_275, 0); +lean_inc(x_304); +x_297 = x_304; +goto block_303; +} +else +{ +lean_object* x_305; +x_305 = lean_unsigned_to_nat(0u); +x_297 = x_305; +goto block_303; +} +block_295: +{ +lean_object* x_286; lean_object* x_287; uint8_t x_288; +x_286 = lean_nat_add(x_284, x_285); +lean_dec(x_285); +lean_dec(x_284); +lean_inc_ref(x_255); +if (lean_is_scalar(x_280)) { + x_287 = lean_alloc_ctor(0, 5, 0); +} else { + x_287 = x_280; +} +lean_ctor_set(x_287, 0, x_286); +lean_ctor_set(x_287, 1, x_256); +lean_ctor_set(x_287, 2, x_257); +lean_ctor_set(x_287, 3, x_276); +lean_ctor_set(x_287, 4, x_255); +x_288 = !lean_is_exclusive(x_255); +if (x_288 == 0) +{ +lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; +x_289 = lean_ctor_get(x_255, 4); +lean_dec(x_289); +x_290 = lean_ctor_get(x_255, 3); +lean_dec(x_290); +x_291 = lean_ctor_get(x_255, 2); +lean_dec(x_291); +x_292 = lean_ctor_get(x_255, 1); +lean_dec(x_292); +x_293 = lean_ctor_get(x_255, 0); +lean_dec(x_293); +lean_ctor_set(x_255, 4, x_287); +lean_ctor_set(x_255, 3, x_283); +lean_ctor_set(x_255, 2, x_274); +lean_ctor_set(x_255, 1, x_273); +lean_ctor_set(x_255, 0, x_282); +return x_255; +} +else +{ +lean_object* x_294; +lean_dec(x_255); +x_294 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_294, 0, x_282); +lean_ctor_set(x_294, 1, x_273); +lean_ctor_set(x_294, 2, x_274); +lean_ctor_set(x_294, 3, x_283); +lean_ctor_set(x_294, 4, x_287); +return x_294; +} +} +block_303: +{ +lean_object* x_298; lean_object* x_299; lean_object* x_300; +x_298 = lean_nat_add(x_296, x_297); +lean_dec(x_297); +lean_dec(x_296); +if (lean_is_scalar(x_253)) { + x_299 = lean_alloc_ctor(0, 5, 0); +} else { + x_299 = x_253; +} +lean_ctor_set(x_299, 0, x_298); +lean_ctor_set(x_299, 1, x_135); +lean_ctor_set(x_299, 2, x_136); +lean_ctor_set(x_299, 3, x_137); +lean_ctor_set(x_299, 4, x_275); +x_300 = lean_nat_add(x_144, x_258); +if (lean_obj_tag(x_276) == 0) +{ +lean_object* x_301; +x_301 = lean_ctor_get(x_276, 0); +lean_inc(x_301); +x_283 = x_299; +x_284 = x_300; +x_285 = x_301; +goto block_295; +} +else +{ +lean_object* x_302; +x_302 = lean_unsigned_to_nat(0u); +x_283 = x_299; +x_284 = x_300; +x_285 = x_302; +goto block_295; +} +} +} +else +{ +lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; +x_306 = lean_nat_add(x_144, x_134); +lean_dec(x_134); +x_307 = lean_nat_add(x_306, x_258); +lean_dec(x_306); +x_308 = lean_nat_add(x_144, x_258); +x_309 = lean_nat_add(x_308, x_272); +lean_dec(x_308); +if (lean_is_scalar(x_253)) { + x_310 = lean_alloc_ctor(0, 5, 0); +} else { + x_310 = x_253; +} +lean_ctor_set(x_310, 0, x_309); +lean_ctor_set(x_310, 1, x_256); +lean_ctor_set(x_310, 2, x_257); +lean_ctor_set(x_310, 3, x_138); +lean_ctor_set(x_310, 4, x_255); +lean_ctor_set(x_5, 4, x_310); +lean_ctor_set(x_5, 0, x_307); +return x_5; +} +} +else +{ +lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; uint8_t x_319; +lean_dec(x_5); +x_311 = lean_ctor_get(x_137, 0); +x_312 = lean_ctor_get(x_138, 0); +x_313 = lean_ctor_get(x_138, 1); +x_314 = lean_ctor_get(x_138, 2); +x_315 = lean_ctor_get(x_138, 3); +x_316 = lean_ctor_get(x_138, 4); +x_317 = lean_unsigned_to_nat(2u); +x_318 = lean_nat_mul(x_317, x_311); +x_319 = lean_nat_dec_lt(x_312, x_318); +lean_dec(x_318); +if (x_319 == 0) +{ +lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_331; lean_object* x_332; +lean_inc(x_316); +lean_inc(x_315); +lean_inc(x_314); +lean_inc(x_313); +if (lean_is_exclusive(x_138)) { + lean_ctor_release(x_138, 0); + lean_ctor_release(x_138, 1); + lean_ctor_release(x_138, 2); + lean_ctor_release(x_138, 3); + lean_ctor_release(x_138, 4); + x_320 = x_138; +} else { + lean_dec_ref(x_138); + x_320 = lean_box(0); +} +x_321 = lean_nat_add(x_144, x_134); +lean_dec(x_134); +x_322 = lean_nat_add(x_321, x_258); +lean_dec(x_321); +x_331 = lean_nat_add(x_144, x_311); +if (lean_obj_tag(x_315) == 0) +{ +lean_object* x_339; +x_339 = lean_ctor_get(x_315, 0); +lean_inc(x_339); +x_332 = x_339; +goto block_338; +} +else +{ +lean_object* x_340; +x_340 = lean_unsigned_to_nat(0u); +x_332 = x_340; +goto block_338; +} +block_330: +{ +lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; +x_326 = lean_nat_add(x_324, x_325); +lean_dec(x_325); +lean_dec(x_324); +lean_inc_ref(x_255); +if (lean_is_scalar(x_320)) { + x_327 = lean_alloc_ctor(0, 5, 0); +} else { + x_327 = x_320; +} +lean_ctor_set(x_327, 0, x_326); +lean_ctor_set(x_327, 1, x_256); +lean_ctor_set(x_327, 2, x_257); +lean_ctor_set(x_327, 3, x_316); +lean_ctor_set(x_327, 4, x_255); +if (lean_is_exclusive(x_255)) { + lean_ctor_release(x_255, 0); + lean_ctor_release(x_255, 1); + lean_ctor_release(x_255, 2); + lean_ctor_release(x_255, 3); + lean_ctor_release(x_255, 4); + x_328 = x_255; +} else { + lean_dec_ref(x_255); + x_328 = lean_box(0); +} +if (lean_is_scalar(x_328)) { + x_329 = lean_alloc_ctor(0, 5, 0); +} else { + x_329 = x_328; +} +lean_ctor_set(x_329, 0, x_322); +lean_ctor_set(x_329, 1, x_313); +lean_ctor_set(x_329, 2, x_314); +lean_ctor_set(x_329, 3, x_323); +lean_ctor_set(x_329, 4, x_327); +return x_329; +} +block_338: +{ +lean_object* x_333; lean_object* x_334; lean_object* x_335; +x_333 = lean_nat_add(x_331, x_332); +lean_dec(x_332); +lean_dec(x_331); +if (lean_is_scalar(x_253)) { + x_334 = lean_alloc_ctor(0, 5, 0); +} else { + x_334 = x_253; +} +lean_ctor_set(x_334, 0, x_333); +lean_ctor_set(x_334, 1, x_135); +lean_ctor_set(x_334, 2, x_136); +lean_ctor_set(x_334, 3, x_137); +lean_ctor_set(x_334, 4, x_315); +x_335 = lean_nat_add(x_144, x_258); +if (lean_obj_tag(x_316) == 0) +{ +lean_object* x_336; +x_336 = lean_ctor_get(x_316, 0); +lean_inc(x_336); +x_323 = x_334; +x_324 = x_335; +x_325 = x_336; +goto block_330; +} +else +{ +lean_object* x_337; +x_337 = lean_unsigned_to_nat(0u); +x_323 = x_334; +x_324 = x_335; +x_325 = x_337; +goto block_330; +} +} +} +else +{ +lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; +x_341 = lean_nat_add(x_144, x_134); +lean_dec(x_134); +x_342 = lean_nat_add(x_341, x_258); +lean_dec(x_341); +x_343 = lean_nat_add(x_144, x_258); +x_344 = lean_nat_add(x_343, x_312); +lean_dec(x_343); +if (lean_is_scalar(x_253)) { + x_345 = lean_alloc_ctor(0, 5, 0); +} else { + x_345 = x_253; +} +lean_ctor_set(x_345, 0, x_344); +lean_ctor_set(x_345, 1, x_256); +lean_ctor_set(x_345, 2, x_257); +lean_ctor_set(x_345, 3, x_138); +lean_ctor_set(x_345, 4, x_255); +x_346 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_346, 0, x_342); +lean_ctor_set(x_346, 1, x_135); +lean_ctor_set(x_346, 2, x_136); +lean_ctor_set(x_346, 3, x_137); +lean_ctor_set(x_346, 4, x_345); +return x_346; +} +} +} +} +else +{ +if (lean_obj_tag(x_137) == 0) +{ +uint8_t x_347; +lean_inc_ref(x_137); +lean_inc(x_136); +lean_inc(x_135); +lean_inc(x_134); +x_347 = !lean_is_exclusive(x_5); +if (x_347 == 0) +{ +lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; +x_348 = lean_ctor_get(x_5, 4); +lean_dec(x_348); +x_349 = lean_ctor_get(x_5, 3); +lean_dec(x_349); +x_350 = lean_ctor_get(x_5, 2); +lean_dec(x_350); +x_351 = lean_ctor_get(x_5, 1); +lean_dec(x_351); +x_352 = lean_ctor_get(x_5, 0); +lean_dec(x_352); +if (lean_obj_tag(x_138) == 0) +{ +lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; +x_353 = lean_ctor_get(x_254, 0); +lean_inc(x_353); +x_354 = lean_ctor_get(x_254, 1); +lean_inc(x_354); +lean_dec_ref(x_254); +x_355 = lean_ctor_get(x_138, 0); +x_356 = lean_nat_add(x_144, x_134); +lean_dec(x_134); +x_357 = lean_nat_add(x_144, x_355); +if (lean_is_scalar(x_253)) { + x_358 = lean_alloc_ctor(0, 5, 0); +} else { + x_358 = x_253; +} +lean_ctor_set(x_358, 0, x_357); +lean_ctor_set(x_358, 1, x_353); +lean_ctor_set(x_358, 2, x_354); +lean_ctor_set(x_358, 3, x_138); +lean_ctor_set(x_358, 4, x_255); +lean_ctor_set(x_5, 4, x_358); +lean_ctor_set(x_5, 0, x_356); +return x_5; +} +else +{ +lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; +lean_dec(x_134); +x_359 = lean_ctor_get(x_254, 0); +lean_inc(x_359); +x_360 = lean_ctor_get(x_254, 1); +lean_inc(x_360); +lean_dec_ref(x_254); +x_361 = lean_unsigned_to_nat(3u); +if (lean_is_scalar(x_253)) { + x_362 = lean_alloc_ctor(0, 5, 0); +} else { + x_362 = x_253; +} +lean_ctor_set(x_362, 0, x_144); +lean_ctor_set(x_362, 1, x_359); +lean_ctor_set(x_362, 2, x_360); +lean_ctor_set(x_362, 3, x_138); +lean_ctor_set(x_362, 4, x_138); +lean_ctor_set(x_5, 4, x_362); +lean_ctor_set(x_5, 0, x_361); +return x_5; +} +} +else +{ +lean_dec(x_5); +if (lean_obj_tag(x_138) == 0) +{ +lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; +x_363 = lean_ctor_get(x_254, 0); +lean_inc(x_363); +x_364 = lean_ctor_get(x_254, 1); +lean_inc(x_364); +lean_dec_ref(x_254); +x_365 = lean_ctor_get(x_138, 0); +x_366 = lean_nat_add(x_144, x_134); +lean_dec(x_134); +x_367 = lean_nat_add(x_144, x_365); +if (lean_is_scalar(x_253)) { + x_368 = lean_alloc_ctor(0, 5, 0); +} else { + x_368 = x_253; +} +lean_ctor_set(x_368, 0, x_367); +lean_ctor_set(x_368, 1, x_363); +lean_ctor_set(x_368, 2, x_364); +lean_ctor_set(x_368, 3, x_138); +lean_ctor_set(x_368, 4, x_255); +x_369 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_369, 0, x_366); +lean_ctor_set(x_369, 1, x_135); +lean_ctor_set(x_369, 2, x_136); +lean_ctor_set(x_369, 3, x_137); +lean_ctor_set(x_369, 4, x_368); +return x_369; +} +else +{ +lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; +lean_dec(x_134); +x_370 = lean_ctor_get(x_254, 0); +lean_inc(x_370); +x_371 = lean_ctor_get(x_254, 1); +lean_inc(x_371); +lean_dec_ref(x_254); +x_372 = lean_unsigned_to_nat(3u); +if (lean_is_scalar(x_253)) { + x_373 = lean_alloc_ctor(0, 5, 0); +} else { + x_373 = x_253; +} +lean_ctor_set(x_373, 0, x_144); +lean_ctor_set(x_373, 1, x_370); +lean_ctor_set(x_373, 2, x_371); +lean_ctor_set(x_373, 3, x_138); +lean_ctor_set(x_373, 4, x_138); +x_374 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_374, 0, x_372); +lean_ctor_set(x_374, 1, x_135); +lean_ctor_set(x_374, 2, x_136); +lean_ctor_set(x_374, 3, x_137); +lean_ctor_set(x_374, 4, x_373); +return x_374; +} +} +} +else +{ +if (lean_obj_tag(x_138) == 0) +{ +uint8_t x_375; +lean_inc(x_137); +lean_inc(x_136); +lean_inc(x_135); +x_375 = !lean_is_exclusive(x_5); +if (x_375 == 0) +{ +lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; uint8_t x_383; +x_376 = lean_ctor_get(x_5, 4); +lean_dec(x_376); +x_377 = lean_ctor_get(x_5, 3); +lean_dec(x_377); +x_378 = lean_ctor_get(x_5, 2); +lean_dec(x_378); +x_379 = lean_ctor_get(x_5, 1); +lean_dec(x_379); +x_380 = lean_ctor_get(x_5, 0); +lean_dec(x_380); +x_381 = lean_ctor_get(x_254, 0); +lean_inc(x_381); +x_382 = lean_ctor_get(x_254, 1); +lean_inc(x_382); +lean_dec_ref(x_254); +x_383 = !lean_is_exclusive(x_138); +if (x_383 == 0) +{ +lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; +x_384 = lean_ctor_get(x_138, 1); +x_385 = lean_ctor_get(x_138, 2); +x_386 = lean_ctor_get(x_138, 4); +lean_dec(x_386); +x_387 = lean_ctor_get(x_138, 3); +lean_dec(x_387); +x_388 = lean_ctor_get(x_138, 0); +lean_dec(x_388); +x_389 = lean_unsigned_to_nat(3u); +lean_ctor_set(x_138, 4, x_137); +lean_ctor_set(x_138, 3, x_137); +lean_ctor_set(x_138, 2, x_136); +lean_ctor_set(x_138, 1, x_135); +lean_ctor_set(x_138, 0, x_144); +if (lean_is_scalar(x_253)) { + x_390 = lean_alloc_ctor(0, 5, 0); +} else { + x_390 = x_253; +} +lean_ctor_set(x_390, 0, x_144); +lean_ctor_set(x_390, 1, x_381); +lean_ctor_set(x_390, 2, x_382); +lean_ctor_set(x_390, 3, x_137); +lean_ctor_set(x_390, 4, x_137); +lean_ctor_set(x_5, 4, x_390); +lean_ctor_set(x_5, 3, x_138); +lean_ctor_set(x_5, 2, x_385); +lean_ctor_set(x_5, 1, x_384); +lean_ctor_set(x_5, 0, x_389); +return x_5; +} +else +{ +lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; +x_391 = lean_ctor_get(x_138, 1); +x_392 = lean_ctor_get(x_138, 2); +lean_inc(x_392); +lean_inc(x_391); +lean_dec(x_138); +x_393 = lean_unsigned_to_nat(3u); +x_394 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_394, 0, x_144); +lean_ctor_set(x_394, 1, x_135); +lean_ctor_set(x_394, 2, x_136); +lean_ctor_set(x_394, 3, x_137); +lean_ctor_set(x_394, 4, x_137); +if (lean_is_scalar(x_253)) { + x_395 = lean_alloc_ctor(0, 5, 0); +} else { + x_395 = x_253; +} +lean_ctor_set(x_395, 0, x_144); +lean_ctor_set(x_395, 1, x_381); +lean_ctor_set(x_395, 2, x_382); +lean_ctor_set(x_395, 3, x_137); +lean_ctor_set(x_395, 4, x_137); +lean_ctor_set(x_5, 4, x_395); +lean_ctor_set(x_5, 3, x_394); +lean_ctor_set(x_5, 2, x_392); +lean_ctor_set(x_5, 1, x_391); +lean_ctor_set(x_5, 0, x_393); +return x_5; +} +} +else +{ +lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; +lean_dec(x_5); +x_396 = lean_ctor_get(x_254, 0); +lean_inc(x_396); +x_397 = lean_ctor_get(x_254, 1); +lean_inc(x_397); +lean_dec_ref(x_254); +x_398 = lean_ctor_get(x_138, 1); +lean_inc(x_398); +x_399 = lean_ctor_get(x_138, 2); +lean_inc(x_399); +if (lean_is_exclusive(x_138)) { + lean_ctor_release(x_138, 0); + lean_ctor_release(x_138, 1); + lean_ctor_release(x_138, 2); + lean_ctor_release(x_138, 3); + lean_ctor_release(x_138, 4); + x_400 = x_138; +} else { + lean_dec_ref(x_138); + x_400 = lean_box(0); +} +x_401 = lean_unsigned_to_nat(3u); +if (lean_is_scalar(x_400)) { + x_402 = lean_alloc_ctor(0, 5, 0); +} else { + x_402 = x_400; +} +lean_ctor_set(x_402, 0, x_144); +lean_ctor_set(x_402, 1, x_135); +lean_ctor_set(x_402, 2, x_136); +lean_ctor_set(x_402, 3, x_137); +lean_ctor_set(x_402, 4, x_137); +if (lean_is_scalar(x_253)) { + x_403 = lean_alloc_ctor(0, 5, 0); +} else { + x_403 = x_253; +} +lean_ctor_set(x_403, 0, x_144); +lean_ctor_set(x_403, 1, x_396); +lean_ctor_set(x_403, 2, x_397); +lean_ctor_set(x_403, 3, x_137); +lean_ctor_set(x_403, 4, x_137); +x_404 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_404, 0, x_401); +lean_ctor_set(x_404, 1, x_398); +lean_ctor_set(x_404, 2, x_399); +lean_ctor_set(x_404, 3, x_402); +lean_ctor_set(x_404, 4, x_403); +return x_404; +} +} +else +{ +lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; +x_405 = lean_ctor_get(x_254, 0); +lean_inc(x_405); +x_406 = lean_ctor_get(x_254, 1); +lean_inc(x_406); +lean_dec_ref(x_254); +x_407 = lean_unsigned_to_nat(2u); +if (lean_is_scalar(x_253)) { + x_408 = lean_alloc_ctor(0, 5, 0); +} else { + x_408 = x_253; +} +lean_ctor_set(x_408, 0, x_407); +lean_ctor_set(x_408, 1, x_405); +lean_ctor_set(x_408, 2, x_406); +lean_ctor_set(x_408, 3, x_5); +lean_ctor_set(x_408, 4, x_138); +return x_408; +} +} +} +} +} +else +{ +return x_5; +} +} +else +{ +return x_6; +} +} +} +else +{ +lean_object* x_409; lean_object* x_410; +x_409 = l_Std_DTreeMap_Internal_Impl_erase___at___00Lean_IR_StructRC_Context_remove_spec__0___redArg(x_1, x_5); +x_410 = lean_unsigned_to_nat(1u); +if (lean_obj_tag(x_409) == 0) +{ +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; uint8_t x_419; +x_411 = lean_ctor_get(x_409, 0); +lean_inc(x_411); +x_412 = lean_ctor_get(x_6, 0); +x_413 = lean_ctor_get(x_6, 1); +x_414 = lean_ctor_get(x_6, 2); +x_415 = lean_ctor_get(x_6, 3); +lean_inc(x_415); +x_416 = lean_ctor_get(x_6, 4); +x_417 = lean_unsigned_to_nat(3u); +x_418 = lean_nat_mul(x_417, x_411); +x_419 = lean_nat_dec_lt(x_418, x_412); +lean_dec(x_418); +if (x_419 == 0) +{ +lean_object* x_420; lean_object* x_421; lean_object* x_422; +lean_dec(x_415); +x_420 = lean_nat_add(x_410, x_411); +lean_dec(x_411); +x_421 = lean_nat_add(x_420, x_412); +lean_dec(x_420); +if (lean_is_scalar(x_7)) { + x_422 = lean_alloc_ctor(0, 5, 0); +} else { + x_422 = x_7; +} +lean_ctor_set(x_422, 0, x_421); +lean_ctor_set(x_422, 1, x_3); +lean_ctor_set(x_422, 2, x_4); +lean_ctor_set(x_422, 3, x_409); +lean_ctor_set(x_422, 4, x_6); +return x_422; +} +else +{ +lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; uint8_t x_432; +lean_inc(x_416); +lean_inc(x_414); +lean_inc(x_413); +lean_inc(x_412); +if (lean_is_exclusive(x_6)) { + lean_ctor_release(x_6, 0); + lean_ctor_release(x_6, 1); + lean_ctor_release(x_6, 2); + lean_ctor_release(x_6, 3); + lean_ctor_release(x_6, 4); + x_423 = x_6; +} else { + lean_dec_ref(x_6); + x_423 = lean_box(0); +} +x_424 = lean_ctor_get(x_415, 0); +x_425 = lean_ctor_get(x_415, 1); +x_426 = lean_ctor_get(x_415, 2); +x_427 = lean_ctor_get(x_415, 3); +x_428 = lean_ctor_get(x_415, 4); +x_429 = lean_ctor_get(x_416, 0); +x_430 = lean_unsigned_to_nat(2u); +x_431 = lean_nat_mul(x_430, x_429); +x_432 = lean_nat_dec_lt(x_424, x_431); +lean_dec(x_431); +if (x_432 == 0) +{ +lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_443; +lean_inc(x_428); +lean_inc(x_427); +lean_inc(x_426); +lean_inc(x_425); +if (lean_is_exclusive(x_415)) { + lean_ctor_release(x_415, 0); + lean_ctor_release(x_415, 1); + lean_ctor_release(x_415, 2); + lean_ctor_release(x_415, 3); + lean_ctor_release(x_415, 4); + x_433 = x_415; +} else { + lean_dec_ref(x_415); + x_433 = lean_box(0); +} +x_434 = lean_nat_add(x_410, x_411); +lean_dec(x_411); +x_435 = lean_nat_add(x_434, x_412); +lean_dec(x_412); +if (lean_obj_tag(x_427) == 0) +{ +lean_object* x_450; +x_450 = lean_ctor_get(x_427, 0); +lean_inc(x_450); +x_443 = x_450; +goto block_449; +} +else +{ +lean_object* x_451; +x_451 = lean_unsigned_to_nat(0u); +x_443 = x_451; +goto block_449; +} +block_442: +{ +lean_object* x_439; lean_object* x_440; lean_object* x_441; +x_439 = lean_nat_add(x_437, x_438); +lean_dec(x_438); +lean_dec(x_437); +if (lean_is_scalar(x_433)) { + x_440 = lean_alloc_ctor(0, 5, 0); +} else { + x_440 = x_433; +} +lean_ctor_set(x_440, 0, x_439); +lean_ctor_set(x_440, 1, x_413); +lean_ctor_set(x_440, 2, x_414); +lean_ctor_set(x_440, 3, x_428); +lean_ctor_set(x_440, 4, x_416); +if (lean_is_scalar(x_423)) { + x_441 = lean_alloc_ctor(0, 5, 0); +} else { + x_441 = x_423; +} +lean_ctor_set(x_441, 0, x_435); +lean_ctor_set(x_441, 1, x_425); +lean_ctor_set(x_441, 2, x_426); +lean_ctor_set(x_441, 3, x_436); +lean_ctor_set(x_441, 4, x_440); +return x_441; +} +block_449: +{ +lean_object* x_444; lean_object* x_445; lean_object* x_446; +x_444 = lean_nat_add(x_434, x_443); +lean_dec(x_443); +lean_dec(x_434); +if (lean_is_scalar(x_7)) { + x_445 = lean_alloc_ctor(0, 5, 0); +} else { + x_445 = x_7; +} +lean_ctor_set(x_445, 0, x_444); +lean_ctor_set(x_445, 1, x_3); +lean_ctor_set(x_445, 2, x_4); +lean_ctor_set(x_445, 3, x_409); +lean_ctor_set(x_445, 4, x_427); +x_446 = lean_nat_add(x_410, x_429); +if (lean_obj_tag(x_428) == 0) +{ +lean_object* x_447; +x_447 = lean_ctor_get(x_428, 0); +lean_inc(x_447); +x_436 = x_445; +x_437 = x_446; +x_438 = x_447; +goto block_442; +} +else +{ +lean_object* x_448; +x_448 = lean_unsigned_to_nat(0u); +x_436 = x_445; +x_437 = x_446; +x_438 = x_448; +goto block_442; +} +} +} +else +{ +lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; uint8_t x_456; +lean_dec(x_7); +x_452 = lean_nat_add(x_410, x_411); +lean_dec(x_411); +x_453 = lean_nat_add(x_452, x_412); +lean_dec(x_412); +x_454 = lean_nat_add(x_452, x_424); +lean_dec(x_452); +lean_inc_ref(x_409); +if (lean_is_scalar(x_423)) { + x_455 = lean_alloc_ctor(0, 5, 0); +} else { + x_455 = x_423; +} +lean_ctor_set(x_455, 0, x_454); +lean_ctor_set(x_455, 1, x_3); +lean_ctor_set(x_455, 2, x_4); +lean_ctor_set(x_455, 3, x_409); +lean_ctor_set(x_455, 4, x_415); +x_456 = !lean_is_exclusive(x_409); +if (x_456 == 0) +{ +lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; +x_457 = lean_ctor_get(x_409, 4); +lean_dec(x_457); +x_458 = lean_ctor_get(x_409, 3); +lean_dec(x_458); +x_459 = lean_ctor_get(x_409, 2); +lean_dec(x_459); +x_460 = lean_ctor_get(x_409, 1); +lean_dec(x_460); +x_461 = lean_ctor_get(x_409, 0); +lean_dec(x_461); +lean_ctor_set(x_409, 4, x_416); +lean_ctor_set(x_409, 3, x_455); +lean_ctor_set(x_409, 2, x_414); +lean_ctor_set(x_409, 1, x_413); +lean_ctor_set(x_409, 0, x_453); +return x_409; +} +else +{ +lean_object* x_462; +lean_dec(x_409); +x_462 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_462, 0, x_453); +lean_ctor_set(x_462, 1, x_413); +lean_ctor_set(x_462, 2, x_414); +lean_ctor_set(x_462, 3, x_455); +lean_ctor_set(x_462, 4, x_416); +return x_462; +} +} +} +} +else +{ +lean_object* x_463; lean_object* x_464; lean_object* x_465; +x_463 = lean_ctor_get(x_409, 0); +lean_inc(x_463); +x_464 = lean_nat_add(x_410, x_463); +lean_dec(x_463); +if (lean_is_scalar(x_7)) { + x_465 = lean_alloc_ctor(0, 5, 0); +} else { + x_465 = x_7; +} +lean_ctor_set(x_465, 0, x_464); +lean_ctor_set(x_465, 1, x_3); +lean_ctor_set(x_465, 2, x_4); +lean_ctor_set(x_465, 3, x_409); +lean_ctor_set(x_465, 4, x_6); +return x_465; +} +} +else +{ +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_466; +x_466 = lean_ctor_get(x_6, 3); +lean_inc(x_466); +if (lean_obj_tag(x_466) == 0) +{ +lean_object* x_467; +x_467 = lean_ctor_get(x_6, 4); +lean_inc(x_467); +if (lean_obj_tag(x_467) == 0) +{ +uint8_t x_468; +x_468 = !lean_is_exclusive(x_6); +if (x_468 == 0) +{ +lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; +x_469 = lean_ctor_get(x_6, 0); +x_470 = lean_ctor_get(x_6, 1); +x_471 = lean_ctor_get(x_6, 2); +x_472 = lean_ctor_get(x_6, 4); +lean_dec(x_472); +x_473 = lean_ctor_get(x_6, 3); +lean_dec(x_473); +x_474 = lean_ctor_get(x_466, 0); +x_475 = lean_nat_add(x_410, x_469); +lean_dec(x_469); +x_476 = lean_nat_add(x_410, x_474); +lean_ctor_set(x_6, 4, x_466); +lean_ctor_set(x_6, 3, x_409); +lean_ctor_set(x_6, 2, x_4); +lean_ctor_set(x_6, 1, x_3); +lean_ctor_set(x_6, 0, x_476); +if (lean_is_scalar(x_7)) { + x_477 = lean_alloc_ctor(0, 5, 0); +} else { + x_477 = x_7; +} +lean_ctor_set(x_477, 0, x_475); +lean_ctor_set(x_477, 1, x_470); +lean_ctor_set(x_477, 2, x_471); +lean_ctor_set(x_477, 3, x_6); +lean_ctor_set(x_477, 4, x_467); +return x_477; +} +else +{ +lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; +x_478 = lean_ctor_get(x_6, 0); +x_479 = lean_ctor_get(x_6, 1); +x_480 = lean_ctor_get(x_6, 2); +lean_inc(x_480); +lean_inc(x_479); +lean_inc(x_478); +lean_dec(x_6); +x_481 = lean_ctor_get(x_466, 0); +x_482 = lean_nat_add(x_410, x_478); +lean_dec(x_478); +x_483 = lean_nat_add(x_410, x_481); +x_484 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_484, 0, x_483); +lean_ctor_set(x_484, 1, x_3); +lean_ctor_set(x_484, 2, x_4); +lean_ctor_set(x_484, 3, x_409); +lean_ctor_set(x_484, 4, x_466); +if (lean_is_scalar(x_7)) { + x_485 = lean_alloc_ctor(0, 5, 0); +} else { + x_485 = x_7; +} +lean_ctor_set(x_485, 0, x_482); +lean_ctor_set(x_485, 1, x_479); +lean_ctor_set(x_485, 2, x_480); +lean_ctor_set(x_485, 3, x_484); +lean_ctor_set(x_485, 4, x_467); +return x_485; +} +} +else +{ +uint8_t x_486; +x_486 = !lean_is_exclusive(x_6); +if (x_486 == 0) +{ +lean_object* x_487; lean_object* x_488; lean_object* x_489; uint8_t x_490; +x_487 = lean_ctor_get(x_6, 4); +lean_dec(x_487); +x_488 = lean_ctor_get(x_6, 3); +lean_dec(x_488); +x_489 = lean_ctor_get(x_6, 0); +lean_dec(x_489); +x_490 = !lean_is_exclusive(x_466); +if (x_490 == 0) +{ +lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; +x_491 = lean_ctor_get(x_466, 1); +x_492 = lean_ctor_get(x_466, 2); +x_493 = lean_ctor_get(x_466, 4); +lean_dec(x_493); +x_494 = lean_ctor_get(x_466, 3); +lean_dec(x_494); +x_495 = lean_ctor_get(x_466, 0); +lean_dec(x_495); +x_496 = lean_unsigned_to_nat(3u); +lean_ctor_set(x_466, 4, x_467); +lean_ctor_set(x_466, 3, x_467); +lean_ctor_set(x_466, 2, x_4); +lean_ctor_set(x_466, 1, x_3); +lean_ctor_set(x_466, 0, x_410); +lean_ctor_set(x_6, 3, x_467); +lean_ctor_set(x_6, 0, x_410); +if (lean_is_scalar(x_7)) { + x_497 = lean_alloc_ctor(0, 5, 0); +} else { + x_497 = x_7; +} +lean_ctor_set(x_497, 0, x_496); +lean_ctor_set(x_497, 1, x_491); +lean_ctor_set(x_497, 2, x_492); +lean_ctor_set(x_497, 3, x_466); +lean_ctor_set(x_497, 4, x_6); +return x_497; +} +else +{ +lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; +x_498 = lean_ctor_get(x_466, 1); +x_499 = lean_ctor_get(x_466, 2); +lean_inc(x_499); +lean_inc(x_498); +lean_dec(x_466); +x_500 = lean_unsigned_to_nat(3u); +x_501 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_501, 0, x_410); +lean_ctor_set(x_501, 1, x_3); +lean_ctor_set(x_501, 2, x_4); +lean_ctor_set(x_501, 3, x_467); +lean_ctor_set(x_501, 4, x_467); +lean_ctor_set(x_6, 3, x_467); +lean_ctor_set(x_6, 0, x_410); +if (lean_is_scalar(x_7)) { + x_502 = lean_alloc_ctor(0, 5, 0); +} else { + x_502 = x_7; +} +lean_ctor_set(x_502, 0, x_500); +lean_ctor_set(x_502, 1, x_498); +lean_ctor_set(x_502, 2, x_499); +lean_ctor_set(x_502, 3, x_501); +lean_ctor_set(x_502, 4, x_6); +return x_502; +} +} +else +{ +lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; +x_503 = lean_ctor_get(x_6, 1); +x_504 = lean_ctor_get(x_6, 2); +lean_inc(x_504); +lean_inc(x_503); +lean_dec(x_6); +x_505 = lean_ctor_get(x_466, 1); +lean_inc(x_505); +x_506 = lean_ctor_get(x_466, 2); +lean_inc(x_506); +if (lean_is_exclusive(x_466)) { + lean_ctor_release(x_466, 0); + lean_ctor_release(x_466, 1); + lean_ctor_release(x_466, 2); + lean_ctor_release(x_466, 3); + lean_ctor_release(x_466, 4); + x_507 = x_466; +} else { + lean_dec_ref(x_466); + x_507 = lean_box(0); +} +x_508 = lean_unsigned_to_nat(3u); +if (lean_is_scalar(x_507)) { + x_509 = lean_alloc_ctor(0, 5, 0); +} else { + x_509 = x_507; +} +lean_ctor_set(x_509, 0, x_410); +lean_ctor_set(x_509, 1, x_3); +lean_ctor_set(x_509, 2, x_4); +lean_ctor_set(x_509, 3, x_467); +lean_ctor_set(x_509, 4, x_467); +x_510 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_510, 0, x_410); +lean_ctor_set(x_510, 1, x_503); +lean_ctor_set(x_510, 2, x_504); +lean_ctor_set(x_510, 3, x_467); +lean_ctor_set(x_510, 4, x_467); +if (lean_is_scalar(x_7)) { + x_511 = lean_alloc_ctor(0, 5, 0); +} else { + x_511 = x_7; +} +lean_ctor_set(x_511, 0, x_508); +lean_ctor_set(x_511, 1, x_505); +lean_ctor_set(x_511, 2, x_506); +lean_ctor_set(x_511, 3, x_509); +lean_ctor_set(x_511, 4, x_510); +return x_511; +} +} +} +else +{ +lean_object* x_512; +x_512 = lean_ctor_get(x_6, 4); +lean_inc(x_512); +if (lean_obj_tag(x_512) == 0) +{ +uint8_t x_513; +x_513 = !lean_is_exclusive(x_6); +if (x_513 == 0) +{ +lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; +x_514 = lean_ctor_get(x_6, 1); +x_515 = lean_ctor_get(x_6, 2); +x_516 = lean_ctor_get(x_6, 4); +lean_dec(x_516); +x_517 = lean_ctor_get(x_6, 3); +lean_dec(x_517); +x_518 = lean_ctor_get(x_6, 0); +lean_dec(x_518); +x_519 = lean_unsigned_to_nat(3u); +lean_ctor_set(x_6, 4, x_466); +lean_ctor_set(x_6, 2, x_4); +lean_ctor_set(x_6, 1, x_3); +lean_ctor_set(x_6, 0, x_410); +if (lean_is_scalar(x_7)) { + x_520 = lean_alloc_ctor(0, 5, 0); +} else { + x_520 = x_7; +} +lean_ctor_set(x_520, 0, x_519); +lean_ctor_set(x_520, 1, x_514); +lean_ctor_set(x_520, 2, x_515); +lean_ctor_set(x_520, 3, x_6); +lean_ctor_set(x_520, 4, x_512); +return x_520; +} +else +{ +lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; +x_521 = lean_ctor_get(x_6, 1); +x_522 = lean_ctor_get(x_6, 2); +lean_inc(x_522); +lean_inc(x_521); +lean_dec(x_6); +x_523 = lean_unsigned_to_nat(3u); +x_524 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_524, 0, x_410); +lean_ctor_set(x_524, 1, x_3); +lean_ctor_set(x_524, 2, x_4); +lean_ctor_set(x_524, 3, x_466); +lean_ctor_set(x_524, 4, x_466); +if (lean_is_scalar(x_7)) { + x_525 = lean_alloc_ctor(0, 5, 0); +} else { + x_525 = x_7; +} +lean_ctor_set(x_525, 0, x_523); +lean_ctor_set(x_525, 1, x_521); +lean_ctor_set(x_525, 2, x_522); +lean_ctor_set(x_525, 3, x_524); +lean_ctor_set(x_525, 4, x_512); +return x_525; +} +} +else +{ +uint8_t x_526; +x_526 = !lean_is_exclusive(x_6); +if (x_526 == 0) +{ +lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; +x_527 = lean_ctor_get(x_6, 4); +lean_dec(x_527); +x_528 = lean_ctor_get(x_6, 3); +lean_dec(x_528); +lean_ctor_set(x_6, 3, x_512); +x_529 = lean_unsigned_to_nat(2u); +if (lean_is_scalar(x_7)) { + x_530 = lean_alloc_ctor(0, 5, 0); +} else { + x_530 = x_7; +} +lean_ctor_set(x_530, 0, x_529); +lean_ctor_set(x_530, 1, x_3); +lean_ctor_set(x_530, 2, x_4); +lean_ctor_set(x_530, 3, x_512); +lean_ctor_set(x_530, 4, x_6); +return x_530; +} +else +{ +lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; +x_531 = lean_ctor_get(x_6, 0); +x_532 = lean_ctor_get(x_6, 1); +x_533 = lean_ctor_get(x_6, 2); +lean_inc(x_533); +lean_inc(x_532); +lean_inc(x_531); +lean_dec(x_6); +x_534 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_534, 0, x_531); +lean_ctor_set(x_534, 1, x_532); +lean_ctor_set(x_534, 2, x_533); +lean_ctor_set(x_534, 3, x_512); +lean_ctor_set(x_534, 4, x_512); +x_535 = lean_unsigned_to_nat(2u); +if (lean_is_scalar(x_7)) { + x_536 = lean_alloc_ctor(0, 5, 0); +} else { + x_536 = x_7; +} +lean_ctor_set(x_536, 0, x_535); +lean_ctor_set(x_536, 1, x_3); +lean_ctor_set(x_536, 2, x_4); +lean_ctor_set(x_536, 3, x_512); +lean_ctor_set(x_536, 4, x_534); +return x_536; +} +} +} +} +else +{ +lean_object* x_537; +if (lean_is_scalar(x_7)) { + x_537 = lean_alloc_ctor(0, 5, 0); +} else { + x_537 = x_7; +} +lean_ctor_set(x_537, 0, x_410); +lean_ctor_set(x_537, 1, x_3); +lean_ctor_set(x_537, 2, x_4); +lean_ctor_set(x_537, 3, x_6); +lean_ctor_set(x_537, 4, x_6); +return x_537; +} +} +} +} +else +{ +return x_2; +} +} +} +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_erase___at___00Lean_IR_StructRC_Context_remove_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Std_DTreeMap_Internal_Impl_erase___at___00Lean_IR_StructRC_Context_remove_spec__0___redArg(x_2, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_remove(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_1, 0); +x_5 = l_Std_DTreeMap_Internal_Impl_erase___at___00Lean_IR_StructRC_Context_remove_spec__0___redArg(x_2, x_4); +lean_ctor_set(x_1, 0, x_5); +return x_1; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_6 = lean_ctor_get(x_1, 0); +x_7 = lean_ctor_get(x_1, 1); +x_8 = lean_ctor_get(x_1, 2); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_dec(x_1); +x_9 = l_Std_DTreeMap_Internal_Impl_erase___at___00Lean_IR_StructRC_Context_remove_spec__0___redArg(x_2, x_6); +x_10 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_7); +lean_ctor_set(x_10, 2, x_8); +return x_10; +} +} +} +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_erase___at___00Lean_IR_StructRC_Context_remove_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Std_DTreeMap_Internal_Impl_erase___at___00Lean_IR_StructRC_Context_remove_spec__0(x_1, x_2, x_3, x_4); +lean_dec(x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_remove___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_IR_StructRC_Context_remove(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_erase___at___00Lean_IR_StructRC_Context_remove_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_DTreeMap_Internal_Impl_erase___at___00Lean_IR_StructRC_Context_remove_spec__0___redArg(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_addInstr(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_1, 1); +x_5 = lean_array_push(x_4, x_2); +lean_ctor_set(x_1, 1, x_5); +return x_1; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_6 = lean_ctor_get(x_1, 0); +x_7 = lean_ctor_get(x_1, 1); +x_8 = lean_ctor_get(x_1, 2); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_dec(x_1); +x_9 = lean_array_push(x_7, x_2); +x_10 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_10, 0, x_6); +lean_ctor_set(x_10, 1, x_9); +lean_ctor_set(x_10, 2, x_8); +return x_10; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_mkFreshVar(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_unsigned_to_nat(1u); +x_3 = lean_nat_add(x_1, x_2); +x_4 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Lean_IR_StructRC_needsRC(lean_object* x_1) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 7: +{ +uint8_t x_2; +x_2 = 1; +return x_2; +} +case 8: +{ +uint8_t x_3; +x_3 = 1; +return x_3; +} +case 11: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_4 = lean_ctor_get(x_1, 1); +x_5 = lean_unsigned_to_nat(0u); +x_6 = lean_array_get_size(x_4); +x_7 = lean_nat_dec_lt(x_5, x_6); +if (x_7 == 0) +{ +return x_7; +} +else +{ +if (x_7 == 0) +{ +return x_7; +} +else +{ +size_t x_8; size_t x_9; uint8_t x_10; +x_8 = 0; +x_9 = lean_usize_of_nat(x_6); +x_10 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_StructRC_needsRC_spec__0(x_4, x_8, x_9); +return x_10; +} +} +} +case 10: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_ctor_get(x_1, 1); +x_12 = lean_unsigned_to_nat(0u); +x_13 = lean_array_get_size(x_11); +x_14 = lean_nat_dec_lt(x_12, x_13); +if (x_14 == 0) +{ +return x_14; +} +else +{ +if (x_14 == 0) +{ +return x_14; +} +else +{ +size_t x_15; size_t x_16; uint8_t x_17; +x_15 = 0; +x_16 = lean_usize_of_nat(x_13); +x_17 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_StructRC_needsRC_spec__0(x_11, x_15, x_16); +return x_17; +} +} +} +default: +{ +uint8_t x_18; +x_18 = 0; +return x_18; +} +} +} +} +LEAN_EXPORT uint8_t l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_StructRC_needsRC_spec__0(lean_object* x_1, size_t x_2, size_t x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_eq(x_2, x_3); +if (x_4 == 0) +{ +lean_object* x_5; uint8_t x_6; +x_5 = lean_array_uget(x_1, x_2); +x_6 = l_Lean_IR_StructRC_needsRC(x_5); +lean_dec(x_5); +if (x_6 == 0) +{ +size_t x_7; size_t x_8; +x_7 = 1; +x_8 = lean_usize_add(x_2, x_7); +x_2 = x_8; +goto _start; +} +else +{ +return x_6; +} +} +else +{ +uint8_t x_10; +x_10 = 0; +return x_10; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_StructRC_needsRC_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; uint8_t x_6; lean_object* x_7; +x_4 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_5 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_6 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_StructRC_needsRC_spec__0(x_1, x_4, x_5); +lean_dec_ref(x_1); +x_7 = lean_box(x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_needsRC___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_IR_StructRC_needsRC(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_IR_StructRC_Entry_ofStruct___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.IR.StructRC.Entry.ofStruct", 31, 31); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_StructRC_Entry_ofStruct___closed__1() { +_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_IR_StructRC_Entry_ofStruct___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_IR_StructRC_Entry_ofStruct___closed__1; +x_2 = lean_unsigned_to_nat(9u); +x_3 = lean_unsigned_to_nat(128u); +x_4 = l_Lean_IR_StructRC_Entry_ofStruct___closed__0; +x_5 = l_Lean_IR_StructRC_Context_renameVar_x21___closed__0; +x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_ofStruct(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_2) == 10) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_array_get_size(x_4); +x_6 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_6, 0, x_3); +x_7 = lean_mk_array(x_5, x_6); +lean_inc_ref(x_4); +x_8 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_8, 0, x_1); +lean_ctor_set(x_8, 1, x_4); +lean_ctor_set(x_8, 2, x_7); +return x_8; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +lean_dec(x_3); +lean_dec(x_1); +x_9 = lean_box(0); +x_10 = l_Lean_IR_StructRC_Entry_ofStruct___closed__2; +x_11 = l_panic___redArg(x_9, x_10); +return x_11; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_ofStruct___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_IR_StructRC_Entry_ofStruct(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Entry_ofType(lean_object* x_1, lean_object* x_2) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 10: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_3 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_3); +lean_dec_ref(x_1); +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_array_get_size(x_3); +x_6 = lean_nat_dec_lt(x_4, x_5); +if (x_6 == 0) +{ +lean_object* x_7; +lean_dec_ref(x_3); +lean_dec(x_2); +x_7 = lean_box(0); +return x_7; +} +else +{ +if (x_6 == 0) +{ +lean_object* x_8; +lean_dec_ref(x_3); +lean_dec(x_2); +x_8 = lean_box(0); +return x_8; +} +else +{ +size_t x_9; size_t x_10; uint8_t x_11; +x_9 = 0; +x_10 = lean_usize_of_nat(x_5); +x_11 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_StructRC_needsRC_spec__0(x_3, x_9, x_10); +if (x_11 == 0) +{ +lean_object* x_12; +lean_dec_ref(x_3); +lean_dec(x_2); +x_12 = lean_box(0); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_13, 0, x_2); +x_14 = lean_mk_array(x_5, x_13); +x_15 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_15, 0, x_4); +lean_ctor_set(x_15, 1, x_3); +lean_ctor_set(x_15, 2, x_14); +return x_15; +} +} +} +} +case 11: +{ +uint8_t x_16; +x_16 = !lean_is_exclusive(x_1); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_17 = lean_ctor_get(x_1, 1); +x_18 = lean_ctor_get(x_1, 0); +lean_dec(x_18); +x_19 = lean_unsigned_to_nat(0u); +x_20 = lean_array_get_size(x_17); +x_21 = lean_nat_dec_lt(x_19, x_20); +if (x_21 == 0) +{ +lean_object* x_22; +lean_free_object(x_1); +lean_dec_ref(x_17); +lean_dec(x_2); +x_22 = lean_box(0); +return x_22; +} +else +{ +if (x_21 == 0) +{ +lean_object* x_23; +lean_free_object(x_1); +lean_dec_ref(x_17); +lean_dec(x_2); +x_23 = lean_box(0); +return x_23; +} +else +{ +size_t x_24; size_t x_25; uint8_t x_26; +x_24 = 0; +x_25 = lean_usize_of_nat(x_20); +x_26 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_StructRC_needsRC_spec__0(x_17, x_24, x_25); +if (x_26 == 0) +{ +lean_object* x_27; +lean_free_object(x_1); +lean_dec_ref(x_17); +lean_dec(x_2); +x_27 = lean_box(0); +return x_27; +} +else +{ +lean_ctor_set_tag(x_1, 1); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set(x_1, 0, x_17); +return x_1; +} +} +} +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_28 = lean_ctor_get(x_1, 1); +lean_inc(x_28); +lean_dec(x_1); +x_29 = lean_unsigned_to_nat(0u); +x_30 = lean_array_get_size(x_28); +x_31 = lean_nat_dec_lt(x_29, x_30); +if (x_31 == 0) +{ +lean_object* x_32; +lean_dec_ref(x_28); +lean_dec(x_2); +x_32 = lean_box(0); +return x_32; +} +else +{ +if (x_31 == 0) +{ +lean_object* x_33; +lean_dec_ref(x_28); +lean_dec(x_2); +x_33 = lean_box(0); +return x_33; +} +else +{ +size_t x_34; size_t x_35; uint8_t x_36; +x_34 = 0; +x_35 = lean_usize_of_nat(x_30); +x_36 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_StructRC_needsRC_spec__0(x_28, x_34, x_35); +if (x_36 == 0) +{ +lean_object* x_37; +lean_dec_ref(x_28); +lean_dec(x_2); +x_37 = lean_box(0); +return x_37; +} +else +{ +lean_object* x_38; +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_28); +lean_ctor_set(x_38, 1, x_2); +return x_38; +} +} +} +} +} +case 7: +{ +uint8_t x_39; lean_object* x_40; +x_39 = 0; +x_40 = lean_alloc_ctor(3, 1, 1); +lean_ctor_set(x_40, 0, x_2); +lean_ctor_set_uint8(x_40, sizeof(void*)*1, x_39); +return x_40; +} +case 8: +{ +uint8_t x_41; lean_object* x_42; +x_41 = 1; +x_42 = lean_alloc_ctor(3, 1, 1); +lean_ctor_set(x_42, 0, x_2); +lean_ctor_set_uint8(x_42, sizeof(void*)*1, x_41); +return x_42; +} +default: +{ +lean_object* x_43; +lean_dec(x_2); +lean_dec(x_1); +x_43 = lean_box(0); +return x_43; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_addVar(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_IR_StructRC_Entry_ofType(x_3, x_4); +x_6 = l_Lean_IR_StructRC_Context_insertIfNew(x_2, x_5, x_1); +return x_6; +} +} +static lean_object* _init_l_Lean_IR_StructRC_Context_addVarBasic___closed__0() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__10; +x_2 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_addVarBasic(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +switch (lean_obj_tag(x_3)) { +case 10: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_10 = lean_ctor_get(x_3, 1); +lean_inc_ref(x_10); +lean_dec_ref(x_3); +x_11 = lean_unsigned_to_nat(0u); +x_12 = lean_array_get_size(x_10); +x_13 = lean_nat_dec_lt(x_11, x_12); +if (x_13 == 0) +{ +lean_dec_ref(x_10); +goto block_6; +} +else +{ +if (x_13 == 0) +{ +lean_dec_ref(x_10); +goto block_6; +} +else +{ +size_t x_14; size_t x_15; uint8_t x_16; +x_14 = 0; +x_15 = lean_usize_of_nat(x_12); +x_16 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_StructRC_needsRC_spec__0(x_10, x_14, x_15); +if (x_16 == 0) +{ +lean_dec_ref(x_10); +goto block_6; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = l_Lean_IR_StructRC_Context_addVarBasic___closed__0; +x_18 = lean_mk_array(x_12, x_17); +x_19 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_19, 0, x_11); +lean_ctor_set(x_19, 1, x_10); +lean_ctor_set(x_19, 2, x_18); +x_20 = l_Lean_IR_StructRC_Context_insert(x_2, x_19, x_1); +return x_20; +} +} +} +} +case 11: +{ +uint8_t x_21; +x_21 = !lean_is_exclusive(x_3); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_22 = lean_ctor_get(x_3, 1); +x_23 = lean_ctor_get(x_3, 0); +lean_dec(x_23); +x_24 = lean_unsigned_to_nat(0u); +x_25 = lean_array_get_size(x_22); +x_26 = lean_nat_dec_lt(x_24, x_25); +if (x_26 == 0) +{ +lean_free_object(x_3); +lean_dec_ref(x_22); +goto block_9; +} +else +{ +if (x_26 == 0) +{ +lean_free_object(x_3); +lean_dec_ref(x_22); +goto block_9; +} +else +{ +size_t x_27; size_t x_28; uint8_t x_29; +x_27 = 0; +x_28 = lean_usize_of_nat(x_25); +x_29 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_StructRC_needsRC_spec__0(x_22, x_27, x_28); +if (x_29 == 0) +{ +lean_free_object(x_3); +lean_dec_ref(x_22); +goto block_9; +} +else +{ +lean_object* x_30; lean_object* x_31; +x_30 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__10; +lean_ctor_set_tag(x_3, 1); +lean_ctor_set(x_3, 1, x_30); +lean_ctor_set(x_3, 0, x_22); +x_31 = l_Lean_IR_StructRC_Context_insert(x_2, x_3, x_1); +return x_31; +} +} +} +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_32 = lean_ctor_get(x_3, 1); +lean_inc(x_32); +lean_dec(x_3); +x_33 = lean_unsigned_to_nat(0u); +x_34 = lean_array_get_size(x_32); +x_35 = lean_nat_dec_lt(x_33, x_34); +if (x_35 == 0) +{ +lean_dec_ref(x_32); +goto block_9; +} +else +{ +if (x_35 == 0) +{ +lean_dec_ref(x_32); +goto block_9; +} +else +{ +size_t x_36; size_t x_37; uint8_t x_38; +x_36 = 0; +x_37 = lean_usize_of_nat(x_34); +x_38 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_StructRC_needsRC_spec__0(x_32, x_36, x_37); +if (x_38 == 0) +{ +lean_dec_ref(x_32); +goto block_9; +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__10; +x_40 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_40, 0, x_32); +lean_ctor_set(x_40, 1, x_39); +x_41 = l_Lean_IR_StructRC_Context_insert(x_2, x_40, x_1); +return x_41; +} +} +} +} +} +default: +{ +lean_dec(x_3); +lean_dec(x_2); +return x_1; +} +} +block_6: +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_box(0); +x_5 = l_Lean_IR_StructRC_Context_insert(x_2, x_4, x_1); +return x_5; +} +block_9: +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_box(0); +x_8 = l_Lean_IR_StructRC_Context_insert(x_2, x_7, x_1); +return x_8; +} +} +} +static lean_object* _init_l_Lean_IR_StructRC_Context_addProjInfo___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Id_instMonad___lam__0), 4, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_StructRC_Context_addProjInfo___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Id_instMonad___lam__1___boxed), 4, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_StructRC_Context_addProjInfo___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Id_instMonad___lam__2___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_StructRC_Context_addProjInfo___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Id_instMonad___lam__3), 4, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_StructRC_Context_addProjInfo___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Id_instMonad___lam__4___boxed), 4, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_StructRC_Context_addProjInfo___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Id_instMonad___lam__5___boxed), 4, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_StructRC_Context_addProjInfo___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Id_instMonad___lam__6), 4, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_StructRC_Context_addProjInfo___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_StructRC_Context_addProjInfo___closed__1; +x_2 = l_Lean_IR_StructRC_Context_addProjInfo___closed__0; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_IR_StructRC_Context_addProjInfo___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_IR_StructRC_Context_addProjInfo___closed__5; +x_2 = l_Lean_IR_StructRC_Context_addProjInfo___closed__4; +x_3 = l_Lean_IR_StructRC_Context_addProjInfo___closed__3; +x_4 = l_Lean_IR_StructRC_Context_addProjInfo___closed__2; +x_5 = l_Lean_IR_StructRC_Context_addProjInfo___closed__7; +x_6 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_4); +lean_ctor_set(x_6, 2, x_3); +lean_ctor_set(x_6, 3, x_2); +lean_ctor_set(x_6, 4, x_1); +return x_6; +} +} +static lean_object* _init_l_Lean_IR_StructRC_Context_addProjInfo___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_StructRC_Context_addProjInfo___closed__6; +x_2 = l_Lean_IR_StructRC_Context_addProjInfo___closed__8; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_IR_StructRC_Context_addProjInfo___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.IR.StructRC.Context.addProjInfo", 36, 36); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_StructRC_Context_addProjInfo___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_IR_StructRC_Entry_ofStruct___closed__1; +x_2 = lean_unsigned_to_nat(47u); +x_3 = lean_unsigned_to_nat(171u); +x_4 = l_Lean_IR_StructRC_Context_addProjInfo___closed__10; +x_5 = l_Lean_IR_StructRC_Context_renameVar_x21___closed__0; +x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_addProjInfo(lean_object* x_1, 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_IR_StructRC_Context_addProjInfo___closed__9; +if (lean_obj_tag(x_6) == 2) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_53; uint8_t x_54; +x_8 = lean_ctor_get(x_6, 0); +x_9 = lean_ctor_get(x_6, 1); +x_10 = lean_ctor_get(x_6, 2); +x_53 = lean_array_get_size(x_9); +x_54 = lean_nat_dec_lt(x_4, x_53); +if (x_54 == 0) +{ +lean_object* x_55; lean_object* x_56; +x_55 = lean_box(0); +x_56 = l_outOfBounds___redArg(x_55); +x_11 = x_56; +goto block_52; +} +else +{ +lean_object* x_57; +x_57 = lean_array_fget_borrowed(x_10, x_4); +lean_inc(x_57); +x_11 = x_57; +goto block_52; +} +block_52: +{ +switch (lean_obj_tag(x_11)) { +case 0: +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +lean_dec(x_4); +lean_dec(x_3); +x_12 = l_Lean_IR_StructRC_Context_insert(x_5, x_6, x_1); +x_13 = lean_box(1); +x_14 = l_Lean_IR_StructRC_Context_insertRename(x_12, x_2, x_13); +return x_14; +} +case 1: +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +lean_dec(x_4); +lean_dec(x_3); +x_15 = lean_ctor_get(x_11, 0); +lean_inc(x_15); +lean_dec_ref(x_11); +x_16 = l_Lean_IR_StructRC_Context_insert(x_5, x_6, x_1); +x_17 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_17, 0, x_15); +x_18 = l_Lean_IR_StructRC_Context_insertRename(x_16, x_2, x_17); +return x_18; +} +default: +{ +uint8_t x_19; +lean_inc_ref(x_10); +lean_inc_ref(x_9); +lean_inc(x_8); +x_19 = !lean_is_exclusive(x_6); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_20 = lean_ctor_get(x_6, 2); +lean_dec(x_20); +x_21 = lean_ctor_get(x_6, 1); +lean_dec(x_21); +x_22 = lean_ctor_get(x_6, 0); +lean_dec(x_22); +x_23 = !lean_is_exclusive(x_11); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_24 = lean_ctor_get(x_11, 0); +lean_inc(x_3); +lean_inc(x_2); +x_25 = l_Lean_IR_StructRC_Context_addVar(x_1, x_2, x_3, x_24); +lean_inc(x_2); +lean_ctor_set_tag(x_11, 1); +lean_ctor_set(x_11, 0, x_2); +x_26 = lean_array_set(x_10, x_4, x_11); +lean_inc(x_8); +lean_ctor_set(x_6, 2, x_26); +lean_inc(x_5); +x_27 = l_Lean_IR_StructRC_Context_insert(x_5, x_6, x_25); +x_28 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_28, 0, x_8); +lean_ctor_set(x_28, 1, x_4); +lean_ctor_set(x_28, 2, x_5); +x_29 = lean_box(12); +x_30 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_30, 0, x_2); +lean_ctor_set(x_30, 1, x_3); +lean_ctor_set(x_30, 2, x_28); +lean_ctor_set(x_30, 3, x_29); +x_31 = l_Lean_IR_StructRC_Context_addInstr(x_27, x_30); +return x_31; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_32 = lean_ctor_get(x_11, 0); +lean_inc(x_32); +lean_dec(x_11); +lean_inc(x_3); +lean_inc(x_2); +x_33 = l_Lean_IR_StructRC_Context_addVar(x_1, x_2, x_3, x_32); +lean_inc(x_2); +x_34 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_34, 0, x_2); +x_35 = lean_array_set(x_10, x_4, x_34); +lean_inc(x_8); +lean_ctor_set(x_6, 2, x_35); +lean_inc(x_5); +x_36 = l_Lean_IR_StructRC_Context_insert(x_5, x_6, x_33); +x_37 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_37, 0, x_8); +lean_ctor_set(x_37, 1, x_4); +lean_ctor_set(x_37, 2, x_5); +x_38 = lean_box(12); +x_39 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_39, 0, x_2); +lean_ctor_set(x_39, 1, x_3); +lean_ctor_set(x_39, 2, x_37); +lean_ctor_set(x_39, 3, x_38); +x_40 = l_Lean_IR_StructRC_Context_addInstr(x_36, x_39); +return x_40; +} +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +lean_dec(x_6); +x_41 = lean_ctor_get(x_11, 0); +lean_inc(x_41); +if (lean_is_exclusive(x_11)) { + lean_ctor_release(x_11, 0); + x_42 = x_11; +} else { + lean_dec_ref(x_11); + x_42 = lean_box(0); +} +lean_inc(x_3); +lean_inc(x_2); +x_43 = l_Lean_IR_StructRC_Context_addVar(x_1, x_2, x_3, x_41); +lean_inc(x_2); +if (lean_is_scalar(x_42)) { + x_44 = lean_alloc_ctor(1, 1, 0); +} else { + x_44 = x_42; + lean_ctor_set_tag(x_44, 1); +} +lean_ctor_set(x_44, 0, x_2); +x_45 = lean_array_set(x_10, x_4, x_44); +lean_inc(x_8); +x_46 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_46, 0, x_8); +lean_ctor_set(x_46, 1, x_9); +lean_ctor_set(x_46, 2, x_45); +lean_inc(x_5); +x_47 = l_Lean_IR_StructRC_Context_insert(x_5, x_46, x_43); +x_48 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_48, 0, x_8); +lean_ctor_set(x_48, 1, x_4); +lean_ctor_set(x_48, 2, x_5); +x_49 = lean_box(12); +x_50 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_50, 0, x_2); +lean_ctor_set(x_50, 1, x_3); +lean_ctor_set(x_50, 2, x_48); +lean_ctor_set(x_50, 3, x_49); +x_51 = l_Lean_IR_StructRC_Context_addInstr(x_47, x_50); +return x_51; +} +} +} +} +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_58 = l_instInhabitedOfMonad___redArg(x_7, x_1); +x_59 = l_Lean_IR_StructRC_Context_addProjInfo___closed__11; +x_60 = l_panic___redArg(x_58, x_59); +return x_60; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_emitRCDiff(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; uint8_t x_6; +x_5 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__10; +x_6 = lean_int_dec_lt(x_5, x_3); +if (x_6 == 0) +{ +uint8_t x_7; +x_7 = lean_int_dec_lt(x_3, x_5); +if (x_7 == 0) +{ +lean_dec(x_1); +return x_4; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_nat_abs(x_3); +x_9 = lean_box(12); +x_10 = lean_alloc_ctor(7, 3, 2); +lean_ctor_set(x_10, 0, x_1); +lean_ctor_set(x_10, 1, x_8); +lean_ctor_set(x_10, 2, x_9); +lean_ctor_set_uint8(x_10, sizeof(void*)*3, x_2); +lean_ctor_set_uint8(x_10, sizeof(void*)*3 + 1, x_6); +x_11 = l_Lean_IR_StructRC_Context_addInstr(x_4, x_10); +return x_11; +} +} +else +{ +lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = lean_nat_abs(x_3); +x_13 = 0; +x_14 = lean_box(12); +x_15 = lean_alloc_ctor(6, 3, 2); +lean_ctor_set(x_15, 0, x_1); +lean_ctor_set(x_15, 1, x_12); +lean_ctor_set(x_15, 2, x_14); +lean_ctor_set_uint8(x_15, sizeof(void*)*3, x_2); +lean_ctor_set_uint8(x_15, sizeof(void*)*3 + 1, x_13); +x_16 = l_Lean_IR_StructRC_Context_addInstr(x_4, x_15); +return x_16; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_emitRCDiff___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; lean_object* x_6; +x_5 = lean_unbox(x_2); +x_6 = l_Lean_IR_StructRC_Context_emitRCDiff(x_1, x_5, x_3, x_4); +lean_dec(x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Vector_Basic_0__Vector_mapM_go___at___00Lean_IR_StructRC_Context_accumulateRCDiff_spec__0___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_nat_dec_lt(x_4, x_1); +if (x_7 == 0) +{ +lean_object* x_8; +lean_dec(x_4); +lean_dec_ref(x_2); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_5); +lean_ctor_set(x_8, 1, x_6); +return x_8; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_9 = lean_array_fget_borrowed(x_3, x_4); +lean_inc_ref(x_2); +lean_inc(x_9); +x_10 = lean_apply_2(x_2, x_9, x_6); +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_ref(x_10); +x_13 = lean_unsigned_to_nat(1u); +x_14 = lean_nat_add(x_4, x_13); +lean_dec(x_4); +x_15 = lean_array_push(x_5, x_11); +x_4 = x_14; +x_5 = x_15; +x_6 = x_12; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Vector_Basic_0__Vector_mapM_go___at___00Lean_IR_StructRC_Context_accumulateRCDiff_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l___private_Init_Data_Vector_Basic_0__Vector_mapM_go___at___00Lean_IR_StructRC_Context_accumulateRCDiff_spec__0___redArg(x_3, x_4, x_5, x_6, x_8, x_9); +return x_10; +} +} +static lean_object* _init_l_Lean_IR_StructRC_Context_accumulateRCDiff___closed__0() { +_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_IR_StructRC_Context_accumulateRCDiff(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_4, 0); +x_6 = l_Std_DTreeMap_Internal_Impl_Const_get_x3f___at___00Lean_IR_StructRC_Context_renameVar_x21_spec__0___redArg(x_5, x_1); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_7; +x_7 = l_Lean_IR_StructRC_Context_emitRCDiff(x_1, x_2, x_3, x_4); +lean_dec(x_3); +return x_7; +} +else +{ +lean_object* x_8; +x_8 = lean_ctor_get(x_6, 0); +lean_inc(x_8); +lean_dec_ref(x_6); +switch (lean_obj_tag(x_8)) { +case 0: +{ +lean_dec(x_3); +lean_dec(x_1); +return x_4; +} +case 1: +{ +uint8_t x_9; +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_ctor_get(x_8, 1); +x_11 = lean_int_add(x_3, x_10); +lean_dec(x_10); +lean_dec(x_3); +lean_ctor_set(x_8, 1, x_11); +x_12 = l_Lean_IR_StructRC_Context_insert(x_1, x_8, x_4); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_13 = lean_ctor_get(x_8, 0); +x_14 = lean_ctor_get(x_8, 1); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_8); +x_15 = lean_int_add(x_3, x_14); +lean_dec(x_14); +lean_dec(x_3); +x_16 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_13); +lean_ctor_set(x_16, 1, x_15); +x_17 = l_Lean_IR_StructRC_Context_insert(x_1, x_16, x_4); +return x_17; +} +} +case 2: +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_8); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_19 = lean_ctor_get(x_8, 1); +x_20 = lean_ctor_get(x_8, 2); +x_21 = lean_alloc_closure((void*)(l_Lean_IR_StructRC_Context_accumulateRCDiff___lam__0), 3, 1); +lean_closure_set(x_21, 0, x_3); +x_22 = lean_array_get_size(x_19); +x_23 = lean_unsigned_to_nat(0u); +x_24 = l_Lean_IR_StructRC_Context_accumulateRCDiff___closed__0; +x_25 = l___private_Init_Data_Vector_Basic_0__Vector_mapM_go___at___00Lean_IR_StructRC_Context_accumulateRCDiff_spec__0___redArg(x_22, x_21, x_20, x_23, x_24, x_4); +lean_dec_ref(x_20); +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +lean_dec_ref(x_25); +lean_ctor_set(x_8, 2, x_26); +x_28 = l_Lean_IR_StructRC_Context_insert(x_1, x_8, x_27); +return x_28; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_29 = lean_ctor_get(x_8, 0); +x_30 = lean_ctor_get(x_8, 1); +x_31 = lean_ctor_get(x_8, 2); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_8); +x_32 = lean_alloc_closure((void*)(l_Lean_IR_StructRC_Context_accumulateRCDiff___lam__0), 3, 1); +lean_closure_set(x_32, 0, x_3); +x_33 = lean_array_get_size(x_30); +x_34 = lean_unsigned_to_nat(0u); +x_35 = l_Lean_IR_StructRC_Context_accumulateRCDiff___closed__0; +x_36 = l___private_Init_Data_Vector_Basic_0__Vector_mapM_go___at___00Lean_IR_StructRC_Context_accumulateRCDiff_spec__0___redArg(x_33, x_32, x_31, x_34, x_35, x_4); +lean_dec_ref(x_31); +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +lean_dec_ref(x_36); +x_39 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_39, 0, x_29); +lean_ctor_set(x_39, 1, x_30); +lean_ctor_set(x_39, 2, x_37); +x_40 = l_Lean_IR_StructRC_Context_insert(x_1, x_39, x_38); +return x_40; +} +} +default: +{ +uint8_t x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_41 = lean_ctor_get_uint8(x_8, sizeof(void*)*1); +x_42 = lean_ctor_get(x_8, 0); +lean_inc(x_42); +if (lean_is_exclusive(x_8)) { + lean_ctor_release(x_8, 0); + x_43 = x_8; +} else { + lean_dec_ref(x_8); + x_43 = lean_box(0); +} +if (x_2 == 0) +{ +x_44 = x_2; +goto block_48; +} +else +{ +x_44 = x_41; +goto block_48; +} +block_48: +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_int_add(x_3, x_42); +lean_dec(x_42); +lean_dec(x_3); +if (lean_is_scalar(x_43)) { + x_46 = lean_alloc_ctor(3, 1, 1); +} else { + x_46 = x_43; +} +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set_uint8(x_46, sizeof(void*)*1, x_44); +x_47 = l_Lean_IR_StructRC_Context_insert(x_1, x_46, x_4); +return x_47; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_accumulateRCDiff___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +switch (lean_obj_tag(x_2)) { +case 0: +{ +lean_object* x_4; +lean_dec(x_1); +x_4 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_4, 0, x_2); +lean_ctor_set(x_4, 1, x_3); +return x_4; +} +case 1: +{ +lean_object* x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; +x_5 = lean_ctor_get(x_2, 0); +x_6 = 1; +lean_inc(x_5); +x_7 = l_Lean_IR_StructRC_Context_accumulateRCDiff(x_5, x_6, x_1, x_3); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_2); +lean_ctor_set(x_8, 1, x_7); +return x_8; +} +default: +{ +uint8_t x_9; +x_9 = !lean_is_exclusive(x_2); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_ctor_get(x_2, 0); +x_11 = lean_int_add(x_1, x_10); +lean_dec(x_10); +lean_dec(x_1); +lean_ctor_set(x_2, 0, x_11); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_2); +lean_ctor_set(x_12, 1, x_3); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_2, 0); +lean_inc(x_13); +lean_dec(x_2); +x_14 = lean_int_add(x_1, x_13); +lean_dec(x_13); +lean_dec(x_1); +x_15 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_15, 0, x_14); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_3); +return x_16; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Vector_Basic_0__Vector_mapM_go___at___00Lean_IR_StructRC_Context_accumulateRCDiff_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l___private_Init_Data_Vector_Basic_0__Vector_mapM_go___at___00Lean_IR_StructRC_Context_accumulateRCDiff_spec__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec_ref(x_5); +lean_dec(x_3); +return x_10; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Vector_Basic_0__Vector_mapM_go___at___00Lean_IR_StructRC_Context_accumulateRCDiff_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l___private_Init_Data_Vector_Basic_0__Vector_mapM_go___at___00Lean_IR_StructRC_Context_accumulateRCDiff_spec__0___redArg(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec_ref(x_3); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_accumulateRCDiff___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; lean_object* x_6; +x_5 = lean_unbox(x_2); +x_6 = l_Lean_IR_StructRC_Context_accumulateRCDiff(x_1, x_5, x_3, x_4); +return x_6; +} +} +LEAN_EXPORT uint8_t l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_StructRC_Context_useVar_spec__0(lean_object* x_1, size_t x_2, size_t x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_eq(x_2, x_3); +if (x_4 == 0) +{ +uint8_t x_5; lean_object* x_6; uint8_t x_7; +x_5 = 1; +x_6 = lean_array_uget(x_1, x_2); +x_7 = l_Lean_IR_IRType_isDefiniteRef(x_6); +lean_dec(x_6); +if (x_7 == 0) +{ +return x_5; +} +else +{ +if (x_4 == 0) +{ +size_t x_8; size_t x_9; +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +x_2 = x_9; +goto _start; +} +else +{ +return x_5; +} +} +} +else +{ +uint8_t x_11; +x_11 = 0; +return x_11; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_useVar(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_1, 0); +x_6 = l_Std_DTreeMap_Internal_Impl_Const_get_x3f___at___00Lean_IR_StructRC_Context_renameVar_x21_spec__0___redArg(x_5, x_2); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_7; +lean_dec(x_2); +x_7 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_7, 0, x_1); +lean_ctor_set(x_7, 1, x_4); +return x_7; +} +else +{ +lean_object* x_8; +x_8 = lean_ctor_get(x_6, 0); +lean_inc(x_8); +lean_dec_ref(x_6); +switch (lean_obj_tag(x_8)) { +case 0: +{ +lean_object* x_9; +lean_dec(x_2); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_1); +lean_ctor_set(x_9, 1, x_4); +return x_9; +} +case 1: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_10 = lean_ctor_get(x_8, 0); +lean_inc_ref(x_10); +x_11 = lean_ctor_get(x_8, 1); +lean_inc(x_11); +if (lean_is_exclusive(x_8)) { + lean_ctor_release(x_8, 0); + lean_ctor_release(x_8, 1); + x_12 = x_8; +} else { + lean_dec_ref(x_8); + x_12 = lean_box(0); +} +if (x_3 == 0) +{ +goto block_27; +} +else +{ +lean_object* x_28; uint8_t x_29; +x_28 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__10; +x_29 = lean_int_dec_le(x_28, x_11); +if (x_29 == 0) +{ +goto block_27; +} +else +{ +lean_object* x_30; +lean_dec(x_12); +lean_dec(x_11); +lean_dec_ref(x_10); +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_4); +return x_30; +} +} +block_20: +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_inc(x_2); +x_15 = l_Lean_IR_StructRC_Context_emitRCDiff(x_2, x_14, x_11, x_1); +lean_dec(x_11); +x_16 = lean_nat_to_int(x_13); +if (lean_is_scalar(x_12)) { + x_17 = lean_alloc_ctor(1, 2, 0); +} else { + x_17 = x_12; +} +lean_ctor_set(x_17, 0, x_10); +lean_ctor_set(x_17, 1, x_16); +x_18 = l_Lean_IR_StructRC_Context_insert(x_2, x_17, x_15); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_4); +return x_19; +} +block_27: +{ +lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_21 = lean_unsigned_to_nat(0u); +x_22 = lean_array_get_size(x_10); +x_23 = lean_nat_dec_lt(x_21, x_22); +if (x_23 == 0) +{ +x_13 = x_21; +x_14 = x_23; +goto block_20; +} +else +{ +if (x_23 == 0) +{ +x_13 = x_21; +x_14 = x_23; +goto block_20; +} +else +{ +size_t x_24; size_t x_25; uint8_t x_26; +x_24 = 0; +x_25 = lean_usize_of_nat(x_22); +x_26 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_StructRC_Context_useVar_spec__0(x_10, x_24, x_25); +x_13 = x_21; +x_14 = x_26; +goto block_20; +} +} +} +} +case 2: +{ +uint8_t x_31; +x_31 = !lean_is_exclusive(x_8); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; +x_32 = lean_ctor_get(x_8, 0); +x_33 = lean_ctor_get(x_8, 1); +x_34 = lean_ctor_get(x_8, 2); +x_35 = lean_array_get_size(x_33); +x_36 = lean_unsigned_to_nat(0u); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_1); +lean_ctor_set(x_37, 1, x_34); +lean_inc(x_2); +lean_inc(x_32); +x_38 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_StructRC_Context_useVar_spec__1___redArg(x_35, x_3, x_33, x_32, x_2, x_36, x_37, x_4); +x_39 = !lean_is_exclusive(x_38); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_40 = lean_ctor_get(x_38, 0); +x_41 = lean_ctor_get(x_40, 0); +lean_inc(x_41); +x_42 = lean_ctor_get(x_40, 1); +lean_inc(x_42); +lean_dec(x_40); +lean_ctor_set(x_8, 2, x_42); +x_43 = l_Lean_IR_StructRC_Context_insert(x_2, x_8, x_41); +lean_ctor_set(x_38, 0, x_43); +return x_38; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_38, 0); +x_45 = lean_ctor_get(x_38, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_38); +x_46 = lean_ctor_get(x_44, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_44, 1); +lean_inc(x_47); +lean_dec(x_44); +lean_ctor_set(x_8, 2, x_47); +x_48 = l_Lean_IR_StructRC_Context_insert(x_2, x_8, 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_45); +return x_49; +} +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_50 = lean_ctor_get(x_8, 0); +x_51 = lean_ctor_get(x_8, 1); +x_52 = lean_ctor_get(x_8, 2); +lean_inc(x_52); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_8); +x_53 = lean_array_get_size(x_51); +x_54 = lean_unsigned_to_nat(0u); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_1); +lean_ctor_set(x_55, 1, x_52); +lean_inc(x_2); +lean_inc(x_50); +x_56 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_StructRC_Context_useVar_spec__1___redArg(x_53, x_3, x_51, x_50, x_2, x_54, x_55, x_4); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_56, 1); +lean_inc(x_58); +if (lean_is_exclusive(x_56)) { + lean_ctor_release(x_56, 0); + lean_ctor_release(x_56, 1); + x_59 = x_56; +} else { + lean_dec_ref(x_56); + x_59 = lean_box(0); +} +x_60 = lean_ctor_get(x_57, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_57, 1); +lean_inc(x_61); +lean_dec(x_57); +x_62 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_62, 0, x_50); +lean_ctor_set(x_62, 1, x_51); +lean_ctor_set(x_62, 2, x_61); +x_63 = l_Lean_IR_StructRC_Context_insert(x_2, x_62, x_60); +if (lean_is_scalar(x_59)) { + x_64 = lean_alloc_ctor(0, 2, 0); +} else { + x_64 = x_59; +} +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_58); +return x_64; +} +} +default: +{ +uint8_t x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get_uint8(x_8, sizeof(void*)*1); +x_66 = lean_ctor_get(x_8, 0); +lean_inc(x_66); +if (lean_is_exclusive(x_8)) { + lean_ctor_release(x_8, 0); + x_67 = x_8; +} else { + lean_dec_ref(x_8); + x_67 = lean_box(0); +} +if (x_3 == 0) +{ +goto block_73; +} +else +{ +lean_object* x_74; uint8_t x_75; +x_74 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__10; +x_75 = lean_int_dec_le(x_74, x_66); +if (x_75 == 0) +{ +goto block_73; +} +else +{ +lean_object* x_76; +lean_dec(x_67); +lean_dec(x_66); +lean_dec(x_2); +x_76 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_76, 0, x_1); +lean_ctor_set(x_76, 1, x_4); +return x_76; +} +} +block_73: +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +lean_inc(x_2); +x_68 = l_Lean_IR_StructRC_Context_emitRCDiff(x_2, x_65, x_66, x_1); +lean_dec(x_66); +x_69 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__10; +if (lean_is_scalar(x_67)) { + x_70 = lean_alloc_ctor(3, 1, 1); +} else { + x_70 = x_67; +} +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set_uint8(x_70, sizeof(void*)*1, x_65); +x_71 = l_Lean_IR_StructRC_Context_insert(x_2, x_70, x_68); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set(x_72, 1, x_4); +return x_72; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_StructRC_Context_useVar_spec__1___redArg(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; uint8_t x_15; +x_15 = lean_nat_dec_lt(x_6, x_1); +if (x_15 == 0) +{ +lean_object* x_16; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_7); +lean_ctor_set(x_16, 1, x_8); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_22; +x_17 = lean_ctor_get(x_7, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_7, 1); +lean_inc(x_18); +if (lean_is_exclusive(x_7)) { + lean_ctor_release(x_7, 0); + lean_ctor_release(x_7, 1); + x_19 = x_7; +} else { + lean_dec_ref(x_7); + x_19 = lean_box(0); +} +x_22 = lean_array_fget(x_18, x_6); +switch (lean_obj_tag(x_22)) { +case 0: +{ +lean_object* x_23; +lean_dec(x_19); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_17); +lean_ctor_set(x_23, 1, x_18); +x_9 = x_23; +x_10 = x_8; +goto block_14; +} +case 1: +{ +lean_object* x_24; lean_object* x_25; uint8_t x_26; +lean_dec(x_19); +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec_ref(x_22); +x_25 = l_Lean_IR_StructRC_Context_useVar(x_17, x_24, x_2, x_8); +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) +{ +lean_object* x_27; +x_27 = lean_ctor_get(x_25, 1); +lean_ctor_set(x_25, 1, x_18); +x_9 = x_25; +x_10 = x_27; +goto block_14; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_25, 0); +x_29 = lean_ctor_get(x_25, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_25); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_18); +x_9 = x_30; +x_10 = x_29; +goto block_14; +} +} +default: +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; +x_31 = lean_ctor_get(x_22, 0); +lean_inc(x_31); +if (lean_is_exclusive(x_22)) { + lean_ctor_release(x_22, 0); + x_32 = x_22; +} else { + lean_dec_ref(x_22); + x_32 = lean_box(0); +} +x_33 = lean_array_fget_borrowed(x_3, x_6); +if (x_2 == 0) +{ +goto block_56; +} +else +{ +lean_object* x_57; uint8_t x_58; +x_57 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__10; +x_58 = lean_int_dec_le(x_57, x_31); +if (x_58 == 0) +{ +goto block_56; +} +else +{ +lean_object* x_59; +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_19); +x_59 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_59, 0, x_17); +lean_ctor_set(x_59, 1, x_18); +x_9 = x_59; +x_10 = x_8; +goto block_14; +} +} +block_44: +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +lean_inc(x_36); +x_39 = l_Lean_IR_StructRC_Context_emitRCDiff(x_36, x_38, x_31, x_35); +lean_dec(x_31); +lean_inc(x_33); +lean_inc(x_36); +x_40 = l_Lean_IR_StructRC_Context_addVar(x_39, x_36, x_33, x_37); +if (lean_is_scalar(x_32)) { + x_41 = lean_alloc_ctor(1, 1, 0); +} else { + x_41 = x_32; + lean_ctor_set_tag(x_41, 1); +} +lean_ctor_set(x_41, 0, x_36); +x_42 = lean_array_fset(x_18, x_6, x_41); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_40); +lean_ctor_set(x_43, 1, x_42); +x_9 = x_43; +x_10 = x_34; +goto block_14; +} +block_56: +{ +lean_object* x_45; uint8_t x_46; +x_45 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__10; +x_46 = lean_int_dec_eq(x_31, x_45); +if (x_46 == 0) +{ +if (x_15 == 0) +{ +lean_dec(x_32); +lean_dec(x_31); +goto block_21; +} +else +{ +uint8_t x_47; +x_47 = l_Lean_IR_StructRC_needsRC(x_33); +if (x_47 == 0) +{ +lean_dec(x_32); +lean_dec(x_31); +goto block_21; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +lean_dec(x_19); +x_48 = l_Lean_IR_StructRC_mkFreshVar(x_8); +x_49 = lean_ctor_get(x_48, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_48, 1); +lean_inc(x_50); +lean_dec_ref(x_48); +lean_inc(x_5); +lean_inc(x_6); +lean_inc(x_4); +x_51 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_51, 0, x_4); +lean_ctor_set(x_51, 1, x_6); +lean_ctor_set(x_51, 2, x_5); +x_52 = lean_box(12); +lean_inc(x_33); +lean_inc(x_49); +x_53 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_53, 0, x_49); +lean_ctor_set(x_53, 1, x_33); +lean_ctor_set(x_53, 2, x_51); +lean_ctor_set(x_53, 3, x_52); +x_54 = l_Lean_IR_StructRC_Context_addInstr(x_17, x_53); +x_55 = l_Lean_IR_IRType_isDefiniteRef(x_33); +if (x_55 == 0) +{ +x_34 = x_50; +x_35 = x_54; +x_36 = x_49; +x_37 = x_45; +x_38 = x_47; +goto block_44; +} +else +{ +x_34 = x_50; +x_35 = x_54; +x_36 = x_49; +x_37 = x_45; +x_38 = x_46; +goto block_44; +} +} +} +} +else +{ +lean_dec(x_32); +lean_dec(x_31); +goto block_21; +} +} +} +} +block_21: +{ +lean_object* x_20; +if (lean_is_scalar(x_19)) { + x_20 = lean_alloc_ctor(0, 2, 0); +} else { + x_20 = x_19; +} +lean_ctor_set(x_20, 0, x_17); +lean_ctor_set(x_20, 1, x_18); +x_9 = x_20; +x_10 = x_8; +goto block_14; +} +} +block_14: +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_add(x_6, x_11); +lean_dec(x_6); +x_6 = x_12; +x_7 = x_9; +x_8 = x_10; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_StructRC_Context_useVar_spec__1(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_StructRC_Context_useVar_spec__1___redArg(x_1, x_2, x_3, x_4, x_5, x_9, x_10, x_12); +return x_13; +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_StructRC_Context_useVar_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_2); +x_14 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_StructRC_Context_useVar_spec__1(x_1, x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_6); +lean_dec_ref(x_3); +lean_dec(x_1); +return x_14; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_StructRC_Context_useVar_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; uint8_t x_6; lean_object* x_7; +x_4 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_5 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_6 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_IR_StructRC_Context_useVar_spec__0(x_1, x_4, x_5); +lean_dec_ref(x_1); +x_7 = lean_box(x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_IR_StructRC_Context_useVar_spec__1___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_2); +x_10 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_StructRC_Context_useVar_spec__1___redArg(x_1, x_9, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec_ref(x_3); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_useVar___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; lean_object* x_6; +x_5 = lean_unbox(x_3); +x_6 = l_Lean_IR_StructRC_Context_useVar(x_1, x_2, x_5, x_4); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_useArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_4; uint8_t x_5; lean_object* x_6; +x_4 = lean_ctor_get(x_2, 0); +lean_inc(x_4); +lean_dec_ref(x_2); +x_5 = 0; +x_6 = l_Lean_IR_StructRC_Context_useVar(x_1, x_4, x_5, x_3); +return x_6; +} +else +{ +lean_object* x_7; +x_7 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_7, 0, x_1); +lean_ctor_set(x_7, 1, x_3); +return x_7; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_StructRC_Context_addParams_spec__0(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_2, x_3); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; size_t x_10; size_t x_11; +x_6 = lean_array_uget(x_1, x_2); +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +lean_dec(x_6); +x_9 = l_Lean_IR_StructRC_Context_addVarBasic(x_4, x_7, x_8); +x_10 = 1; +x_11 = lean_usize_add(x_2, x_10); +x_2 = x_11; +x_4 = x_9; +goto _start; +} +else +{ +return x_4; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_addParams(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = lean_unsigned_to_nat(0u); +x_4 = lean_array_get_size(x_1); +x_5 = lean_nat_dec_lt(x_3, x_4); +if (x_5 == 0) +{ +return x_2; +} +else +{ +uint8_t x_6; +x_6 = lean_nat_dec_le(x_4, x_4); +if (x_6 == 0) +{ +if (x_5 == 0) +{ +return x_2; +} +else +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = 0; +x_8 = lean_usize_of_nat(x_4); +x_9 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_StructRC_Context_addParams_spec__0(x_1, x_7, x_8, x_2); +return x_9; +} +} +else +{ +size_t x_10; size_t x_11; lean_object* x_12; +x_10 = 0; +x_11 = lean_usize_of_nat(x_4); +x_12 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_StructRC_Context_addParams_spec__0(x_1, x_10, x_11, x_2); +return x_12; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_StructRC_Context_addParams_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; lean_object* x_7; +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_StructRC_Context_addParams_spec__0(x_1, x_5, x_6, x_4); +lean_dec_ref(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_addParams___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_IR_StructRC_Context_addParams(x_1, x_2); +lean_dec_ref(x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_IR_StructRC_Context_resetRC___closed__0() { +_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_Std_DTreeMap_Internal_Impl_map___at___00Lean_IR_StructRC_Context_resetRC_spec__0(lean_object* x_1) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +x_5 = lean_ctor_get(x_1, 3); +lean_inc(x_5); +x_6 = lean_ctor_get(x_1, 4); +lean_inc(x_6); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + lean_ctor_release(x_1, 4); + x_7 = x_1; +} else { + lean_dec_ref(x_1); + x_7 = lean_box(0); +} +switch (lean_obj_tag(x_4)) { +case 1: +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_4); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_4, 1); +lean_dec(x_14); +x_15 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__10; +lean_ctor_set(x_4, 1, x_15); +x_8 = x_4; +goto block_12; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_4, 0); +lean_inc(x_16); +lean_dec(x_4); +x_17 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__10; +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +x_8 = x_18; +goto block_12; +} +} +case 3: +{ +uint8_t x_19; +x_19 = !lean_is_exclusive(x_4); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_4, 0); +lean_dec(x_20); +x_21 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__10; +lean_ctor_set(x_4, 0, x_21); +x_8 = x_4; +goto block_12; +} +else +{ +uint8_t x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get_uint8(x_4, sizeof(void*)*1); +lean_dec(x_4); +x_23 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__10; +x_24 = lean_alloc_ctor(3, 1, 1); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set_uint8(x_24, sizeof(void*)*1, x_22); +x_8 = x_24; +goto block_12; +} +} +default: +{ +x_8 = x_4; +goto block_12; +} +} +block_12: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = l_Std_DTreeMap_Internal_Impl_map___at___00Lean_IR_StructRC_Context_resetRC_spec__0(x_5); +x_10 = l_Std_DTreeMap_Internal_Impl_map___at___00Lean_IR_StructRC_Context_resetRC_spec__0(x_6); +if (lean_is_scalar(x_7)) { + x_11 = lean_alloc_ctor(0, 5, 0); +} else { + x_11 = x_7; +} +lean_ctor_set(x_11, 0, x_2); +lean_ctor_set(x_11, 1, x_3); +lean_ctor_set(x_11, 2, x_8); +lean_ctor_set(x_11, 3, x_9); +lean_ctor_set(x_11, 4, x_10); +return x_11; +} +} +else +{ +return x_1; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_resetRC(lean_object* x_1) { +_start: +{ +uint8_t x_2; +x_2 = !lean_is_exclusive(x_1); +if (x_2 == 0) +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_ctor_get(x_1, 1); +lean_dec(x_4); +x_5 = l_Std_DTreeMap_Internal_Impl_map___at___00Lean_IR_StructRC_Context_resetRC_spec__0(x_3); +x_6 = l_Lean_IR_StructRC_Context_resetRC___closed__0; +lean_ctor_set(x_1, 1, x_6); +lean_ctor_set(x_1, 0, x_5); +return x_1; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_7 = lean_ctor_get(x_1, 0); +x_8 = lean_ctor_get(x_1, 2); +lean_inc(x_8); +lean_inc(x_7); +lean_dec(x_1); +x_9 = l_Std_DTreeMap_Internal_Impl_map___at___00Lean_IR_StructRC_Context_resetRC_spec__0(x_7); +x_10 = l_Lean_IR_StructRC_Context_resetRC___closed__0; +x_11 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_11, 0, x_9); +lean_ctor_set(x_11, 1, x_10); +lean_ctor_set(x_11, 2, x_8); +return x_11; +} +} +} +LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_forInStep___at___00Lean_IR_StructRC_Context_finish_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_4 = lean_ctor_get(x_2, 1); +lean_inc(x_4); +x_5 = lean_ctor_get(x_2, 3); +lean_inc(x_5); +x_6 = lean_ctor_get(x_2, 4); +lean_inc(x_6); +lean_dec_ref(x_2); +x_7 = l_Std_DTreeMap_Internal_Impl_forInStep___at___00Lean_IR_StructRC_Context_finish_spec__0(x_1, x_5, x_3); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +lean_dec_ref(x_7); +x_10 = lean_ctor_get(x_8, 0); +lean_inc(x_10); +lean_dec(x_8); +x_11 = 0; +x_12 = l_Lean_IR_StructRC_Context_useVar(x_10, x_4, x_11, x_9); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec_ref(x_12); +x_1 = x_13; +x_2 = x_6; +x_3 = x_14; +goto _start; +} +else +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_16, 0, x_1); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_3); +return x_17; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_finish(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_19; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = l_Std_DTreeMap_Internal_Impl_forInStep___at___00Lean_IR_StructRC_Context_finish_spec__0(x_1, x_3, x_2); +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_4, 1); +lean_inc(x_6); +if (lean_is_exclusive(x_4)) { + lean_ctor_release(x_4, 0); + lean_ctor_release(x_4, 1); + x_7 = x_4; +} else { + lean_dec_ref(x_4); + x_7 = lean_box(0); +} +x_19 = lean_ctor_get(x_5, 0); +lean_inc(x_19); +lean_dec(x_5); +x_8 = x_19; +goto block_18; +block_18: +{ +uint8_t x_9; +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_ctor_get(x_8, 0); +lean_dec(x_10); +x_11 = lean_box(1); +lean_ctor_set(x_8, 0, x_11); +if (lean_is_scalar(x_7)) { + x_12 = lean_alloc_ctor(0, 2, 0); +} else { + x_12 = x_7; +} +lean_ctor_set(x_12, 0, x_8); +lean_ctor_set(x_12, 1, x_6); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_13 = lean_ctor_get(x_8, 1); +x_14 = lean_ctor_get(x_8, 2); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_8); +x_15 = lean_box(1); +x_16 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_13); +lean_ctor_set(x_16, 2, x_14); +if (lean_is_scalar(x_7)) { + x_17 = lean_alloc_ctor(0, 2, 0); +} else { + x_17 = x_7; +} +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_6); +return x_17; +} +} +} +} +LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_StructRC_Context_addCtorKnowledge_spec__0(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_box(0); +x_3 = lean_panic_fn(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_Context_addCtorKnowledge(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_1, 0); +x_5 = l_Std_DTreeMap_Internal_Impl_Const_get_x3f___at___00Lean_IR_StructRC_Context_renameVar_x21_spec__0___redArg(x_4, x_2); +if (lean_obj_tag(x_5) == 1) +{ +uint8_t x_6; +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; +x_7 = lean_ctor_get(x_5, 0); +if (lean_obj_tag(x_7) == 1) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_7, 0); +lean_inc_ref(x_8); +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +lean_dec_ref(x_7); +x_10 = lean_box(0); +x_11 = lean_array_get(x_10, x_8, x_3); +lean_dec_ref(x_8); +if (lean_obj_tag(x_11) == 10) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = lean_ctor_get(x_11, 1); +lean_inc_ref(x_12); +lean_dec_ref(x_11); +x_13 = lean_array_get_size(x_12); +lean_ctor_set_tag(x_5, 2); +lean_ctor_set(x_5, 0, x_9); +x_14 = lean_mk_array(x_13, x_5); +x_15 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_15, 0, x_3); +lean_ctor_set(x_15, 1, x_12); +lean_ctor_set(x_15, 2, x_14); +x_16 = l_Lean_IR_StructRC_Context_insert(x_2, x_15, x_1); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_dec(x_11); +lean_dec(x_9); +lean_free_object(x_5); +lean_dec(x_3); +x_17 = l_Lean_IR_StructRC_Entry_ofStruct___closed__2; +x_18 = l_panic___at___00Lean_IR_StructRC_Context_addCtorKnowledge_spec__0(x_17); +x_19 = l_Lean_IR_StructRC_Context_insert(x_2, x_18, x_1); +return x_19; +} +} +else +{ +lean_free_object(x_5); +lean_dec(x_7); +lean_dec(x_3); +lean_dec(x_2); +return x_1; +} +} +else +{ +lean_object* x_20; +x_20 = lean_ctor_get(x_5, 0); +lean_inc(x_20); +lean_dec(x_5); +if (lean_obj_tag(x_20) == 1) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_21 = lean_ctor_get(x_20, 0); +lean_inc_ref(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec_ref(x_20); +x_23 = lean_box(0); +x_24 = lean_array_get(x_23, x_21, x_3); +lean_dec_ref(x_21); +if (lean_obj_tag(x_24) == 10) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_25 = lean_ctor_get(x_24, 1); +lean_inc_ref(x_25); +lean_dec_ref(x_24); +x_26 = lean_array_get_size(x_25); +x_27 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_27, 0, x_22); +x_28 = lean_mk_array(x_26, x_27); +x_29 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_29, 0, x_3); +lean_ctor_set(x_29, 1, x_25); +lean_ctor_set(x_29, 2, x_28); +x_30 = l_Lean_IR_StructRC_Context_insert(x_2, x_29, x_1); +return x_30; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +lean_dec(x_24); +lean_dec(x_22); +lean_dec(x_3); +x_31 = l_Lean_IR_StructRC_Entry_ofStruct___closed__2; +x_32 = l_panic___at___00Lean_IR_StructRC_Context_addCtorKnowledge_spec__0(x_31); +x_33 = l_Lean_IR_StructRC_Context_insert(x_2, x_32, x_1); +return x_33; +} +} +else +{ +lean_dec(x_20); +lean_dec(x_3); +lean_dec(x_2); +return x_1; +} +} +} +else +{ +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +return x_1; +} +} +} +static lean_object* _init_l_panic___at___00Lean_IR_StructRC_atConstructorIndex_spec__0___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = l_Array_empty(lean_box(0)); +return x_1; +} +} +LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_StructRC_atConstructorIndex_spec__0(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_panic___at___00Lean_IR_StructRC_atConstructorIndex_spec__0___closed__0; +x_3 = lean_panic_fn(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_IR_StructRC_atConstructorIndex___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.IR.StructRC.atConstructorIndex", 35, 35); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_StructRC_atConstructorIndex___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_IR_StructRC_Entry_ofStruct___closed__1; +x_2 = lean_unsigned_to_nat(11u); +x_3 = lean_unsigned_to_nat(281u); +x_4 = l_Lean_IR_StructRC_atConstructorIndex___closed__0; +x_5 = l_Lean_IR_StructRC_Context_renameVar_x21___closed__0; +x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); +return x_6; +} +} +static lean_object* _init_l_Lean_IR_StructRC_atConstructorIndex___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_IR_StructRC_Entry_ofStruct___closed__1; +x_2 = lean_unsigned_to_nat(9u); +x_3 = lean_unsigned_to_nat(282u); +x_4 = l_Lean_IR_StructRC_atConstructorIndex___closed__0; +x_5 = l_Lean_IR_StructRC_Context_renameVar_x21___closed__0; +x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_atConstructorIndex(lean_object* x_1, lean_object* x_2) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 10: +{ +lean_object* x_3; +x_3 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_3); +return x_3; +} +case 11: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_4 = lean_ctor_get(x_1, 1); +x_5 = lean_box(0); +x_6 = lean_array_get_borrowed(x_5, x_4, x_2); +if (lean_obj_tag(x_6) == 10) +{ +lean_object* x_7; +x_7 = lean_ctor_get(x_6, 1); +lean_inc_ref(x_7); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; +x_8 = l_Lean_IR_StructRC_atConstructorIndex___closed__1; +x_9 = l_panic___at___00Lean_IR_StructRC_atConstructorIndex_spec__0(x_8); +return x_9; +} +} +default: +{ +lean_object* x_10; lean_object* x_11; +x_10 = l_Lean_IR_StructRC_atConstructorIndex___closed__2; +x_11 = l_panic___at___00Lean_IR_StructRC_atConstructorIndex_spec__0(x_10); +return x_11; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_atConstructorIndex___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_IR_StructRC_atConstructorIndex(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_IR_StructRC_visitExpr___lam__0___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_IR_instInhabitedArg_default; +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00Lean_IR_StructRC_visitExpr_spec__1___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_nat_dec_eq(x_4, x_6); +if (x_7 == 1) +{ +lean_dec(x_4); +return x_5; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_8 = lean_unsigned_to_nat(1u); +x_9 = lean_nat_sub(x_4, x_8); +x_10 = lean_nat_sub(x_3, x_4); +lean_dec(x_4); +x_11 = l_Lean_IR_StructRC_visitExpr___lam__0___closed__0; +x_12 = lean_array_get_borrowed(x_11, x_1, x_10); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_12, 0); +x_14 = lean_array_fget_borrowed(x_2, x_10); +lean_dec(x_10); +x_15 = l_Lean_IR_StructRC_instReprProjEntry_repr___closed__10; +lean_inc(x_14); +lean_inc(x_13); +x_16 = l_Lean_IR_StructRC_Context_addVar(x_5, x_13, x_14, x_15); +x_4 = x_9; +x_5 = x_16; +goto _start; +} +else +{ +lean_dec(x_10); +x_4 = x_9; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00Lean_IR_StructRC_visitExpr_spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00Lean_IR_StructRC_visitExpr_spec__1___redArg(x_1, x_2, x_3, x_4, x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_visitExpr___lam__0(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = l_Lean_IR_StructRC_visitExpr___lam__0___closed__0; +x_4 = lean_array_get(x_3, x_1, x_2); +if (lean_obj_tag(x_4) == 0) +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) +{ +lean_ctor_set_tag(x_4, 1); +return x_4; +} +else +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_4, 0); +lean_inc(x_6); +lean_dec(x_4); +x_7 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_7, 0, x_6); +return x_7; +} +} +else +{ +lean_object* x_8; +x_8 = lean_box(0); +return x_8; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_StructRC_visitExpr_spec__0(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_2, x_3); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; size_t x_11; size_t x_12; +x_7 = lean_array_uget(x_1, x_2); +x_8 = l_Lean_IR_StructRC_Context_useArg(x_4, x_7, x_5); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +lean_dec_ref(x_8); +x_11 = 1; +x_12 = lean_usize_add(x_2, x_11); +x_2 = x_12; +x_4 = x_9; +x_5 = x_10; +goto _start; +} +else +{ +lean_object* x_14; +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_4); +lean_ctor_set(x_14, 1, x_5); +return x_14; +} +} +} +LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_StructRC_visitExpr_spec__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Lean_IR_StructRC_Context_addProjInfo___closed__9; +x_4 = l_instInhabitedOfMonad___redArg(x_3, x_1); +x_5 = lean_panic_fn(x_4, x_2); +return x_5; +} +} +static lean_object* _init_l_Lean_IR_StructRC_visitExpr___closed__0() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_IR_StructRC_visitExpr___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_IR_StructRC_visitExpr___closed__0; +x_2 = lean_alloc_ctor(11, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_visitExpr___lam__0___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_IR_StructRC_visitExpr___lam__0(x_1, x_2); +lean_dec(x_2); +lean_dec_ref(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_visitExpr(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +switch (lean_obj_tag(x_3)) { +case 0: +{ +uint8_t x_6; +x_6 = !lean_is_exclusive(x_3); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_17; uint8_t x_21; +x_7 = lean_ctor_get(x_3, 0); +x_8 = lean_ctor_get(x_3, 1); +x_9 = l_Lean_IR_StructRC_Context_renameArgs(x_4, x_8); +lean_inc_ref(x_9); +lean_inc_ref(x_7); +lean_ctor_set(x_3, 1, x_9); +x_21 = l_Lean_IR_IRType_isStruct(x_2); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; uint8_t x_24; +lean_dec_ref(x_7); +x_22 = lean_unsigned_to_nat(0u); +x_23 = lean_array_get_size(x_9); +x_24 = lean_nat_dec_lt(x_22, x_23); +if (x_24 == 0) +{ +lean_dec_ref(x_9); +x_10 = x_4; +x_11 = x_5; +goto block_16; +} +else +{ +uint8_t x_25; +x_25 = lean_nat_dec_le(x_23, x_23); +if (x_25 == 0) +{ +if (x_24 == 0) +{ +lean_dec_ref(x_9); +x_10 = x_4; +x_11 = x_5; +goto block_16; +} +else +{ +size_t x_26; size_t x_27; lean_object* x_28; +x_26 = 0; +x_27 = lean_usize_of_nat(x_23); +x_28 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_StructRC_visitExpr_spec__0(x_9, x_26, x_27, x_4, x_5); +lean_dec_ref(x_9); +x_17 = x_28; +goto block_20; +} +} +else +{ +size_t x_29; size_t x_30; lean_object* x_31; +x_29 = 0; +x_30 = lean_usize_of_nat(x_23); +x_31 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_StructRC_visitExpr_spec__0(x_9, x_29, x_30, x_4, x_5); +lean_dec_ref(x_9); +x_17 = x_31; +goto block_20; +} +} +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_32 = lean_ctor_get(x_7, 1); +lean_inc(x_32); +lean_dec_ref(x_7); +lean_inc_ref(x_9); +x_33 = lean_alloc_closure((void*)(l_Lean_IR_StructRC_visitExpr___lam__0___boxed), 2, 1); +lean_closure_set(x_33, 0, x_9); +x_34 = l_Lean_IR_StructRC_atConstructorIndex(x_2, x_32); +x_35 = lean_array_get_size(x_34); +x_36 = l_Array_ofFn___redArg(x_35, x_33); +lean_inc_ref(x_34); +x_37 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_37, 0, x_32); +lean_ctor_set(x_37, 1, x_34); +lean_ctor_set(x_37, 2, x_36); +x_38 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00Lean_IR_StructRC_visitExpr_spec__1___redArg(x_9, x_34, x_35, x_35, x_4); +lean_dec_ref(x_34); +lean_dec_ref(x_9); +lean_inc(x_1); +x_39 = l_Lean_IR_StructRC_Context_insert(x_1, x_37, x_38); +x_40 = lean_box(12); +x_41 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_41, 0, x_1); +lean_ctor_set(x_41, 1, x_2); +lean_ctor_set(x_41, 2, x_3); +lean_ctor_set(x_41, 3, x_40); +x_42 = l_Lean_IR_StructRC_Context_addInstr(x_39, x_41); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_5); +return x_43; +} +block_16: +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_box(12); +x_13 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_13, 0, x_1); +lean_ctor_set(x_13, 1, x_2); +lean_ctor_set(x_13, 2, x_3); +lean_ctor_set(x_13, 3, x_12); +x_14 = l_Lean_IR_StructRC_Context_addInstr(x_10, x_13); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_11); +return x_15; +} +block_20: +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec_ref(x_17); +x_10 = x_18; +x_11 = x_19; +goto block_16; +} +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_55; uint8_t x_59; +x_44 = lean_ctor_get(x_3, 0); +x_45 = lean_ctor_get(x_3, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_3); +x_46 = l_Lean_IR_StructRC_Context_renameArgs(x_4, x_45); +lean_inc_ref(x_46); +lean_inc_ref(x_44); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_44); +lean_ctor_set(x_47, 1, x_46); +x_59 = l_Lean_IR_IRType_isStruct(x_2); +if (x_59 == 0) +{ +lean_object* x_60; lean_object* x_61; uint8_t x_62; +lean_dec_ref(x_44); +x_60 = lean_unsigned_to_nat(0u); +x_61 = lean_array_get_size(x_46); +x_62 = lean_nat_dec_lt(x_60, x_61); +if (x_62 == 0) +{ +lean_dec_ref(x_46); +x_48 = x_4; +x_49 = x_5; +goto block_54; +} +else +{ +uint8_t x_63; +x_63 = lean_nat_dec_le(x_61, x_61); +if (x_63 == 0) +{ +if (x_62 == 0) +{ +lean_dec_ref(x_46); +x_48 = x_4; +x_49 = x_5; +goto block_54; +} +else +{ +size_t x_64; size_t x_65; lean_object* x_66; +x_64 = 0; +x_65 = lean_usize_of_nat(x_61); +x_66 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_StructRC_visitExpr_spec__0(x_46, x_64, x_65, x_4, x_5); +lean_dec_ref(x_46); +x_55 = x_66; +goto block_58; +} +} +else +{ +size_t x_67; size_t x_68; lean_object* x_69; +x_67 = 0; +x_68 = lean_usize_of_nat(x_61); +x_69 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_StructRC_visitExpr_spec__0(x_46, x_67, x_68, x_4, x_5); +lean_dec_ref(x_46); +x_55 = x_69; +goto block_58; +} +} +} +else +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_70 = lean_ctor_get(x_44, 1); +lean_inc(x_70); +lean_dec_ref(x_44); +lean_inc_ref(x_46); +x_71 = lean_alloc_closure((void*)(l_Lean_IR_StructRC_visitExpr___lam__0___boxed), 2, 1); +lean_closure_set(x_71, 0, x_46); +x_72 = l_Lean_IR_StructRC_atConstructorIndex(x_2, x_70); +x_73 = lean_array_get_size(x_72); +x_74 = l_Array_ofFn___redArg(x_73, x_71); +lean_inc_ref(x_72); +x_75 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_75, 0, x_70); +lean_ctor_set(x_75, 1, x_72); +lean_ctor_set(x_75, 2, x_74); +x_76 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00Lean_IR_StructRC_visitExpr_spec__1___redArg(x_46, x_72, x_73, x_73, x_4); +lean_dec_ref(x_72); +lean_dec_ref(x_46); +lean_inc(x_1); +x_77 = l_Lean_IR_StructRC_Context_insert(x_1, x_75, x_76); +x_78 = lean_box(12); +x_79 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_79, 0, x_1); +lean_ctor_set(x_79, 1, x_2); +lean_ctor_set(x_79, 2, x_47); +lean_ctor_set(x_79, 3, x_78); +x_80 = l_Lean_IR_StructRC_Context_addInstr(x_77, x_79); +x_81 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_5); +return x_81; +} +block_54: +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_50 = lean_box(12); +x_51 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_51, 0, x_1); +lean_ctor_set(x_51, 1, x_2); +lean_ctor_set(x_51, 2, x_47); +lean_ctor_set(x_51, 3, x_50); +x_52 = l_Lean_IR_StructRC_Context_addInstr(x_48, x_51); +x_53 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_49); +return x_53; +} +block_58: +{ +lean_object* x_56; lean_object* x_57; +x_56 = lean_ctor_get(x_55, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_55, 1); +lean_inc(x_57); +lean_dec_ref(x_55); +x_48 = x_56; +x_49 = x_57; +goto block_54; +} +} +} +case 1: +{ +uint8_t x_82; +x_82 = !lean_is_exclusive(x_3); +if (x_82 == 0) +{ +lean_object* x_83; lean_object* x_84; uint8_t x_85; lean_object* x_86; uint8_t x_87; +x_83 = lean_ctor_get(x_3, 1); +x_84 = l_Lean_IR_StructRC_Context_renameVar_x21(x_4, x_83); +lean_dec(x_83); +x_85 = 0; +lean_inc(x_84); +x_86 = l_Lean_IR_StructRC_Context_useVar(x_4, x_84, x_85, x_5); +x_87 = !lean_is_exclusive(x_86); +if (x_87 == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_88 = lean_ctor_get(x_86, 0); +lean_ctor_set(x_3, 1, x_84); +x_89 = lean_box(12); +x_90 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_90, 0, x_1); +lean_ctor_set(x_90, 1, x_2); +lean_ctor_set(x_90, 2, x_3); +lean_ctor_set(x_90, 3, x_89); +x_91 = l_Lean_IR_StructRC_Context_addInstr(x_88, x_90); +lean_ctor_set(x_86, 0, x_91); +return x_86; +} +else +{ +lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; +x_92 = lean_ctor_get(x_86, 0); +x_93 = lean_ctor_get(x_86, 1); +lean_inc(x_93); +lean_inc(x_92); +lean_dec(x_86); +lean_ctor_set(x_3, 1, x_84); +x_94 = lean_box(12); +x_95 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_95, 0, x_1); +lean_ctor_set(x_95, 1, x_2); +lean_ctor_set(x_95, 2, x_3); +lean_ctor_set(x_95, 3, x_94); +x_96 = l_Lean_IR_StructRC_Context_addInstr(x_92, x_95); +x_97 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_97, 0, x_96); +lean_ctor_set(x_97, 1, x_93); +return x_97; +} +} +else +{ +lean_object* x_98; lean_object* x_99; lean_object* x_100; uint8_t x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; +x_98 = lean_ctor_get(x_3, 0); +x_99 = lean_ctor_get(x_3, 1); +lean_inc(x_99); +lean_inc(x_98); +lean_dec(x_3); +x_100 = l_Lean_IR_StructRC_Context_renameVar_x21(x_4, x_99); +lean_dec(x_99); +x_101 = 0; +lean_inc(x_100); +x_102 = l_Lean_IR_StructRC_Context_useVar(x_4, x_100, x_101, x_5); +x_103 = lean_ctor_get(x_102, 0); +lean_inc(x_103); +x_104 = lean_ctor_get(x_102, 1); +lean_inc(x_104); +if (lean_is_exclusive(x_102)) { + lean_ctor_release(x_102, 0); + lean_ctor_release(x_102, 1); + x_105 = x_102; +} else { + lean_dec_ref(x_102); + x_105 = lean_box(0); +} +x_106 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_106, 0, x_98); +lean_ctor_set(x_106, 1, x_100); +x_107 = lean_box(12); +x_108 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_108, 0, x_1); +lean_ctor_set(x_108, 1, x_2); +lean_ctor_set(x_108, 2, x_106); +lean_ctor_set(x_108, 3, x_107); +x_109 = l_Lean_IR_StructRC_Context_addInstr(x_103, x_108); +if (lean_is_scalar(x_105)) { + x_110 = lean_alloc_ctor(0, 2, 0); +} else { + x_110 = x_105; +} +lean_ctor_set(x_110, 0, x_109); +lean_ctor_set(x_110, 1, x_104); +return x_110; +} +} +case 2: +{ +lean_object* x_111; lean_object* x_112; uint8_t x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; uint8_t x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; uint8_t x_139; +x_111 = lean_ctor_get(x_3, 0); +lean_inc(x_111); +x_112 = lean_ctor_get(x_3, 1); +lean_inc_ref(x_112); +x_113 = lean_ctor_get_uint8(x_3, sizeof(void*)*3); +x_114 = lean_ctor_get(x_3, 2); +lean_inc_ref(x_114); +if (lean_is_exclusive(x_3)) { + lean_ctor_release(x_3, 0); + lean_ctor_release(x_3, 1); + lean_ctor_release(x_3, 2); + x_115 = x_3; +} else { + lean_dec_ref(x_3); + x_115 = lean_box(0); +} +x_133 = 0; +lean_inc(x_1); +lean_inc_ref(x_4); +x_134 = l_Lean_IR_StructRC_Context_useVar(x_4, x_1, x_133, x_5); +x_135 = lean_ctor_get(x_134, 0); +lean_inc(x_135); +x_136 = lean_ctor_get(x_134, 1); +lean_inc(x_136); +x_137 = lean_unsigned_to_nat(0u); +x_138 = lean_array_get_size(x_114); +x_139 = lean_nat_dec_lt(x_137, x_138); +if (x_139 == 0) +{ +lean_dec(x_136); +lean_dec(x_135); +x_116 = x_134; +goto block_132; +} +else +{ +uint8_t x_140; +x_140 = lean_nat_dec_le(x_138, x_138); +if (x_140 == 0) +{ +if (x_139 == 0) +{ +lean_dec(x_136); +lean_dec(x_135); +x_116 = x_134; +goto block_132; +} +else +{ +size_t x_141; size_t x_142; lean_object* x_143; +lean_dec_ref(x_134); +x_141 = 0; +x_142 = lean_usize_of_nat(x_138); +x_143 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_StructRC_visitExpr_spec__0(x_114, x_141, x_142, x_135, x_136); +x_116 = x_143; +goto block_132; +} +} +else +{ +size_t x_144; size_t x_145; lean_object* x_146; +lean_dec_ref(x_134); +x_144 = 0; +x_145 = lean_usize_of_nat(x_138); +x_146 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_StructRC_visitExpr_spec__0(x_114, x_144, x_145, x_135, x_136); +x_116 = x_146; +goto block_132; +} +} +block_132: +{ +uint8_t x_117; +x_117 = !lean_is_exclusive(x_116); +if (x_117 == 0) +{ +lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_118 = lean_ctor_get(x_116, 0); +x_119 = l_Lean_IR_StructRC_Context_renameVar_x21(x_4, x_111); +lean_dec(x_111); +lean_dec_ref(x_4); +if (lean_is_scalar(x_115)) { + x_120 = lean_alloc_ctor(2, 3, 1); +} else { + x_120 = x_115; +} +lean_ctor_set(x_120, 0, x_119); +lean_ctor_set(x_120, 1, x_112); +lean_ctor_set(x_120, 2, x_114); +lean_ctor_set_uint8(x_120, sizeof(void*)*3, x_113); +x_121 = lean_box(12); +x_122 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_122, 0, x_1); +lean_ctor_set(x_122, 1, x_2); +lean_ctor_set(x_122, 2, x_120); +lean_ctor_set(x_122, 3, x_121); +x_123 = l_Lean_IR_StructRC_Context_addInstr(x_118, x_122); +lean_ctor_set(x_116, 0, x_123); +return x_116; +} +else +{ +lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; +x_124 = lean_ctor_get(x_116, 0); +x_125 = lean_ctor_get(x_116, 1); +lean_inc(x_125); +lean_inc(x_124); +lean_dec(x_116); +x_126 = l_Lean_IR_StructRC_Context_renameVar_x21(x_4, x_111); +lean_dec(x_111); +lean_dec_ref(x_4); +if (lean_is_scalar(x_115)) { + x_127 = lean_alloc_ctor(2, 3, 1); +} else { + x_127 = x_115; +} +lean_ctor_set(x_127, 0, x_126); +lean_ctor_set(x_127, 1, x_112); +lean_ctor_set(x_127, 2, x_114); +lean_ctor_set_uint8(x_127, sizeof(void*)*3, x_113); +x_128 = lean_box(12); +x_129 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_129, 0, x_1); +lean_ctor_set(x_129, 1, x_2); +lean_ctor_set(x_129, 2, x_127); +lean_ctor_set(x_129, 3, x_128); +x_130 = l_Lean_IR_StructRC_Context_addInstr(x_124, x_129); +x_131 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_131, 0, x_130); +lean_ctor_set(x_131, 1, x_125); +return x_131; +} +} +} +case 3: +{ +lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_200; lean_object* x_201; +x_147 = lean_ctor_get(x_3, 0); +lean_inc(x_147); +x_148 = lean_ctor_get(x_3, 1); +lean_inc(x_148); +x_149 = lean_ctor_get(x_3, 2); +lean_inc(x_149); +if (lean_is_exclusive(x_3)) { + lean_ctor_release(x_3, 0); + lean_ctor_release(x_3, 1); + lean_ctor_release(x_3, 2); + x_150 = x_3; +} else { + lean_dec_ref(x_3); + x_150 = lean_box(0); +} +x_151 = lean_ctor_get(x_4, 0); +x_152 = l_Lean_IR_StructRC_Context_renameVar_x21(x_4, x_149); +lean_dec(x_149); +lean_inc(x_152); +lean_inc(x_148); +lean_inc(x_147); +x_200 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_200, 0, x_147); +lean_ctor_set(x_200, 1, x_148); +lean_ctor_set(x_200, 2, x_152); +x_201 = l_Std_DTreeMap_Internal_Impl_Const_get_x3f___at___00Lean_IR_StructRC_Context_renameVar_x21_spec__0___redArg(x_151, x_152); +if (lean_obj_tag(x_201) == 0) +{ +lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; +lean_dec(x_152); +lean_dec(x_150); +lean_dec(x_148); +lean_dec(x_147); +x_202 = lean_box(12); +x_203 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_203, 0, x_1); +lean_ctor_set(x_203, 1, x_2); +lean_ctor_set(x_203, 2, x_200); +lean_ctor_set(x_203, 3, x_202); +x_204 = l_Lean_IR_StructRC_Context_addInstr(x_4, x_203); +x_205 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_205, 0, x_204); +lean_ctor_set(x_205, 1, x_5); +return x_205; +} +else +{ +lean_object* x_206; lean_object* x_207; +x_206 = lean_ctor_get(x_201, 0); +lean_inc(x_206); +if (lean_is_exclusive(x_201)) { + lean_ctor_release(x_201, 0); + x_207 = x_201; +} else { + lean_dec_ref(x_201); + x_207 = lean_box(0); +} +switch (lean_obj_tag(x_206)) { +case 0: +{ +lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; +lean_dec(x_207); +lean_dec(x_152); +lean_dec(x_150); +lean_dec(x_148); +lean_dec(x_147); +lean_inc(x_1); +x_208 = l_Lean_IR_StructRC_Context_insert(x_1, x_206, x_4); +x_209 = lean_box(12); +x_210 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_210, 0, x_1); +lean_ctor_set(x_210, 1, x_2); +lean_ctor_set(x_210, 2, x_200); +lean_ctor_set(x_210, 3, x_209); +x_211 = l_Lean_IR_StructRC_Context_addInstr(x_208, x_210); +x_212 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_212, 0, x_211); +lean_ctor_set(x_212, 1, x_5); +return x_212; +} +case 1: +{ +lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; +lean_dec_ref(x_200); +x_213 = lean_ctor_get(x_206, 0); +lean_inc_ref(x_213); +x_214 = lean_ctor_get(x_206, 1); +lean_inc(x_214); +lean_dec_ref(x_206); +x_215 = lean_box(0); +x_216 = lean_array_get(x_215, x_213, x_147); +lean_dec_ref(x_213); +if (lean_obj_tag(x_216) == 10) +{ +lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; +x_217 = lean_ctor_get(x_216, 1); +lean_inc_ref(x_217); +lean_dec_ref(x_216); +x_218 = lean_array_get_size(x_217); +if (lean_is_scalar(x_207)) { + x_219 = lean_alloc_ctor(2, 1, 0); +} else { + x_219 = x_207; + lean_ctor_set_tag(x_219, 2); +} +lean_ctor_set(x_219, 0, x_214); +x_220 = lean_mk_array(x_218, x_219); +lean_inc_ref(x_220); +lean_inc_ref(x_217); +lean_inc(x_147); +x_221 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_221, 0, x_147); +lean_ctor_set(x_221, 1, x_217); +lean_ctor_set(x_221, 2, x_220); +x_190 = x_221; +x_191 = x_147; +x_192 = x_217; +x_193 = x_220; +goto block_199; +} +else +{ +lean_object* x_222; lean_object* x_223; +lean_dec(x_216); +lean_dec(x_214); +lean_dec(x_207); +lean_dec(x_147); +x_222 = l_Lean_IR_StructRC_Entry_ofStruct___closed__2; +x_223 = l_panic___at___00Lean_IR_StructRC_Context_addCtorKnowledge_spec__0(x_222); +if (lean_obj_tag(x_223) == 2) +{ +lean_object* x_224; lean_object* x_225; lean_object* x_226; +x_224 = lean_ctor_get(x_223, 0); +lean_inc(x_224); +x_225 = lean_ctor_get(x_223, 1); +lean_inc_ref(x_225); +x_226 = lean_ctor_get(x_223, 2); +lean_inc_ref(x_226); +x_190 = x_223; +x_191 = x_224; +x_192 = x_225; +x_193 = x_226; +goto block_199; +} +else +{ +lean_object* x_227; lean_object* x_228; lean_object* x_229; +lean_dec(x_223); +lean_dec(x_152); +lean_dec(x_150); +lean_dec(x_148); +lean_dec(x_2); +lean_dec(x_1); +x_227 = l_Lean_IR_StructRC_Context_addProjInfo___closed__11; +x_228 = l_panic___at___00Lean_IR_StructRC_visitExpr_spec__2(x_4, x_227); +x_229 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_229, 0, x_228); +lean_ctor_set(x_229, 1, x_5); +return x_229; +} +} +} +case 2: +{ +lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; uint8_t x_280; +lean_dec(x_150); +x_230 = lean_ctor_get(x_206, 0); +x_231 = lean_ctor_get(x_206, 1); +x_232 = lean_ctor_get(x_206, 2); +x_280 = lean_nat_dec_eq(x_147, x_230); +lean_dec(x_147); +if (x_280 == 0) +{ +lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; +lean_dec(x_207); +lean_dec_ref(x_206); +lean_dec(x_152); +lean_dec(x_148); +x_281 = lean_box(12); +x_282 = l_Lean_IR_StructRC_Context_addInstr(x_4, x_281); +x_283 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_283, 0, x_1); +lean_ctor_set(x_283, 1, x_2); +lean_ctor_set(x_283, 2, x_200); +lean_ctor_set(x_283, 3, x_281); +x_284 = l_Lean_IR_StructRC_Context_addInstr(x_282, x_283); +x_285 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_285, 0, x_284); +lean_ctor_set(x_285, 1, x_5); +return x_285; +} +else +{ +lean_object* x_286; uint8_t x_287; +lean_dec_ref(x_200); +x_286 = lean_array_get_size(x_231); +x_287 = lean_nat_dec_lt(x_148, x_286); +if (x_287 == 0) +{ +lean_object* x_288; lean_object* x_289; +x_288 = lean_box(0); +x_289 = l_outOfBounds___redArg(x_288); +x_233 = x_289; +goto block_279; +} +else +{ +lean_object* x_290; +x_290 = lean_array_fget_borrowed(x_232, x_148); +lean_inc(x_290); +x_233 = x_290; +goto block_279; +} +} +block_279: +{ +switch (lean_obj_tag(x_233)) { +case 0: +{ +lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; +lean_dec(x_207); +lean_dec(x_148); +lean_dec(x_2); +x_234 = l_Lean_IR_StructRC_Context_insert(x_152, x_206, x_4); +x_235 = lean_box(1); +x_236 = l_Lean_IR_StructRC_Context_insertRename(x_234, x_1, x_235); +x_237 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_237, 0, x_236); +lean_ctor_set(x_237, 1, x_5); +return x_237; +} +case 1: +{ +lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; +lean_dec(x_148); +lean_dec(x_2); +x_238 = lean_ctor_get(x_233, 0); +lean_inc(x_238); +lean_dec_ref(x_233); +x_239 = l_Lean_IR_StructRC_Context_insert(x_152, x_206, x_4); +if (lean_is_scalar(x_207)) { + x_240 = lean_alloc_ctor(0, 1, 0); +} else { + x_240 = x_207; + lean_ctor_set_tag(x_240, 0); +} +lean_ctor_set(x_240, 0, x_238); +x_241 = l_Lean_IR_StructRC_Context_insertRename(x_239, x_1, x_240); +x_242 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_242, 0, x_241); +lean_ctor_set(x_242, 1, x_5); +return x_242; +} +default: +{ +uint8_t x_243; +lean_inc_ref(x_232); +lean_inc_ref(x_231); +lean_inc(x_230); +lean_dec(x_207); +x_243 = !lean_is_exclusive(x_206); +if (x_243 == 0) +{ +lean_object* x_244; lean_object* x_245; lean_object* x_246; uint8_t x_247; +x_244 = lean_ctor_get(x_206, 2); +lean_dec(x_244); +x_245 = lean_ctor_get(x_206, 1); +lean_dec(x_245); +x_246 = lean_ctor_get(x_206, 0); +lean_dec(x_246); +x_247 = !lean_is_exclusive(x_233); +if (x_247 == 0) +{ +lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; +x_248 = lean_ctor_get(x_233, 0); +lean_inc(x_2); +lean_inc(x_1); +x_249 = l_Lean_IR_StructRC_Context_addVar(x_4, x_1, x_2, x_248); +lean_inc(x_1); +lean_ctor_set_tag(x_233, 1); +lean_ctor_set(x_233, 0, x_1); +x_250 = lean_array_set(x_232, x_148, x_233); +lean_inc(x_230); +lean_ctor_set(x_206, 2, x_250); +lean_inc(x_152); +x_251 = l_Lean_IR_StructRC_Context_insert(x_152, x_206, x_249); +x_252 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_252, 0, x_230); +lean_ctor_set(x_252, 1, x_148); +lean_ctor_set(x_252, 2, x_152); +x_253 = lean_box(12); +x_254 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_254, 0, x_1); +lean_ctor_set(x_254, 1, x_2); +lean_ctor_set(x_254, 2, x_252); +lean_ctor_set(x_254, 3, x_253); +x_255 = l_Lean_IR_StructRC_Context_addInstr(x_251, x_254); +x_256 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_256, 0, x_255); +lean_ctor_set(x_256, 1, x_5); +return x_256; +} +else +{ +lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; +x_257 = lean_ctor_get(x_233, 0); +lean_inc(x_257); +lean_dec(x_233); +lean_inc(x_2); +lean_inc(x_1); +x_258 = l_Lean_IR_StructRC_Context_addVar(x_4, x_1, x_2, x_257); +lean_inc(x_1); +x_259 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_259, 0, x_1); +x_260 = lean_array_set(x_232, x_148, x_259); +lean_inc(x_230); +lean_ctor_set(x_206, 2, x_260); +lean_inc(x_152); +x_261 = l_Lean_IR_StructRC_Context_insert(x_152, x_206, x_258); +x_262 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_262, 0, x_230); +lean_ctor_set(x_262, 1, x_148); +lean_ctor_set(x_262, 2, x_152); +x_263 = lean_box(12); +x_264 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_264, 0, x_1); +lean_ctor_set(x_264, 1, x_2); +lean_ctor_set(x_264, 2, x_262); +lean_ctor_set(x_264, 3, x_263); +x_265 = l_Lean_IR_StructRC_Context_addInstr(x_261, x_264); +x_266 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_266, 0, x_265); +lean_ctor_set(x_266, 1, x_5); +return x_266; +} +} +else +{ +lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; +lean_dec(x_206); +x_267 = lean_ctor_get(x_233, 0); +lean_inc(x_267); +if (lean_is_exclusive(x_233)) { + lean_ctor_release(x_233, 0); + x_268 = x_233; +} else { + lean_dec_ref(x_233); + x_268 = lean_box(0); +} +lean_inc(x_2); +lean_inc(x_1); +x_269 = l_Lean_IR_StructRC_Context_addVar(x_4, x_1, x_2, x_267); +lean_inc(x_1); +if (lean_is_scalar(x_268)) { + x_270 = lean_alloc_ctor(1, 1, 0); +} else { + x_270 = x_268; + lean_ctor_set_tag(x_270, 1); +} +lean_ctor_set(x_270, 0, x_1); +x_271 = lean_array_set(x_232, x_148, x_270); +lean_inc(x_230); +x_272 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_272, 0, x_230); +lean_ctor_set(x_272, 1, x_231); +lean_ctor_set(x_272, 2, x_271); +lean_inc(x_152); +x_273 = l_Lean_IR_StructRC_Context_insert(x_152, x_272, x_269); +x_274 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_274, 0, x_230); +lean_ctor_set(x_274, 1, x_148); +lean_ctor_set(x_274, 2, x_152); +x_275 = lean_box(12); +x_276 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_276, 0, x_1); +lean_ctor_set(x_276, 1, x_2); +lean_ctor_set(x_276, 2, x_274); +lean_ctor_set(x_276, 3, x_275); +x_277 = l_Lean_IR_StructRC_Context_addInstr(x_273, x_276); +x_278 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_278, 0, x_277); +lean_ctor_set(x_278, 1, x_5); +return x_278; +} +} +} +} +} +default: +{ +lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; +lean_dec(x_207); +lean_dec_ref(x_206); +lean_dec(x_152); +lean_dec(x_150); +lean_dec(x_148); +lean_dec(x_147); +lean_inc(x_2); +lean_inc(x_1); +x_291 = l_Lean_IR_StructRC_Context_addVarBasic(x_4, x_1, x_2); +x_292 = lean_box(12); +x_293 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_293, 0, x_1); +lean_ctor_set(x_293, 1, x_2); +lean_ctor_set(x_293, 2, x_200); +lean_ctor_set(x_293, 3, x_292); +x_294 = l_Lean_IR_StructRC_Context_addInstr(x_291, x_293); +x_295 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_295, 0, x_294); +lean_ctor_set(x_295, 1, x_5); +return x_295; +} +} +} +block_189: +{ +switch (lean_obj_tag(x_157)) { +case 0: +{ +lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; +lean_dec_ref(x_156); +lean_dec_ref(x_155); +lean_dec(x_153); +lean_dec(x_150); +lean_dec(x_148); +lean_dec(x_2); +x_158 = l_Lean_IR_StructRC_Context_insert(x_152, x_154, x_4); +x_159 = lean_box(1); +x_160 = l_Lean_IR_StructRC_Context_insertRename(x_158, x_1, x_159); +x_161 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_161, 0, x_160); +lean_ctor_set(x_161, 1, x_5); +return x_161; +} +case 1: +{ +lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; +lean_dec_ref(x_156); +lean_dec_ref(x_155); +lean_dec(x_153); +lean_dec(x_150); +lean_dec(x_148); +lean_dec(x_2); +x_162 = lean_ctor_get(x_157, 0); +lean_inc(x_162); +lean_dec_ref(x_157); +x_163 = l_Lean_IR_StructRC_Context_insert(x_152, x_154, x_4); +x_164 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_164, 0, x_162); +x_165 = l_Lean_IR_StructRC_Context_insertRename(x_163, x_1, x_164); +x_166 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_166, 0, x_165); +lean_ctor_set(x_166, 1, x_5); +return x_166; +} +default: +{ +uint8_t x_167; +lean_dec(x_154); +x_167 = !lean_is_exclusive(x_157); +if (x_167 == 0) +{ +lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; +x_168 = lean_ctor_get(x_157, 0); +lean_inc(x_2); +lean_inc(x_1); +x_169 = l_Lean_IR_StructRC_Context_addVar(x_4, x_1, x_2, x_168); +lean_inc(x_1); +lean_ctor_set_tag(x_157, 1); +lean_ctor_set(x_157, 0, x_1); +x_170 = lean_array_set(x_156, x_148, x_157); +lean_inc(x_153); +x_171 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_171, 0, x_153); +lean_ctor_set(x_171, 1, x_155); +lean_ctor_set(x_171, 2, x_170); +lean_inc(x_152); +x_172 = l_Lean_IR_StructRC_Context_insert(x_152, x_171, x_169); +if (lean_is_scalar(x_150)) { + x_173 = lean_alloc_ctor(3, 3, 0); +} else { + x_173 = x_150; +} +lean_ctor_set(x_173, 0, x_153); +lean_ctor_set(x_173, 1, x_148); +lean_ctor_set(x_173, 2, x_152); +x_174 = lean_box(12); +x_175 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_175, 0, x_1); +lean_ctor_set(x_175, 1, x_2); +lean_ctor_set(x_175, 2, x_173); +lean_ctor_set(x_175, 3, x_174); +x_176 = l_Lean_IR_StructRC_Context_addInstr(x_172, x_175); +x_177 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_177, 0, x_176); +lean_ctor_set(x_177, 1, x_5); +return x_177; +} +else +{ +lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; +x_178 = lean_ctor_get(x_157, 0); +lean_inc(x_178); +lean_dec(x_157); +lean_inc(x_2); +lean_inc(x_1); +x_179 = l_Lean_IR_StructRC_Context_addVar(x_4, x_1, x_2, x_178); +lean_inc(x_1); +x_180 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_180, 0, x_1); +x_181 = lean_array_set(x_156, x_148, x_180); +lean_inc(x_153); +x_182 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_182, 0, x_153); +lean_ctor_set(x_182, 1, x_155); +lean_ctor_set(x_182, 2, x_181); +lean_inc(x_152); +x_183 = l_Lean_IR_StructRC_Context_insert(x_152, x_182, x_179); +if (lean_is_scalar(x_150)) { + x_184 = lean_alloc_ctor(3, 3, 0); +} else { + x_184 = x_150; +} +lean_ctor_set(x_184, 0, x_153); +lean_ctor_set(x_184, 1, x_148); +lean_ctor_set(x_184, 2, x_152); +x_185 = lean_box(12); +x_186 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_186, 0, x_1); +lean_ctor_set(x_186, 1, x_2); +lean_ctor_set(x_186, 2, x_184); +lean_ctor_set(x_186, 3, x_185); +x_187 = l_Lean_IR_StructRC_Context_addInstr(x_183, x_186); +x_188 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_188, 0, x_187); +lean_ctor_set(x_188, 1, x_5); +return x_188; +} +} +} +} +block_199: +{ +lean_object* x_194; uint8_t x_195; +x_194 = lean_array_get_size(x_192); +x_195 = lean_nat_dec_lt(x_148, x_194); +if (x_195 == 0) +{ +lean_object* x_196; lean_object* x_197; +x_196 = lean_box(0); +x_197 = l_outOfBounds___redArg(x_196); +x_153 = x_191; +x_154 = x_190; +x_155 = x_192; +x_156 = x_193; +x_157 = x_197; +goto block_189; +} +else +{ +lean_object* x_198; +x_198 = lean_array_fget(x_193, x_148); +x_153 = x_191; +x_154 = x_190; +x_155 = x_192; +x_156 = x_193; +x_157 = x_198; +goto block_189; +} +} +} +case 4: +{ +uint8_t x_296; +x_296 = !lean_is_exclusive(x_3); +if (x_296 == 0) +{ +lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; +x_297 = lean_ctor_get(x_3, 2); +x_298 = l_Lean_IR_StructRC_Context_renameVar_x21(x_4, x_297); +lean_dec(x_297); +lean_ctor_set(x_3, 2, x_298); +x_299 = lean_box(12); +x_300 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_300, 0, x_1); +lean_ctor_set(x_300, 1, x_2); +lean_ctor_set(x_300, 2, x_3); +lean_ctor_set(x_300, 3, x_299); +x_301 = l_Lean_IR_StructRC_Context_addInstr(x_4, x_300); +x_302 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_302, 0, x_301); +lean_ctor_set(x_302, 1, x_5); +return x_302; +} +else +{ +lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; +x_303 = lean_ctor_get(x_3, 0); +x_304 = lean_ctor_get(x_3, 1); +x_305 = lean_ctor_get(x_3, 2); +lean_inc(x_305); +lean_inc(x_304); +lean_inc(x_303); +lean_dec(x_3); +x_306 = l_Lean_IR_StructRC_Context_renameVar_x21(x_4, x_305); +lean_dec(x_305); +x_307 = lean_alloc_ctor(4, 3, 0); +lean_ctor_set(x_307, 0, x_303); +lean_ctor_set(x_307, 1, x_304); +lean_ctor_set(x_307, 2, x_306); +x_308 = lean_box(12); +x_309 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_309, 0, x_1); +lean_ctor_set(x_309, 1, x_2); +lean_ctor_set(x_309, 2, x_307); +lean_ctor_set(x_309, 3, x_308); +x_310 = l_Lean_IR_StructRC_Context_addInstr(x_4, x_309); +x_311 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_311, 0, x_310); +lean_ctor_set(x_311, 1, x_5); +return x_311; +} +} +case 5: +{ +uint8_t x_312; +x_312 = !lean_is_exclusive(x_3); +if (x_312 == 0) +{ +lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; +x_313 = lean_ctor_get(x_3, 3); +x_314 = l_Lean_IR_StructRC_Context_renameVar_x21(x_4, x_313); +lean_dec(x_313); +lean_ctor_set(x_3, 3, x_314); +x_315 = lean_box(12); +x_316 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_316, 0, x_1); +lean_ctor_set(x_316, 1, x_2); +lean_ctor_set(x_316, 2, x_3); +lean_ctor_set(x_316, 3, x_315); +x_317 = l_Lean_IR_StructRC_Context_addInstr(x_4, x_316); +x_318 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_318, 0, x_317); +lean_ctor_set(x_318, 1, x_5); +return x_318; +} +else +{ +lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; +x_319 = lean_ctor_get(x_3, 0); +x_320 = lean_ctor_get(x_3, 1); +x_321 = lean_ctor_get(x_3, 2); +x_322 = lean_ctor_get(x_3, 3); +lean_inc(x_322); +lean_inc(x_321); +lean_inc(x_320); +lean_inc(x_319); +lean_dec(x_3); +x_323 = l_Lean_IR_StructRC_Context_renameVar_x21(x_4, x_322); +lean_dec(x_322); +x_324 = lean_alloc_ctor(5, 4, 0); +lean_ctor_set(x_324, 0, x_319); +lean_ctor_set(x_324, 1, x_320); +lean_ctor_set(x_324, 2, x_321); +lean_ctor_set(x_324, 3, x_323); +x_325 = lean_box(12); +x_326 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_326, 0, x_1); +lean_ctor_set(x_326, 1, x_2); +lean_ctor_set(x_326, 2, x_324); +lean_ctor_set(x_326, 3, x_325); +x_327 = l_Lean_IR_StructRC_Context_addInstr(x_4, x_326); +x_328 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_328, 0, x_327); +lean_ctor_set(x_328, 1, x_5); +return x_328; +} +} +case 6: +{ +uint8_t x_329; +x_329 = !lean_is_exclusive(x_3); +if (x_329 == 0) +{ +lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_339; lean_object* x_343; lean_object* x_344; uint8_t x_345; +x_330 = lean_ctor_get(x_3, 1); +x_331 = l_Lean_IR_StructRC_Context_renameArgs(x_4, x_330); +lean_inc_ref(x_331); +lean_ctor_set(x_3, 1, x_331); +x_343 = lean_array_get_size(x_331); +x_344 = lean_unsigned_to_nat(0u); +x_345 = lean_nat_dec_eq(x_343, x_344); +if (x_345 == 0) +{ +lean_object* x_346; uint8_t x_347; +lean_inc(x_2); +lean_inc(x_1); +x_346 = l_Lean_IR_StructRC_Context_addVarBasic(x_4, x_1, x_2); +x_347 = lean_nat_dec_lt(x_344, x_343); +if (x_347 == 0) +{ +lean_dec_ref(x_331); +x_332 = x_346; +x_333 = x_5; +goto block_338; +} +else +{ +uint8_t x_348; +x_348 = lean_nat_dec_le(x_343, x_343); +if (x_348 == 0) +{ +if (x_347 == 0) +{ +lean_dec_ref(x_331); +x_332 = x_346; +x_333 = x_5; +goto block_338; +} +else +{ +size_t x_349; size_t x_350; lean_object* x_351; +x_349 = 0; +x_350 = lean_usize_of_nat(x_343); +x_351 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_StructRC_visitExpr_spec__0(x_331, x_349, x_350, x_346, x_5); +lean_dec_ref(x_331); +x_339 = x_351; +goto block_342; +} +} +else +{ +size_t x_352; size_t x_353; lean_object* x_354; +x_352 = 0; +x_353 = lean_usize_of_nat(x_343); +x_354 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_StructRC_visitExpr_spec__0(x_331, x_352, x_353, x_346, x_5); +lean_dec_ref(x_331); +x_339 = x_354; +goto block_342; +} +} +} +else +{ +lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; +lean_dec_ref(x_331); +x_355 = lean_box(0); +lean_inc(x_1); +x_356 = l_Lean_IR_StructRC_Context_insert(x_1, x_355, x_4); +x_357 = lean_box(12); +x_358 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_358, 0, x_1); +lean_ctor_set(x_358, 1, x_2); +lean_ctor_set(x_358, 2, x_3); +lean_ctor_set(x_358, 3, x_357); +x_359 = l_Lean_IR_StructRC_Context_addInstr(x_356, x_358); +x_360 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_360, 0, x_359); +lean_ctor_set(x_360, 1, x_5); +return x_360; +} +block_338: +{ +lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; +x_334 = lean_box(12); +x_335 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_335, 0, x_1); +lean_ctor_set(x_335, 1, x_2); +lean_ctor_set(x_335, 2, x_3); +lean_ctor_set(x_335, 3, x_334); +x_336 = l_Lean_IR_StructRC_Context_addInstr(x_332, x_335); +x_337 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_337, 0, x_336); +lean_ctor_set(x_337, 1, x_333); +return x_337; +} +block_342: +{ +lean_object* x_340; lean_object* x_341; +x_340 = lean_ctor_get(x_339, 0); +lean_inc(x_340); +x_341 = lean_ctor_get(x_339, 1); +lean_inc(x_341); +lean_dec_ref(x_339); +x_332 = x_340; +x_333 = x_341; +goto block_338; +} +} +else +{ +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_372; lean_object* x_376; lean_object* x_377; uint8_t x_378; +x_361 = lean_ctor_get(x_3, 0); +x_362 = lean_ctor_get(x_3, 1); +lean_inc(x_362); +lean_inc(x_361); +lean_dec(x_3); +x_363 = l_Lean_IR_StructRC_Context_renameArgs(x_4, x_362); +lean_inc_ref(x_363); +x_364 = lean_alloc_ctor(6, 2, 0); +lean_ctor_set(x_364, 0, x_361); +lean_ctor_set(x_364, 1, x_363); +x_376 = lean_array_get_size(x_363); +x_377 = lean_unsigned_to_nat(0u); +x_378 = lean_nat_dec_eq(x_376, x_377); +if (x_378 == 0) +{ +lean_object* x_379; uint8_t x_380; +lean_inc(x_2); +lean_inc(x_1); +x_379 = l_Lean_IR_StructRC_Context_addVarBasic(x_4, x_1, x_2); +x_380 = lean_nat_dec_lt(x_377, x_376); +if (x_380 == 0) +{ +lean_dec_ref(x_363); +x_365 = x_379; +x_366 = x_5; +goto block_371; +} +else +{ +uint8_t x_381; +x_381 = lean_nat_dec_le(x_376, x_376); +if (x_381 == 0) +{ +if (x_380 == 0) +{ +lean_dec_ref(x_363); +x_365 = x_379; +x_366 = x_5; +goto block_371; +} +else +{ +size_t x_382; size_t x_383; lean_object* x_384; +x_382 = 0; +x_383 = lean_usize_of_nat(x_376); +x_384 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_StructRC_visitExpr_spec__0(x_363, x_382, x_383, x_379, x_5); +lean_dec_ref(x_363); +x_372 = x_384; +goto block_375; +} +} +else +{ +size_t x_385; size_t x_386; lean_object* x_387; +x_385 = 0; +x_386 = lean_usize_of_nat(x_376); +x_387 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_StructRC_visitExpr_spec__0(x_363, x_385, x_386, x_379, x_5); +lean_dec_ref(x_363); +x_372 = x_387; +goto block_375; +} +} +} +else +{ +lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; +lean_dec_ref(x_363); +x_388 = lean_box(0); +lean_inc(x_1); +x_389 = l_Lean_IR_StructRC_Context_insert(x_1, x_388, x_4); +x_390 = lean_box(12); +x_391 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_391, 0, x_1); +lean_ctor_set(x_391, 1, x_2); +lean_ctor_set(x_391, 2, x_364); +lean_ctor_set(x_391, 3, x_390); +x_392 = l_Lean_IR_StructRC_Context_addInstr(x_389, x_391); +x_393 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_393, 0, x_392); +lean_ctor_set(x_393, 1, x_5); +return x_393; +} +block_371: +{ +lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; +x_367 = lean_box(12); +x_368 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_368, 0, x_1); +lean_ctor_set(x_368, 1, x_2); +lean_ctor_set(x_368, 2, x_364); +lean_ctor_set(x_368, 3, x_367); +x_369 = l_Lean_IR_StructRC_Context_addInstr(x_365, x_368); +x_370 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_370, 0, x_369); +lean_ctor_set(x_370, 1, x_366); +return x_370; +} +block_375: +{ +lean_object* x_373; lean_object* x_374; +x_373 = lean_ctor_get(x_372, 0); +lean_inc(x_373); +x_374 = lean_ctor_get(x_372, 1); +lean_inc(x_374); +lean_dec_ref(x_372); +x_365 = x_373; +x_366 = x_374; +goto block_371; +} +} +} +case 7: +{ +lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_406; lean_object* x_410; lean_object* x_411; uint8_t x_412; +x_394 = lean_ctor_get(x_3, 0); +lean_inc(x_394); +x_395 = lean_ctor_get(x_3, 1); +lean_inc_ref(x_395); +if (lean_is_exclusive(x_3)) { + lean_ctor_release(x_3, 0); + lean_ctor_release(x_3, 1); + x_396 = x_3; +} else { + lean_dec_ref(x_3); + x_396 = lean_box(0); +} +x_397 = l_Lean_IR_StructRC_Context_renameArgs(x_4, x_395); +x_410 = lean_unsigned_to_nat(0u); +x_411 = lean_array_get_size(x_397); +x_412 = lean_nat_dec_lt(x_410, x_411); +if (x_412 == 0) +{ +x_398 = x_4; +x_399 = x_5; +goto block_405; +} +else +{ +uint8_t x_413; +x_413 = lean_nat_dec_le(x_411, x_411); +if (x_413 == 0) +{ +if (x_412 == 0) +{ +x_398 = x_4; +x_399 = x_5; +goto block_405; +} +else +{ +size_t x_414; size_t x_415; lean_object* x_416; +x_414 = 0; +x_415 = lean_usize_of_nat(x_411); +x_416 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_StructRC_visitExpr_spec__0(x_397, x_414, x_415, x_4, x_5); +x_406 = x_416; +goto block_409; +} +} +else +{ +size_t x_417; size_t x_418; lean_object* x_419; +x_417 = 0; +x_418 = lean_usize_of_nat(x_411); +x_419 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_StructRC_visitExpr_spec__0(x_397, x_417, x_418, x_4, x_5); +x_406 = x_419; +goto block_409; +} +} +block_405: +{ +lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; +if (lean_is_scalar(x_396)) { + x_400 = lean_alloc_ctor(7, 2, 0); +} else { + x_400 = x_396; +} +lean_ctor_set(x_400, 0, x_394); +lean_ctor_set(x_400, 1, x_397); +x_401 = lean_box(12); +x_402 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_402, 0, x_1); +lean_ctor_set(x_402, 1, x_2); +lean_ctor_set(x_402, 2, x_400); +lean_ctor_set(x_402, 3, x_401); +x_403 = l_Lean_IR_StructRC_Context_addInstr(x_398, x_402); +x_404 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_404, 0, x_403); +lean_ctor_set(x_404, 1, x_399); +return x_404; +} +block_409: +{ +lean_object* x_407; lean_object* x_408; +x_407 = lean_ctor_get(x_406, 0); +lean_inc(x_407); +x_408 = lean_ctor_get(x_406, 1); +lean_inc(x_408); +lean_dec_ref(x_406); +x_398 = x_407; +x_399 = x_408; +goto block_405; +} +} +case 8: +{ +lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; +x_420 = lean_ctor_get(x_3, 0); +lean_inc(x_420); +x_421 = lean_ctor_get(x_3, 1); +lean_inc_ref(x_421); +if (lean_is_exclusive(x_3)) { + lean_ctor_release(x_3, 0); + lean_ctor_release(x_3, 1); + x_422 = x_3; +} else { + lean_dec_ref(x_3); + x_422 = lean_box(0); +} +x_423 = l_Lean_IR_StructRC_Context_renameVar(x_4, x_420); +if (lean_obj_tag(x_423) == 0) +{ +lean_object* x_424; uint8_t x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_445; lean_object* x_446; uint8_t x_447; +x_424 = lean_ctor_get(x_423, 0); +lean_inc(x_424); +lean_dec_ref(x_423); +x_425 = 0; +lean_inc(x_424); +lean_inc_ref(x_4); +x_426 = l_Lean_IR_StructRC_Context_useVar(x_4, x_424, x_425, x_5); +x_427 = lean_ctor_get(x_426, 0); +lean_inc(x_427); +x_428 = lean_ctor_get(x_426, 1); +lean_inc(x_428); +x_429 = l_Lean_IR_StructRC_Context_renameArgs(x_4, x_421); +lean_dec_ref(x_4); +x_445 = lean_unsigned_to_nat(0u); +x_446 = lean_array_get_size(x_429); +x_447 = lean_nat_dec_lt(x_445, x_446); +if (x_447 == 0) +{ +lean_dec(x_428); +lean_dec(x_427); +x_430 = x_426; +goto block_444; +} +else +{ +uint8_t x_448; +x_448 = lean_nat_dec_le(x_446, x_446); +if (x_448 == 0) +{ +if (x_447 == 0) +{ +lean_dec(x_428); +lean_dec(x_427); +x_430 = x_426; +goto block_444; +} +else +{ +size_t x_449; size_t x_450; lean_object* x_451; +lean_dec_ref(x_426); +x_449 = 0; +x_450 = lean_usize_of_nat(x_446); +x_451 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_StructRC_visitExpr_spec__0(x_429, x_449, x_450, x_427, x_428); +x_430 = x_451; +goto block_444; +} +} +else +{ +size_t x_452; size_t x_453; lean_object* x_454; +lean_dec_ref(x_426); +x_452 = 0; +x_453 = lean_usize_of_nat(x_446); +x_454 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_StructRC_visitExpr_spec__0(x_429, x_452, x_453, x_427, x_428); +x_430 = x_454; +goto block_444; +} +} +block_444: +{ +uint8_t x_431; +x_431 = !lean_is_exclusive(x_430); +if (x_431 == 0) +{ +lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; +x_432 = lean_ctor_get(x_430, 0); +if (lean_is_scalar(x_422)) { + x_433 = lean_alloc_ctor(8, 2, 0); +} else { + x_433 = x_422; +} +lean_ctor_set(x_433, 0, x_424); +lean_ctor_set(x_433, 1, x_429); +x_434 = lean_box(12); +x_435 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_435, 0, x_1); +lean_ctor_set(x_435, 1, x_2); +lean_ctor_set(x_435, 2, x_433); +lean_ctor_set(x_435, 3, x_434); +x_436 = l_Lean_IR_StructRC_Context_addInstr(x_432, x_435); +lean_ctor_set(x_430, 0, x_436); +return x_430; +} +else +{ +lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; +x_437 = lean_ctor_get(x_430, 0); +x_438 = lean_ctor_get(x_430, 1); +lean_inc(x_438); +lean_inc(x_437); +lean_dec(x_430); +if (lean_is_scalar(x_422)) { + x_439 = lean_alloc_ctor(8, 2, 0); +} else { + x_439 = x_422; +} +lean_ctor_set(x_439, 0, x_424); +lean_ctor_set(x_439, 1, x_429); +x_440 = lean_box(12); +x_441 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_441, 0, x_1); +lean_ctor_set(x_441, 1, x_2); +lean_ctor_set(x_441, 2, x_439); +lean_ctor_set(x_441, 3, x_440); +x_442 = l_Lean_IR_StructRC_Context_addInstr(x_437, x_441); +x_443 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_443, 0, x_442); +lean_ctor_set(x_443, 1, x_438); +return x_443; +} +} +} +else +{ +lean_object* x_455; lean_object* x_456; +lean_dec(x_422); +lean_dec_ref(x_421); +lean_dec(x_2); +x_455 = l_Lean_IR_StructRC_Context_insertRename(x_4, x_1, x_423); +x_456 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_456, 0, x_455); +lean_ctor_set(x_456, 1, x_5); +return x_456; +} +} +case 9: +{ +uint8_t x_457; +x_457 = !lean_is_exclusive(x_3); +if (x_457 == 0) +{ +lean_object* x_458; lean_object* x_459; lean_object* x_460; uint8_t x_461; lean_object* x_462; uint8_t x_463; +x_458 = lean_ctor_get(x_3, 1); +x_459 = l_Lean_IR_StructRC_Context_renameVar_x21(x_4, x_458); +lean_dec(x_458); +lean_inc(x_2); +lean_inc(x_1); +x_460 = l_Lean_IR_StructRC_Context_addVarBasic(x_4, x_1, x_2); +x_461 = 0; +lean_inc(x_459); +x_462 = l_Lean_IR_StructRC_Context_useVar(x_460, x_459, x_461, x_5); +x_463 = !lean_is_exclusive(x_462); +if (x_463 == 0) +{ +lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; +x_464 = lean_ctor_get(x_462, 0); +lean_ctor_set(x_3, 1, x_459); +x_465 = lean_box(12); +x_466 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_466, 0, x_1); +lean_ctor_set(x_466, 1, x_2); +lean_ctor_set(x_466, 2, x_3); +lean_ctor_set(x_466, 3, x_465); +x_467 = l_Lean_IR_StructRC_Context_addInstr(x_464, x_466); +lean_ctor_set(x_462, 0, x_467); +return x_462; +} +else +{ +lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; +x_468 = lean_ctor_get(x_462, 0); +x_469 = lean_ctor_get(x_462, 1); +lean_inc(x_469); +lean_inc(x_468); +lean_dec(x_462); +lean_ctor_set(x_3, 1, x_459); +x_470 = lean_box(12); +x_471 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_471, 0, x_1); +lean_ctor_set(x_471, 1, x_2); +lean_ctor_set(x_471, 2, x_3); +lean_ctor_set(x_471, 3, x_470); +x_472 = l_Lean_IR_StructRC_Context_addInstr(x_468, x_471); +x_473 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_473, 0, x_472); +lean_ctor_set(x_473, 1, x_469); +return x_473; +} +} +else +{ +lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; uint8_t x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; +x_474 = lean_ctor_get(x_3, 0); +x_475 = lean_ctor_get(x_3, 1); +lean_inc(x_475); +lean_inc(x_474); +lean_dec(x_3); +x_476 = l_Lean_IR_StructRC_Context_renameVar_x21(x_4, x_475); +lean_dec(x_475); +lean_inc(x_2); +lean_inc(x_1); +x_477 = l_Lean_IR_StructRC_Context_addVarBasic(x_4, x_1, x_2); +x_478 = 0; +lean_inc(x_476); +x_479 = l_Lean_IR_StructRC_Context_useVar(x_477, x_476, x_478, x_5); +x_480 = lean_ctor_get(x_479, 0); +lean_inc(x_480); +x_481 = lean_ctor_get(x_479, 1); +lean_inc(x_481); +if (lean_is_exclusive(x_479)) { + lean_ctor_release(x_479, 0); + lean_ctor_release(x_479, 1); + x_482 = x_479; +} else { + lean_dec_ref(x_479); + x_482 = lean_box(0); +} +x_483 = lean_alloc_ctor(9, 2, 0); +lean_ctor_set(x_483, 0, x_474); +lean_ctor_set(x_483, 1, x_476); +x_484 = lean_box(12); +x_485 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_485, 0, x_1); +lean_ctor_set(x_485, 1, x_2); +lean_ctor_set(x_485, 2, x_483); +lean_ctor_set(x_485, 3, x_484); +x_486 = l_Lean_IR_StructRC_Context_addInstr(x_480, x_485); +if (lean_is_scalar(x_482)) { + x_487 = lean_alloc_ctor(0, 2, 0); +} else { + x_487 = x_482; +} +lean_ctor_set(x_487, 0, x_486); +lean_ctor_set(x_487, 1, x_481); +return x_487; +} +} +case 10: +{ +uint8_t x_488; +x_488 = !lean_is_exclusive(x_3); +if (x_488 == 0) +{ +lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; +x_489 = lean_ctor_get(x_3, 0); +x_490 = l_Lean_IR_StructRC_Context_renameVar_x21(x_4, x_489); +lean_dec(x_489); +lean_ctor_set(x_3, 0, x_490); +x_491 = lean_box(12); +x_492 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_492, 0, x_1); +lean_ctor_set(x_492, 1, x_2); +lean_ctor_set(x_492, 2, x_3); +lean_ctor_set(x_492, 3, x_491); +x_493 = l_Lean_IR_StructRC_Context_addInstr(x_4, x_492); +x_494 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_494, 0, x_493); +lean_ctor_set(x_494, 1, x_5); +return x_494; +} +else +{ +lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; +x_495 = lean_ctor_get(x_3, 0); +lean_inc(x_495); +lean_dec(x_3); +x_496 = l_Lean_IR_StructRC_Context_renameVar_x21(x_4, x_495); +lean_dec(x_495); +x_497 = lean_alloc_ctor(10, 1, 0); +lean_ctor_set(x_497, 0, x_496); +x_498 = lean_box(12); +x_499 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_499, 0, x_1); +lean_ctor_set(x_499, 1, x_2); +lean_ctor_set(x_499, 2, x_497); +lean_ctor_set(x_499, 3, x_498); +x_500 = l_Lean_IR_StructRC_Context_addInstr(x_4, x_499); +x_501 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_501, 0, x_500); +lean_ctor_set(x_501, 1, x_5); +return x_501; +} +} +case 11: +{ +lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; +x_502 = lean_box(12); +x_503 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_503, 0, x_1); +lean_ctor_set(x_503, 1, x_2); +lean_ctor_set(x_503, 2, x_3); +lean_ctor_set(x_503, 3, x_502); +x_504 = l_Lean_IR_StructRC_Context_addInstr(x_4, x_503); +x_505 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_505, 0, x_504); +lean_ctor_set(x_505, 1, x_5); +return x_505; +} +default: +{ +uint8_t x_506; +x_506 = !lean_is_exclusive(x_3); +if (x_506 == 0) +{ +lean_object* x_507; lean_object* x_508; +x_507 = lean_ctor_get(x_3, 0); +x_508 = l_Lean_IR_StructRC_Context_renameVar(x_4, x_507); +if (lean_obj_tag(x_508) == 0) +{ +lean_object* x_509; uint8_t x_510; lean_object* x_511; uint8_t x_512; +x_509 = lean_ctor_get(x_508, 0); +lean_inc(x_509); +lean_dec_ref(x_508); +x_510 = 0; +lean_inc(x_509); +x_511 = l_Lean_IR_StructRC_Context_useVar(x_4, x_509, x_510, x_5); +x_512 = !lean_is_exclusive(x_511); +if (x_512 == 0) +{ +lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; +x_513 = lean_ctor_get(x_511, 0); +lean_ctor_set(x_3, 0, x_509); +x_514 = lean_box(12); +x_515 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_515, 0, x_1); +lean_ctor_set(x_515, 1, x_2); +lean_ctor_set(x_515, 2, x_3); +lean_ctor_set(x_515, 3, x_514); +x_516 = l_Lean_IR_StructRC_Context_addInstr(x_513, x_515); +lean_ctor_set(x_511, 0, x_516); +return x_511; +} +else +{ +lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; +x_517 = lean_ctor_get(x_511, 0); +x_518 = lean_ctor_get(x_511, 1); +lean_inc(x_518); +lean_inc(x_517); +lean_dec(x_511); +lean_ctor_set(x_3, 0, x_509); +x_519 = lean_box(12); +x_520 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_520, 0, x_1); +lean_ctor_set(x_520, 1, x_2); +lean_ctor_set(x_520, 2, x_3); +lean_ctor_set(x_520, 3, x_519); +x_521 = l_Lean_IR_StructRC_Context_addInstr(x_517, x_520); +x_522 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_522, 0, x_521); +lean_ctor_set(x_522, 1, x_518); +return x_522; +} +} +else +{ +lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; +lean_free_object(x_3); +x_523 = l_Lean_IR_StructRC_visitExpr___closed__1; +x_524 = lean_box(12); +x_525 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_525, 0, x_1); +lean_ctor_set(x_525, 1, x_2); +lean_ctor_set(x_525, 2, x_523); +lean_ctor_set(x_525, 3, x_524); +x_526 = l_Lean_IR_StructRC_Context_addInstr(x_4, x_525); +x_527 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_527, 0, x_526); +lean_ctor_set(x_527, 1, x_5); +return x_527; +} +} +else +{ +lean_object* x_528; lean_object* x_529; +x_528 = lean_ctor_get(x_3, 0); +lean_inc(x_528); +lean_dec(x_3); +x_529 = l_Lean_IR_StructRC_Context_renameVar(x_4, x_528); +if (lean_obj_tag(x_529) == 0) +{ +lean_object* x_530; uint8_t x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; +x_530 = lean_ctor_get(x_529, 0); +lean_inc(x_530); +lean_dec_ref(x_529); +x_531 = 0; +lean_inc(x_530); +x_532 = l_Lean_IR_StructRC_Context_useVar(x_4, x_530, x_531, x_5); +x_533 = lean_ctor_get(x_532, 0); +lean_inc(x_533); +x_534 = lean_ctor_get(x_532, 1); +lean_inc(x_534); +if (lean_is_exclusive(x_532)) { + lean_ctor_release(x_532, 0); + lean_ctor_release(x_532, 1); + x_535 = x_532; +} else { + lean_dec_ref(x_532); + x_535 = lean_box(0); +} +x_536 = lean_alloc_ctor(12, 1, 0); +lean_ctor_set(x_536, 0, x_530); +x_537 = lean_box(12); +x_538 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_538, 0, x_1); +lean_ctor_set(x_538, 1, x_2); +lean_ctor_set(x_538, 2, x_536); +lean_ctor_set(x_538, 3, x_537); +x_539 = l_Lean_IR_StructRC_Context_addInstr(x_533, x_538); +if (lean_is_scalar(x_535)) { + x_540 = lean_alloc_ctor(0, 2, 0); +} else { + x_540 = x_535; +} +lean_ctor_set(x_540, 0, x_539); +lean_ctor_set(x_540, 1, x_534); +return x_540; +} +else +{ +lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; +x_541 = l_Lean_IR_StructRC_visitExpr___closed__1; +x_542 = lean_box(12); +x_543 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_543, 0, x_1); +lean_ctor_set(x_543, 1, x_2); +lean_ctor_set(x_543, 2, x_541); +lean_ctor_set(x_543, 3, x_542); +x_544 = l_Lean_IR_StructRC_Context_addInstr(x_4, x_543); +x_545 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_545, 0, x_544); +lean_ctor_set(x_545, 1, x_5); +return x_545; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00Lean_IR_StructRC_visitExpr_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00Lean_IR_StructRC_visitExpr_spec__1(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_StructRC_visitExpr_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_7 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_8 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_StructRC_visitExpr_spec__0(x_1, x_6, x_7, x_4, x_5); +lean_dec_ref(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00Lean_IR_StructRC_visitExpr_spec__1___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00Lean_IR_StructRC_visitExpr_spec__1___redArg(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_StructRC_visitFnBody_spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; uint8_t x_12; +x_12 = lean_usize_dec_lt(x_6, x_5); +if (x_12 == 0) +{ +lean_inc_ref(x_7); +return x_7; +} +else +{ +lean_object* x_13; +x_13 = lean_array_uget(x_4, x_6); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_14 = lean_ctor_get(x_13, 0); +lean_inc_ref(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec_ref(x_13); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec_ref(x_14); +x_17 = lean_nat_dec_eq(x_16, x_2); +lean_dec(x_16); +if (x_17 == 0) +{ +size_t x_18; size_t x_19; +lean_dec(x_15); +x_18 = 1; +x_19 = lean_usize_add(x_6, x_18); +{ +size_t _tmp_5 = x_19; +lean_object* _tmp_6 = x_3; +x_6 = _tmp_5; +x_7 = _tmp_6; +} +goto _start; +} +else +{ +lean_object* x_21; +x_21 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_21, 0, x_15); +x_8 = x_21; +goto block_11; +} +} +else +{ +uint8_t x_22; +x_22 = !lean_is_exclusive(x_13); +if (x_22 == 0) +{ +x_8 = x_13; +goto block_11; +} +else +{ +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_13, 0); +lean_inc(x_23); +lean_dec(x_13); +x_24 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_24, 0, x_23); +x_8 = x_24; +goto block_11; +} +} +} +block_11: +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_9, 0, x_8); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_1); +return x_10; +} +} +} +static lean_object* _init_l_Lean_IR_StructRC_visitFnBody___closed__0() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = lean_box(0); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_visitFnBody(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 0: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_1, 1); +lean_inc(x_5); +x_6 = lean_ctor_get(x_1, 2); +lean_inc_ref(x_6); +x_7 = lean_ctor_get(x_1, 3); +lean_inc(x_7); +lean_dec_ref(x_1); +x_8 = l_Lean_IR_StructRC_visitExpr(x_4, x_5, x_6, x_2, x_3); +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_ref(x_8); +x_1 = x_7; +x_2 = x_9; +x_3 = x_10; +goto _start; +} +case 1: +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_1); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_13 = lean_ctor_get(x_1, 1); +x_14 = lean_ctor_get(x_1, 2); +x_15 = lean_ctor_get(x_1, 3); +lean_inc_ref(x_2); +x_16 = l_Lean_IR_StructRC_Context_resetRC(x_2); +x_17 = l_Lean_IR_StructRC_Context_addParams(x_13, x_16); +x_18 = l_Lean_IR_StructRC_visitFnBody(x_14, x_17, x_3); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec_ref(x_18); +x_21 = lean_box(12); +lean_ctor_set(x_1, 3, x_21); +lean_ctor_set(x_1, 2, x_19); +x_22 = l_Lean_IR_StructRC_Context_addInstr(x_2, x_1); +x_1 = x_15; +x_2 = x_22; +x_3 = x_20; +goto _start; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_24 = lean_ctor_get(x_1, 0); +x_25 = lean_ctor_get(x_1, 1); +x_26 = lean_ctor_get(x_1, 2); +x_27 = lean_ctor_get(x_1, 3); +lean_inc(x_27); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_1); +lean_inc_ref(x_2); +x_28 = l_Lean_IR_StructRC_Context_resetRC(x_2); +x_29 = l_Lean_IR_StructRC_Context_addParams(x_25, x_28); +x_30 = l_Lean_IR_StructRC_visitFnBody(x_26, x_29, x_3); +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec_ref(x_30); +x_33 = lean_box(12); +x_34 = lean_alloc_ctor(1, 4, 0); +lean_ctor_set(x_34, 0, x_24); +lean_ctor_set(x_34, 1, x_25); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_34, 3, x_33); +x_35 = l_Lean_IR_StructRC_Context_addInstr(x_2, x_34); +x_1 = x_27; +x_2 = x_35; +x_3 = x_32; +goto _start; +} +} +case 2: +{ +uint8_t x_37; +x_37 = !lean_is_exclusive(x_1); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_38 = lean_ctor_get(x_1, 0); +x_39 = lean_ctor_get(x_1, 2); +x_40 = lean_ctor_get(x_1, 3); +x_41 = l_Lean_IR_StructRC_Context_renameVar_x21(x_2, x_38); +lean_dec(x_38); +x_42 = 0; +lean_inc(x_41); +lean_inc_ref(x_2); +x_43 = l_Lean_IR_StructRC_Context_useVar(x_2, x_41, x_42, x_3); +x_44 = lean_ctor_get(x_43, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_43, 1); +lean_inc(x_45); +lean_dec_ref(x_43); +x_46 = l_Lean_IR_StructRC_Context_renameArg(x_2, x_39); +lean_dec(x_39); +lean_dec_ref(x_2); +lean_inc(x_46); +x_47 = l_Lean_IR_StructRC_Context_useArg(x_44, x_46, x_45); +x_48 = lean_ctor_get(x_47, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_47, 1); +lean_inc(x_49); +lean_dec_ref(x_47); +x_50 = lean_box(12); +lean_ctor_set(x_1, 3, x_50); +lean_ctor_set(x_1, 2, x_46); +lean_ctor_set(x_1, 0, x_41); +x_51 = l_Lean_IR_StructRC_Context_addInstr(x_48, x_1); +x_1 = x_40; +x_2 = x_51; +x_3 = x_49; +goto _start; +} +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; uint8_t x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_53 = lean_ctor_get(x_1, 0); +x_54 = lean_ctor_get(x_1, 1); +x_55 = lean_ctor_get(x_1, 2); +x_56 = lean_ctor_get(x_1, 3); +lean_inc(x_56); +lean_inc(x_55); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_1); +x_57 = l_Lean_IR_StructRC_Context_renameVar_x21(x_2, x_53); +lean_dec(x_53); +x_58 = 0; +lean_inc(x_57); +lean_inc_ref(x_2); +x_59 = l_Lean_IR_StructRC_Context_useVar(x_2, x_57, x_58, x_3); +x_60 = lean_ctor_get(x_59, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_59, 1); +lean_inc(x_61); +lean_dec_ref(x_59); +x_62 = l_Lean_IR_StructRC_Context_renameArg(x_2, x_55); +lean_dec(x_55); +lean_dec_ref(x_2); +lean_inc(x_62); +x_63 = l_Lean_IR_StructRC_Context_useArg(x_60, x_62, x_61); +x_64 = lean_ctor_get(x_63, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_63, 1); +lean_inc(x_65); +lean_dec_ref(x_63); +x_66 = lean_box(12); +x_67 = lean_alloc_ctor(2, 4, 0); +lean_ctor_set(x_67, 0, x_57); +lean_ctor_set(x_67, 1, x_54); +lean_ctor_set(x_67, 2, x_62); +lean_ctor_set(x_67, 3, x_66); +x_68 = l_Lean_IR_StructRC_Context_addInstr(x_64, x_67); +x_1 = x_56; +x_2 = x_68; +x_3 = x_65; +goto _start; +} +} +case 3: +{ +uint8_t x_70; +x_70 = !lean_is_exclusive(x_1); +if (x_70 == 0) +{ +lean_object* x_71; lean_object* x_72; lean_object* x_73; uint8_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_71 = lean_ctor_get(x_1, 0); +x_72 = lean_ctor_get(x_1, 2); +x_73 = l_Lean_IR_StructRC_Context_renameVar_x21(x_2, x_71); +lean_dec(x_71); +x_74 = 0; +lean_inc(x_73); +x_75 = l_Lean_IR_StructRC_Context_useVar(x_2, x_73, x_74, x_3); +x_76 = lean_ctor_get(x_75, 0); +lean_inc(x_76); +x_77 = lean_ctor_get(x_75, 1); +lean_inc(x_77); +lean_dec_ref(x_75); +x_78 = lean_box(12); +lean_ctor_set(x_1, 2, x_78); +lean_ctor_set(x_1, 0, x_73); +x_79 = l_Lean_IR_StructRC_Context_addInstr(x_76, x_1); +x_1 = x_72; +x_2 = x_79; +x_3 = x_77; +goto _start; +} +else +{ +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; uint8_t x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_81 = lean_ctor_get(x_1, 0); +x_82 = lean_ctor_get(x_1, 1); +x_83 = lean_ctor_get(x_1, 2); +lean_inc(x_83); +lean_inc(x_82); +lean_inc(x_81); +lean_dec(x_1); +x_84 = l_Lean_IR_StructRC_Context_renameVar_x21(x_2, x_81); +lean_dec(x_81); +x_85 = 0; +lean_inc(x_84); +x_86 = l_Lean_IR_StructRC_Context_useVar(x_2, x_84, x_85, x_3); +x_87 = lean_ctor_get(x_86, 0); +lean_inc(x_87); +x_88 = lean_ctor_get(x_86, 1); +lean_inc(x_88); +lean_dec_ref(x_86); +x_89 = lean_box(12); +x_90 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_90, 0, x_84); +lean_ctor_set(x_90, 1, x_82); +lean_ctor_set(x_90, 2, x_89); +x_91 = l_Lean_IR_StructRC_Context_addInstr(x_87, x_90); +x_1 = x_83; +x_2 = x_91; +x_3 = x_88; +goto _start; +} +} +case 4: +{ +uint8_t x_93; +x_93 = !lean_is_exclusive(x_1); +if (x_93 == 0) +{ +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; uint8_t x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_94 = lean_ctor_get(x_1, 0); +x_95 = lean_ctor_get(x_1, 3); +x_96 = lean_ctor_get(x_1, 4); +x_97 = l_Lean_IR_StructRC_Context_renameVar_x21(x_2, x_94); +lean_dec(x_94); +x_98 = 0; +lean_inc(x_97); +lean_inc_ref(x_2); +x_99 = l_Lean_IR_StructRC_Context_useVar(x_2, x_97, x_98, x_3); +x_100 = lean_ctor_get(x_99, 0); +lean_inc(x_100); +x_101 = lean_ctor_get(x_99, 1); +lean_inc(x_101); +lean_dec_ref(x_99); +x_102 = l_Lean_IR_StructRC_Context_renameVar_x21(x_2, x_95); +lean_dec(x_95); +lean_dec_ref(x_2); +x_103 = lean_box(12); +lean_ctor_set(x_1, 4, x_103); +lean_ctor_set(x_1, 3, x_102); +lean_ctor_set(x_1, 0, x_97); +x_104 = l_Lean_IR_StructRC_Context_addInstr(x_100, x_1); +x_1 = x_96; +x_2 = x_104; +x_3 = x_101; +goto _start; +} +else +{ +lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; uint8_t x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; +x_106 = lean_ctor_get(x_1, 0); +x_107 = lean_ctor_get(x_1, 1); +x_108 = lean_ctor_get(x_1, 2); +x_109 = lean_ctor_get(x_1, 3); +x_110 = lean_ctor_get(x_1, 4); +lean_inc(x_110); +lean_inc(x_109); +lean_inc(x_108); +lean_inc(x_107); +lean_inc(x_106); +lean_dec(x_1); +x_111 = l_Lean_IR_StructRC_Context_renameVar_x21(x_2, x_106); +lean_dec(x_106); +x_112 = 0; +lean_inc(x_111); +lean_inc_ref(x_2); +x_113 = l_Lean_IR_StructRC_Context_useVar(x_2, x_111, x_112, x_3); +x_114 = lean_ctor_get(x_113, 0); +lean_inc(x_114); +x_115 = lean_ctor_get(x_113, 1); +lean_inc(x_115); +lean_dec_ref(x_113); +x_116 = l_Lean_IR_StructRC_Context_renameVar_x21(x_2, x_109); +lean_dec(x_109); +lean_dec_ref(x_2); +x_117 = lean_box(12); +x_118 = lean_alloc_ctor(4, 5, 0); +lean_ctor_set(x_118, 0, x_111); +lean_ctor_set(x_118, 1, x_107); +lean_ctor_set(x_118, 2, x_108); +lean_ctor_set(x_118, 3, x_116); +lean_ctor_set(x_118, 4, x_117); +x_119 = l_Lean_IR_StructRC_Context_addInstr(x_114, x_118); +x_1 = x_110; +x_2 = x_119; +x_3 = x_115; +goto _start; +} +} +case 5: +{ +uint8_t x_121; +x_121 = !lean_is_exclusive(x_1); +if (x_121 == 0) +{ +lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; uint8_t x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; +x_122 = lean_ctor_get(x_1, 0); +x_123 = lean_ctor_get(x_1, 4); +x_124 = lean_ctor_get(x_1, 6); +x_125 = l_Lean_IR_StructRC_Context_renameVar_x21(x_2, x_122); +lean_dec(x_122); +x_126 = 0; +lean_inc(x_125); +lean_inc_ref(x_2); +x_127 = l_Lean_IR_StructRC_Context_useVar(x_2, x_125, x_126, x_3); +x_128 = lean_ctor_get(x_127, 0); +lean_inc(x_128); +x_129 = lean_ctor_get(x_127, 1); +lean_inc(x_129); +lean_dec_ref(x_127); +x_130 = l_Lean_IR_StructRC_Context_renameVar_x21(x_2, x_123); +lean_dec(x_123); +lean_dec_ref(x_2); +x_131 = lean_box(12); +lean_ctor_set(x_1, 6, x_131); +lean_ctor_set(x_1, 4, x_130); +lean_ctor_set(x_1, 0, x_125); +x_132 = l_Lean_IR_StructRC_Context_addInstr(x_128, x_1); +x_1 = x_124; +x_2 = x_132; +x_3 = x_129; +goto _start; +} +else +{ +lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; uint8_t x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; +x_134 = lean_ctor_get(x_1, 0); +x_135 = lean_ctor_get(x_1, 1); +x_136 = lean_ctor_get(x_1, 2); +x_137 = lean_ctor_get(x_1, 3); +x_138 = lean_ctor_get(x_1, 4); +x_139 = lean_ctor_get(x_1, 5); +x_140 = lean_ctor_get(x_1, 6); +lean_inc(x_140); +lean_inc(x_139); +lean_inc(x_138); +lean_inc(x_137); +lean_inc(x_136); +lean_inc(x_135); +lean_inc(x_134); +lean_dec(x_1); +x_141 = l_Lean_IR_StructRC_Context_renameVar_x21(x_2, x_134); +lean_dec(x_134); +x_142 = 0; +lean_inc(x_141); +lean_inc_ref(x_2); +x_143 = l_Lean_IR_StructRC_Context_useVar(x_2, x_141, x_142, x_3); +x_144 = lean_ctor_get(x_143, 0); +lean_inc(x_144); +x_145 = lean_ctor_get(x_143, 1); +lean_inc(x_145); +lean_dec_ref(x_143); +x_146 = l_Lean_IR_StructRC_Context_renameVar_x21(x_2, x_138); +lean_dec(x_138); +lean_dec_ref(x_2); +x_147 = lean_box(12); +x_148 = lean_alloc_ctor(5, 7, 0); +lean_ctor_set(x_148, 0, x_141); +lean_ctor_set(x_148, 1, x_135); +lean_ctor_set(x_148, 2, x_136); +lean_ctor_set(x_148, 3, x_137); +lean_ctor_set(x_148, 4, x_146); +lean_ctor_set(x_148, 5, x_139); +lean_ctor_set(x_148, 6, x_147); +x_149 = l_Lean_IR_StructRC_Context_addInstr(x_144, x_148); +x_1 = x_140; +x_2 = x_149; +x_3 = x_145; +goto _start; +} +} +case 6: +{ +uint8_t x_151; +x_151 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 1); +if (x_151 == 0) +{ +lean_object* x_152; lean_object* x_153; uint8_t x_154; lean_object* x_155; lean_object* x_156; +x_152 = lean_ctor_get(x_1, 0); +lean_inc(x_152); +x_153 = lean_ctor_get(x_1, 1); +lean_inc(x_153); +x_154 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); +x_155 = lean_ctor_get(x_1, 2); +lean_inc(x_155); +lean_dec_ref(x_1); +x_156 = l_Lean_IR_StructRC_Context_renameVar(x_2, x_152); +if (lean_obj_tag(x_156) == 0) +{ +lean_object* x_157; lean_object* x_158; lean_object* x_159; +x_157 = lean_ctor_get(x_156, 0); +lean_inc(x_157); +lean_dec_ref(x_156); +x_158 = lean_nat_to_int(x_153); +x_159 = l_Lean_IR_StructRC_Context_accumulateRCDiff(x_157, x_154, x_158, x_2); +x_1 = x_155; +x_2 = x_159; +goto _start; +} +else +{ +lean_dec(x_153); +x_1 = x_155; +goto _start; +} +} +else +{ +lean_object* x_162; +x_162 = lean_ctor_get(x_1, 2); +lean_inc(x_162); +lean_dec_ref(x_1); +x_1 = x_162; +goto _start; +} +} +case 7: +{ +uint8_t x_164; +x_164 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 1); +if (x_164 == 0) +{ +lean_object* x_165; lean_object* x_166; uint8_t x_167; lean_object* x_168; lean_object* x_169; +x_165 = lean_ctor_get(x_1, 0); +lean_inc(x_165); +x_166 = lean_ctor_get(x_1, 1); +lean_inc(x_166); +x_167 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); +x_168 = lean_ctor_get(x_1, 2); +lean_inc(x_168); +lean_dec_ref(x_1); +x_169 = l_Lean_IR_StructRC_Context_renameVar(x_2, x_165); +if (lean_obj_tag(x_169) == 0) +{ +lean_object* x_170; uint8_t x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; +x_170 = lean_ctor_get(x_169, 0); +lean_inc(x_170); +lean_dec_ref(x_169); +x_171 = 1; +x_172 = lean_nat_to_int(x_166); +x_173 = lean_int_neg(x_172); +lean_dec(x_172); +lean_inc(x_170); +x_174 = l_Lean_IR_StructRC_Context_accumulateRCDiff(x_170, x_167, x_173, x_2); +x_175 = l_Lean_IR_StructRC_Context_useVar(x_174, x_170, x_171, x_3); +x_176 = lean_ctor_get(x_175, 0); +lean_inc(x_176); +x_177 = lean_ctor_get(x_175, 1); +lean_inc(x_177); +lean_dec_ref(x_175); +x_1 = x_168; +x_2 = x_176; +x_3 = x_177; +goto _start; +} +else +{ +lean_dec(x_166); +x_1 = x_168; +goto _start; +} +} +else +{ +lean_object* x_180; +x_180 = lean_ctor_get(x_1, 2); +lean_inc(x_180); +lean_dec_ref(x_1); +x_1 = x_180; +goto _start; +} +} +case 8: +{ +uint8_t x_182; +x_182 = !lean_is_exclusive(x_1); +if (x_182 == 0) +{ +lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; +x_183 = lean_ctor_get(x_1, 0); +x_184 = lean_ctor_get(x_1, 1); +x_185 = l_Lean_IR_StructRC_Context_remove(x_2, x_183); +x_186 = lean_box(12); +lean_ctor_set(x_1, 1, x_186); +x_187 = l_Lean_IR_StructRC_Context_addInstr(x_185, x_1); +x_1 = x_184; +x_2 = x_187; +goto _start; +} +else +{ +lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; +x_189 = lean_ctor_get(x_1, 0); +x_190 = lean_ctor_get(x_1, 1); +lean_inc(x_190); +lean_inc(x_189); +lean_dec(x_1); +x_191 = l_Lean_IR_StructRC_Context_remove(x_2, x_189); +x_192 = lean_box(12); +x_193 = lean_alloc_ctor(8, 2, 0); +lean_ctor_set(x_193, 0, x_189); +lean_ctor_set(x_193, 1, x_192); +x_194 = l_Lean_IR_StructRC_Context_addInstr(x_191, x_193); +x_1 = x_190; +x_2 = x_194; +goto _start; +} +} +case 9: +{ +lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_208; lean_object* x_209; lean_object* x_225; +x_196 = lean_ctor_get(x_1, 0); +lean_inc(x_196); +x_197 = lean_ctor_get(x_1, 1); +lean_inc(x_197); +x_198 = lean_ctor_get(x_1, 2); +lean_inc(x_198); +x_199 = lean_ctor_get(x_1, 3); +lean_inc_ref(x_199); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_200 = x_1; +} else { + lean_dec_ref(x_1); + x_200 = lean_box(0); +} +x_201 = lean_ctor_get(x_2, 0); +x_202 = lean_ctor_get(x_2, 1); +x_203 = lean_ctor_get(x_2, 2); +x_208 = l_Lean_IR_StructRC_Context_renameVar_x21(x_2, x_197); +lean_dec(x_197); +x_225 = l_Std_DTreeMap_Internal_Impl_Const_get_x3f___at___00Lean_IR_StructRC_Context_renameVar_x21_spec__0___redArg(x_201, x_208); +if (lean_obj_tag(x_225) == 1) +{ +lean_object* x_226; +x_226 = lean_ctor_get(x_225, 0); +lean_inc(x_226); +lean_dec_ref(x_225); +if (lean_obj_tag(x_226) == 2) +{ +lean_object* x_227; lean_object* x_228; lean_object* x_229; size_t x_230; size_t x_231; lean_object* x_232; lean_object* x_233; +lean_dec(x_208); +lean_dec(x_200); +lean_dec(x_198); +lean_dec(x_196); +x_227 = lean_ctor_get(x_226, 0); +lean_inc(x_227); +lean_dec_ref(x_226); +x_228 = lean_box(0); +x_229 = l_Lean_IR_StructRC_visitFnBody___closed__0; +x_230 = lean_array_size(x_199); +x_231 = 0; +x_232 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_StructRC_visitFnBody_spec__1(x_228, x_227, x_229, x_199, x_230, x_231, x_229); +lean_dec_ref(x_199); +lean_dec(x_227); +x_233 = lean_ctor_get(x_232, 0); +lean_inc(x_233); +lean_dec_ref(x_232); +if (lean_obj_tag(x_233) == 0) +{ +lean_inc_ref(x_202); +lean_dec_ref(x_2); +goto block_207; +} +else +{ +lean_object* x_234; +x_234 = lean_ctor_get(x_233, 0); +lean_inc(x_234); +lean_dec_ref(x_233); +if (lean_obj_tag(x_234) == 0) +{ +lean_inc_ref(x_202); +lean_dec_ref(x_2); +goto block_207; +} +else +{ +lean_object* x_235; +x_235 = lean_ctor_get(x_234, 0); +lean_inc(x_235); +lean_dec_ref(x_234); +x_1 = x_235; +goto _start; +} +} +} +else +{ +lean_dec(x_226); +lean_inc(x_203); +lean_inc_ref(x_202); +lean_inc(x_201); +lean_dec_ref(x_2); +x_209 = x_3; +goto block_224; +} +} +else +{ +lean_dec(x_225); +lean_inc(x_203); +lean_inc_ref(x_202); +lean_inc(x_201); +lean_dec_ref(x_2); +x_209 = x_3; +goto block_224; +} +block_207: +{ +lean_object* x_204; lean_object* x_205; lean_object* x_206; +x_204 = lean_box(12); +x_205 = l_Lean_IR_reshape(x_202, x_204); +x_206 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_206, 0, x_205); +lean_ctor_set(x_206, 1, x_3); +return x_206; +} +block_224: +{ +lean_object* x_210; lean_object* x_211; size_t x_212; size_t x_213; lean_object* x_214; uint8_t x_215; +x_210 = l_Lean_IR_StructRC_Context_resetRC___closed__0; +x_211 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_211, 0, x_201); +lean_ctor_set(x_211, 1, x_210); +lean_ctor_set(x_211, 2, x_203); +x_212 = lean_array_size(x_199); +x_213 = 0; +lean_inc(x_208); +x_214 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_StructRC_visitFnBody_spec__0(x_211, x_208, x_212, x_213, x_199, x_209); +x_215 = !lean_is_exclusive(x_214); +if (x_215 == 0) +{ +lean_object* x_216; lean_object* x_217; lean_object* x_218; +x_216 = lean_ctor_get(x_214, 0); +if (lean_is_scalar(x_200)) { + x_217 = lean_alloc_ctor(9, 4, 0); +} else { + x_217 = x_200; +} +lean_ctor_set(x_217, 0, x_196); +lean_ctor_set(x_217, 1, x_208); +lean_ctor_set(x_217, 2, x_198); +lean_ctor_set(x_217, 3, x_216); +x_218 = l_Lean_IR_reshape(x_202, x_217); +lean_ctor_set(x_214, 0, x_218); +return x_214; +} +else +{ +lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; +x_219 = lean_ctor_get(x_214, 0); +x_220 = lean_ctor_get(x_214, 1); +lean_inc(x_220); +lean_inc(x_219); +lean_dec(x_214); +if (lean_is_scalar(x_200)) { + x_221 = lean_alloc_ctor(9, 4, 0); +} else { + x_221 = x_200; +} +lean_ctor_set(x_221, 0, x_196); +lean_ctor_set(x_221, 1, x_208); +lean_ctor_set(x_221, 2, x_198); +lean_ctor_set(x_221, 3, x_219); +x_222 = l_Lean_IR_reshape(x_202, x_221); +x_223 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_223, 0, x_222); +lean_ctor_set(x_223, 1, x_220); +return x_223; +} +} +} +case 10: +{ +uint8_t x_237; +x_237 = !lean_is_exclusive(x_1); +if (x_237 == 0) +{ +lean_object* x_238; lean_object* x_239; uint8_t x_240; +x_238 = lean_ctor_get(x_1, 0); +lean_inc_ref(x_2); +x_239 = l_Lean_IR_StructRC_Context_finish(x_2, x_3); +x_240 = !lean_is_exclusive(x_239); +if (x_240 == 0) +{ +lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; +x_241 = lean_ctor_get(x_239, 0); +x_242 = lean_ctor_get(x_241, 1); +lean_inc_ref(x_242); +lean_dec(x_241); +x_243 = l_Lean_IR_StructRC_Context_renameArg(x_2, x_238); +lean_dec(x_238); +lean_dec_ref(x_2); +lean_ctor_set(x_1, 0, x_243); +x_244 = l_Lean_IR_reshape(x_242, x_1); +lean_ctor_set(x_239, 0, x_244); +return x_239; +} +else +{ +lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; +x_245 = lean_ctor_get(x_239, 0); +x_246 = lean_ctor_get(x_239, 1); +lean_inc(x_246); +lean_inc(x_245); +lean_dec(x_239); +x_247 = lean_ctor_get(x_245, 1); +lean_inc_ref(x_247); +lean_dec(x_245); +x_248 = l_Lean_IR_StructRC_Context_renameArg(x_2, x_238); +lean_dec(x_238); +lean_dec_ref(x_2); +lean_ctor_set(x_1, 0, x_248); +x_249 = l_Lean_IR_reshape(x_247, x_1); +x_250 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_250, 0, x_249); +lean_ctor_set(x_250, 1, x_246); +return x_250; +} +} +else +{ +lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; +x_251 = lean_ctor_get(x_1, 0); +lean_inc(x_251); +lean_dec(x_1); +lean_inc_ref(x_2); +x_252 = l_Lean_IR_StructRC_Context_finish(x_2, x_3); +x_253 = lean_ctor_get(x_252, 0); +lean_inc(x_253); +x_254 = lean_ctor_get(x_252, 1); +lean_inc(x_254); +if (lean_is_exclusive(x_252)) { + lean_ctor_release(x_252, 0); + lean_ctor_release(x_252, 1); + x_255 = x_252; +} else { + lean_dec_ref(x_252); + x_255 = lean_box(0); +} +x_256 = lean_ctor_get(x_253, 1); +lean_inc_ref(x_256); +lean_dec(x_253); +x_257 = l_Lean_IR_StructRC_Context_renameArg(x_2, x_251); +lean_dec(x_251); +lean_dec_ref(x_2); +x_258 = lean_alloc_ctor(10, 1, 0); +lean_ctor_set(x_258, 0, x_257); +x_259 = l_Lean_IR_reshape(x_256, x_258); +if (lean_is_scalar(x_255)) { + x_260 = lean_alloc_ctor(0, 2, 0); +} else { + x_260 = x_255; +} +lean_ctor_set(x_260, 0, x_259); +lean_ctor_set(x_260, 1, x_254); +return x_260; +} +} +case 11: +{ +uint8_t x_261; +x_261 = !lean_is_exclusive(x_1); +if (x_261 == 0) +{ +lean_object* x_262; lean_object* x_263; uint8_t x_264; +x_262 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_2); +x_263 = l_Lean_IR_StructRC_Context_finish(x_2, x_3); +x_264 = !lean_is_exclusive(x_263); +if (x_264 == 0) +{ +lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; +x_265 = lean_ctor_get(x_263, 0); +x_266 = lean_ctor_get(x_265, 1); +lean_inc_ref(x_266); +lean_dec(x_265); +x_267 = l_Lean_IR_StructRC_Context_renameArgs(x_2, x_262); +lean_dec_ref(x_2); +lean_ctor_set(x_1, 1, x_267); +x_268 = l_Lean_IR_reshape(x_266, x_1); +lean_ctor_set(x_263, 0, x_268); +return x_263; +} +else +{ +lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; +x_269 = lean_ctor_get(x_263, 0); +x_270 = lean_ctor_get(x_263, 1); +lean_inc(x_270); +lean_inc(x_269); +lean_dec(x_263); +x_271 = lean_ctor_get(x_269, 1); +lean_inc_ref(x_271); +lean_dec(x_269); +x_272 = l_Lean_IR_StructRC_Context_renameArgs(x_2, x_262); +lean_dec_ref(x_2); +lean_ctor_set(x_1, 1, x_272); +x_273 = l_Lean_IR_reshape(x_271, x_1); +x_274 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_274, 0, x_273); +lean_ctor_set(x_274, 1, x_270); +return x_274; +} +} +else +{ +lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; +x_275 = lean_ctor_get(x_1, 0); +x_276 = lean_ctor_get(x_1, 1); +lean_inc(x_276); +lean_inc(x_275); +lean_dec(x_1); +lean_inc_ref(x_2); +x_277 = l_Lean_IR_StructRC_Context_finish(x_2, x_3); +x_278 = lean_ctor_get(x_277, 0); +lean_inc(x_278); +x_279 = lean_ctor_get(x_277, 1); +lean_inc(x_279); +if (lean_is_exclusive(x_277)) { + lean_ctor_release(x_277, 0); + lean_ctor_release(x_277, 1); + x_280 = x_277; +} else { + lean_dec_ref(x_277); + x_280 = lean_box(0); +} +x_281 = lean_ctor_get(x_278, 1); +lean_inc_ref(x_281); +lean_dec(x_278); +x_282 = l_Lean_IR_StructRC_Context_renameArgs(x_2, x_276); +lean_dec_ref(x_2); +x_283 = lean_alloc_ctor(11, 2, 0); +lean_ctor_set(x_283, 0, x_275); +lean_ctor_set(x_283, 1, x_282); +x_284 = l_Lean_IR_reshape(x_281, x_283); +if (lean_is_scalar(x_280)) { + x_285 = lean_alloc_ctor(0, 2, 0); +} else { + x_285 = x_280; +} +lean_ctor_set(x_285, 0, x_284); +lean_ctor_set(x_285, 1, x_279); +return x_285; +} +} +default: +{ +lean_object* x_286; lean_object* x_287; lean_object* x_288; +x_286 = lean_ctor_get(x_2, 1); +lean_inc_ref(x_286); +lean_dec_ref(x_2); +x_287 = l_Lean_IR_reshape(x_286, x_1); +x_288 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_288, 0, x_287); +lean_ctor_set(x_288, 1, x_3); +return x_288; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_StructRC_visitFnBody_spec__0(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_4, x_3); +if (x_7 == 0) +{ +lean_object* x_8; +lean_dec(x_2); +lean_dec_ref(x_1); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_5); +lean_ctor_set(x_8, 1, x_6); +return x_8; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_9 = lean_array_uget(x_5, x_4); +x_10 = lean_unsigned_to_nat(0u); +x_11 = lean_array_uset(x_5, x_4, x_10); +if (lean_obj_tag(x_9) == 0) +{ +uint8_t x_19; +x_19 = !lean_is_exclusive(x_9); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_20 = lean_ctor_get(x_9, 0); +x_21 = lean_ctor_get(x_9, 1); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_inc(x_2); +lean_inc_ref(x_1); +x_23 = l_Lean_IR_StructRC_Context_addCtorKnowledge(x_1, x_2, x_22); +x_24 = l_Lean_IR_StructRC_visitFnBody(x_21, x_23, x_6); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec_ref(x_24); +lean_ctor_set(x_9, 1, x_25); +x_12 = x_9; +x_13 = x_26; +goto block_18; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_27 = lean_ctor_get(x_9, 0); +x_28 = lean_ctor_get(x_9, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_9); +x_29 = lean_ctor_get(x_27, 1); +lean_inc(x_29); +lean_inc(x_2); +lean_inc_ref(x_1); +x_30 = l_Lean_IR_StructRC_Context_addCtorKnowledge(x_1, x_2, x_29); +x_31 = l_Lean_IR_StructRC_visitFnBody(x_28, x_30, x_6); +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec_ref(x_31); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_27); +lean_ctor_set(x_34, 1, x_32); +x_12 = x_34; +x_13 = x_33; +goto block_18; +} +} +else +{ +uint8_t x_35; +x_35 = !lean_is_exclusive(x_9); +if (x_35 == 0) +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_36 = lean_ctor_get(x_9, 0); +lean_inc_ref(x_1); +x_37 = l_Lean_IR_StructRC_visitFnBody(x_36, x_1, x_6); +x_38 = lean_ctor_get(x_37, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_37, 1); +lean_inc(x_39); +lean_dec_ref(x_37); +lean_ctor_set(x_9, 0, x_38); +x_12 = x_9; +x_13 = x_39; +goto block_18; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_40 = lean_ctor_get(x_9, 0); +lean_inc(x_40); +lean_dec(x_9); +lean_inc_ref(x_1); +x_41 = l_Lean_IR_StructRC_visitFnBody(x_40, x_1, x_6); +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +lean_dec_ref(x_41); +x_44 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_44, 0, x_42); +x_12 = x_44; +x_13 = x_43; +goto block_18; +} +} +block_18: +{ +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = 1; +x_15 = lean_usize_add(x_4, x_14); +x_16 = lean_array_uset(x_11, x_4, x_12); +x_4 = x_15; +x_5 = x_16; +x_6 = x_13; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_StructRC_visitFnBody_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +size_t x_8; size_t x_9; lean_object* x_10; +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_10 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_StructRC_visitFnBody_spec__1(x_1, x_2, x_3, x_4, x_8, x_9, x_7); +lean_dec_ref(x_7); +lean_dec_ref(x_4); +lean_dec_ref(x_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_StructRC_visitFnBody_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_8 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_9 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_StructRC_visitFnBody_spec__0(x_1, x_2, x_7, x_8, x_5, x_6); +return x_9; +} +} +static lean_object* _init_l_Lean_IR_StructRC_visitDecl___closed__0() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_StructRC_Context_resetRC___closed__0; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 2, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_StructRC_visitDecl(lean_object* x_1) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_3 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +x_5 = lean_ctor_get(x_1, 3); +lean_inc(x_5); +x_6 = lean_ctor_get(x_1, 4); +lean_inc(x_6); +x_7 = l_Lean_IR_StructRC_visitDecl___closed__0; +x_8 = l_Lean_IR_StructRC_Context_addParams(x_3, x_7); +x_9 = l_Lean_IR_Decl_maxIndex(x_1); +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_add(x_9, x_10); +lean_dec(x_9); +x_12 = l_Lean_IR_StructRC_visitFnBody(x_5, x_8, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +lean_dec_ref(x_12); +x_14 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_14, 0, x_2); +lean_ctor_set(x_14, 1, x_3); +lean_ctor_set(x_14, 2, x_4); +lean_ctor_set(x_14, 3, x_13); +lean_ctor_set(x_14, 4, x_6); +return x_14; +} +else +{ +return x_1; +} +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("compiler", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ir", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__2_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("struct_rc", 9, 9); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__3_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__2_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_3 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_4 = l_Lean_Name_mkStr3(x_3, x_2, x_1); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__4_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_private", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__5_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__4_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_2 = lean_box(0); +x_3 = l_Lean_Name_str___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__6_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__7_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__6_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__5_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_3 = l_Lean_Name_str___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__8_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Compiler", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__9_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__8_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__7_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_3 = l_Lean_Name_str___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__10_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("IR", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__11_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__10_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__9_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_3 = l_Lean_Name_str___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__12_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("StructRC", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__13_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__12_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__11_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_3 = l_Lean_Name_str___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__14_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(0u); +x_2 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__13_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_3 = l_Lean_Name_num___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__15_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__6_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__14_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_3 = l_Lean_Name_str___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__16_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__10_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__15_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_3 = l_Lean_Name_str___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__17_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("initFn", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__18_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__17_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__16_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_3 = l_Lean_Name_str___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__19_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_@", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__20_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__19_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__18_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_3 = l_Lean_Name_str___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__21_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__6_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__20_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_3 = l_Lean_Name_str___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__22_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__8_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__21_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_3 = l_Lean_Name_str___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__23_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__10_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__22_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_3 = l_Lean_Name_str___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__24_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__12_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__23_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_3 = l_Lean_Name_str___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__25_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(2818627513u); +x_2 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__24_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_3 = l_Lean_Name_num___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__26_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_hygCtx", 7, 7); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__27_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__26_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__25_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_3 = l_Lean_Name_str___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__28_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_hyg", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__29_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__28_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_2 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__27_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_3 = l_Lean_Name_str___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__30_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(2u); +x_2 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__29_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_3 = l_Lean_Name_num___override(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; +x_2 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__3_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_3 = 1; +x_4 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__30_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_; +x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2____boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +return x_2; +} +} +lean_object* initialize_Lean_Compiler_IR_FreeVars(uint8_t builtin); +lean_object* initialize_Lean_Compiler_IR_Format(uint8_t builtin); +lean_object* initialize_Std_Data_TreeMap_AdditionalOperations(uint8_t builtin); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Compiler_IR_StructRC(uint8_t builtin) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Lean_Compiler_IR_FreeVars(builtin); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Compiler_IR_Format(builtin); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Data_TreeMap_AdditionalOperations(builtin); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_IR_StructRC_instInhabitedProjEntry_default = _init_l_Lean_IR_StructRC_instInhabitedProjEntry_default(); +lean_mark_persistent(l_Lean_IR_StructRC_instInhabitedProjEntry_default); +l_Lean_IR_StructRC_instInhabitedProjEntry = _init_l_Lean_IR_StructRC_instInhabitedProjEntry(); +lean_mark_persistent(l_Lean_IR_StructRC_instInhabitedProjEntry); +l_Lean_IR_StructRC_instReprProjEntry_repr___closed__0 = _init_l_Lean_IR_StructRC_instReprProjEntry_repr___closed__0(); +lean_mark_persistent(l_Lean_IR_StructRC_instReprProjEntry_repr___closed__0); +l_Lean_IR_StructRC_instReprProjEntry_repr___closed__1 = _init_l_Lean_IR_StructRC_instReprProjEntry_repr___closed__1(); +lean_mark_persistent(l_Lean_IR_StructRC_instReprProjEntry_repr___closed__1); +l_Lean_IR_StructRC_instReprProjEntry_repr___closed__2 = _init_l_Lean_IR_StructRC_instReprProjEntry_repr___closed__2(); +lean_mark_persistent(l_Lean_IR_StructRC_instReprProjEntry_repr___closed__2); +l_Lean_IR_StructRC_instReprProjEntry_repr___closed__3 = _init_l_Lean_IR_StructRC_instReprProjEntry_repr___closed__3(); +lean_mark_persistent(l_Lean_IR_StructRC_instReprProjEntry_repr___closed__3); +l_Lean_IR_StructRC_instReprProjEntry_repr___closed__4 = _init_l_Lean_IR_StructRC_instReprProjEntry_repr___closed__4(); +lean_mark_persistent(l_Lean_IR_StructRC_instReprProjEntry_repr___closed__4); +l_Lean_IR_StructRC_instReprProjEntry_repr___closed__5 = _init_l_Lean_IR_StructRC_instReprProjEntry_repr___closed__5(); +lean_mark_persistent(l_Lean_IR_StructRC_instReprProjEntry_repr___closed__5); +l_Lean_IR_StructRC_instReprProjEntry_repr___closed__6 = _init_l_Lean_IR_StructRC_instReprProjEntry_repr___closed__6(); +lean_mark_persistent(l_Lean_IR_StructRC_instReprProjEntry_repr___closed__6); +l_Lean_IR_StructRC_instReprProjEntry_repr___closed__7 = _init_l_Lean_IR_StructRC_instReprProjEntry_repr___closed__7(); +lean_mark_persistent(l_Lean_IR_StructRC_instReprProjEntry_repr___closed__7); +l_Lean_IR_StructRC_instReprProjEntry_repr___closed__8 = _init_l_Lean_IR_StructRC_instReprProjEntry_repr___closed__8(); +lean_mark_persistent(l_Lean_IR_StructRC_instReprProjEntry_repr___closed__8); +l_Lean_IR_StructRC_instReprProjEntry_repr___closed__9 = _init_l_Lean_IR_StructRC_instReprProjEntry_repr___closed__9(); +lean_mark_persistent(l_Lean_IR_StructRC_instReprProjEntry_repr___closed__9); +l_Lean_IR_StructRC_instReprProjEntry_repr___closed__10 = _init_l_Lean_IR_StructRC_instReprProjEntry_repr___closed__10(); +lean_mark_persistent(l_Lean_IR_StructRC_instReprProjEntry_repr___closed__10); +l_Lean_IR_StructRC_instReprProjEntry___closed__0 = _init_l_Lean_IR_StructRC_instReprProjEntry___closed__0(); +lean_mark_persistent(l_Lean_IR_StructRC_instReprProjEntry___closed__0); +l_Lean_IR_StructRC_instReprProjEntry = _init_l_Lean_IR_StructRC_instReprProjEntry(); +lean_mark_persistent(l_Lean_IR_StructRC_instReprProjEntry); +l_Lean_IR_StructRC_instInhabitedEntry_default = _init_l_Lean_IR_StructRC_instInhabitedEntry_default(); +lean_mark_persistent(l_Lean_IR_StructRC_instInhabitedEntry_default); +l_Lean_IR_StructRC_instInhabitedEntry = _init_l_Lean_IR_StructRC_instInhabitedEntry(); +lean_mark_persistent(l_Lean_IR_StructRC_instInhabitedEntry); +l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__4 = _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__4(); +lean_mark_persistent(l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__4); +l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__5 = _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__5(); +lean_mark_persistent(l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__5); +l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__1 = _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__1(); +lean_mark_persistent(l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__1); +l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__2 = _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__2(); +lean_mark_persistent(l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__2); +l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__3 = _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__3(); +lean_mark_persistent(l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__3); +l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__6 = _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__6(); +lean_mark_persistent(l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__6); +l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__7 = _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__7(); +lean_mark_persistent(l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__7); +l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__1 = _init_l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__1(); +lean_mark_persistent(l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__1); +l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__2 = _init_l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__2(); +lean_mark_persistent(l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__2); +l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__3 = _init_l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__3(); +lean_mark_persistent(l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__3); +l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__0 = _init_l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__0(); +lean_mark_persistent(l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__0); +l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__5 = _init_l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__5(); +lean_mark_persistent(l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__5); +l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__6 = _init_l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__6(); +lean_mark_persistent(l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__6); +l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__7 = _init_l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__7(); +lean_mark_persistent(l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__7); +l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__4 = _init_l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__4(); +lean_mark_persistent(l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__4); +l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__8 = _init_l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__8(); +lean_mark_persistent(l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__8); +l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__9 = _init_l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__9(); +lean_mark_persistent(l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__9); +l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__10 = _init_l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__10(); +lean_mark_persistent(l_Array_Array_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__0___closed__10); +l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__8 = _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__8(); +lean_mark_persistent(l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__8); +l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__9 = _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__9(); +lean_mark_persistent(l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__9); +l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__10 = _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__10(); +lean_mark_persistent(l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__10); +l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__11 = _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__11(); +lean_mark_persistent(l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__11); +l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__0 = _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__0(); +lean_mark_persistent(l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__0); +l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__13 = _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__13(); +lean_mark_persistent(l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__13); +l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__14 = _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__14(); +lean_mark_persistent(l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__14); +l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__15 = _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__15(); +lean_mark_persistent(l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__15); +l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__12 = _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__12(); +lean_mark_persistent(l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__12); +l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__16 = _init_l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__16(); +lean_mark_persistent(l_instReprVector_repr___at___00Lean_IR_StructRC_instReprEntry_repr_spec__1___redArg___closed__16); +l_Lean_IR_StructRC_instReprEntry_repr___closed__0 = _init_l_Lean_IR_StructRC_instReprEntry_repr___closed__0(); +lean_mark_persistent(l_Lean_IR_StructRC_instReprEntry_repr___closed__0); +l_Lean_IR_StructRC_instReprEntry_repr___closed__1 = _init_l_Lean_IR_StructRC_instReprEntry_repr___closed__1(); +lean_mark_persistent(l_Lean_IR_StructRC_instReprEntry_repr___closed__1); +l_Lean_IR_StructRC_instReprEntry_repr___closed__2 = _init_l_Lean_IR_StructRC_instReprEntry_repr___closed__2(); +lean_mark_persistent(l_Lean_IR_StructRC_instReprEntry_repr___closed__2); +l_Lean_IR_StructRC_instReprEntry_repr___closed__3 = _init_l_Lean_IR_StructRC_instReprEntry_repr___closed__3(); +lean_mark_persistent(l_Lean_IR_StructRC_instReprEntry_repr___closed__3); +l_Lean_IR_StructRC_instReprEntry_repr___closed__4 = _init_l_Lean_IR_StructRC_instReprEntry_repr___closed__4(); +lean_mark_persistent(l_Lean_IR_StructRC_instReprEntry_repr___closed__4); +l_Lean_IR_StructRC_instReprEntry_repr___closed__5 = _init_l_Lean_IR_StructRC_instReprEntry_repr___closed__5(); +lean_mark_persistent(l_Lean_IR_StructRC_instReprEntry_repr___closed__5); +l_Lean_IR_StructRC_instReprEntry_repr___closed__6 = _init_l_Lean_IR_StructRC_instReprEntry_repr___closed__6(); +lean_mark_persistent(l_Lean_IR_StructRC_instReprEntry_repr___closed__6); +l_Lean_IR_StructRC_instReprEntry_repr___closed__7 = _init_l_Lean_IR_StructRC_instReprEntry_repr___closed__7(); +lean_mark_persistent(l_Lean_IR_StructRC_instReprEntry_repr___closed__7); +l_Lean_IR_StructRC_instReprEntry_repr___closed__8 = _init_l_Lean_IR_StructRC_instReprEntry_repr___closed__8(); +lean_mark_persistent(l_Lean_IR_StructRC_instReprEntry_repr___closed__8); +l_Lean_IR_StructRC_instReprEntry_repr___closed__9 = _init_l_Lean_IR_StructRC_instReprEntry_repr___closed__9(); +lean_mark_persistent(l_Lean_IR_StructRC_instReprEntry_repr___closed__9); +l_Lean_IR_StructRC_instReprEntry_repr___closed__10 = _init_l_Lean_IR_StructRC_instReprEntry_repr___closed__10(); +lean_mark_persistent(l_Lean_IR_StructRC_instReprEntry_repr___closed__10); +l_Lean_IR_StructRC_instReprEntry___closed__0 = _init_l_Lean_IR_StructRC_instReprEntry___closed__0(); +lean_mark_persistent(l_Lean_IR_StructRC_instReprEntry___closed__0); +l_Lean_IR_StructRC_instReprEntry = _init_l_Lean_IR_StructRC_instReprEntry(); +lean_mark_persistent(l_Lean_IR_StructRC_instReprEntry); +l_Lean_IR_StructRC_Context_renameVar_x21___closed__0 = _init_l_Lean_IR_StructRC_Context_renameVar_x21___closed__0(); +lean_mark_persistent(l_Lean_IR_StructRC_Context_renameVar_x21___closed__0); +l_Lean_IR_StructRC_Context_renameVar_x21___closed__1 = _init_l_Lean_IR_StructRC_Context_renameVar_x21___closed__1(); +lean_mark_persistent(l_Lean_IR_StructRC_Context_renameVar_x21___closed__1); +l_Lean_IR_StructRC_Context_renameVar_x21___closed__2 = _init_l_Lean_IR_StructRC_Context_renameVar_x21___closed__2(); +lean_mark_persistent(l_Lean_IR_StructRC_Context_renameVar_x21___closed__2); +l_Lean_IR_StructRC_Context_renameVar_x21___closed__3 = _init_l_Lean_IR_StructRC_Context_renameVar_x21___closed__3(); +lean_mark_persistent(l_Lean_IR_StructRC_Context_renameVar_x21___closed__3); +l_Lean_IR_StructRC_Entry_ofStruct___closed__0 = _init_l_Lean_IR_StructRC_Entry_ofStruct___closed__0(); +lean_mark_persistent(l_Lean_IR_StructRC_Entry_ofStruct___closed__0); +l_Lean_IR_StructRC_Entry_ofStruct___closed__1 = _init_l_Lean_IR_StructRC_Entry_ofStruct___closed__1(); +lean_mark_persistent(l_Lean_IR_StructRC_Entry_ofStruct___closed__1); +l_Lean_IR_StructRC_Entry_ofStruct___closed__2 = _init_l_Lean_IR_StructRC_Entry_ofStruct___closed__2(); +lean_mark_persistent(l_Lean_IR_StructRC_Entry_ofStruct___closed__2); +l_Lean_IR_StructRC_Context_addVarBasic___closed__0 = _init_l_Lean_IR_StructRC_Context_addVarBasic___closed__0(); +lean_mark_persistent(l_Lean_IR_StructRC_Context_addVarBasic___closed__0); +l_Lean_IR_StructRC_Context_addProjInfo___closed__0 = _init_l_Lean_IR_StructRC_Context_addProjInfo___closed__0(); +lean_mark_persistent(l_Lean_IR_StructRC_Context_addProjInfo___closed__0); +l_Lean_IR_StructRC_Context_addProjInfo___closed__1 = _init_l_Lean_IR_StructRC_Context_addProjInfo___closed__1(); +lean_mark_persistent(l_Lean_IR_StructRC_Context_addProjInfo___closed__1); +l_Lean_IR_StructRC_Context_addProjInfo___closed__2 = _init_l_Lean_IR_StructRC_Context_addProjInfo___closed__2(); +lean_mark_persistent(l_Lean_IR_StructRC_Context_addProjInfo___closed__2); +l_Lean_IR_StructRC_Context_addProjInfo___closed__3 = _init_l_Lean_IR_StructRC_Context_addProjInfo___closed__3(); +lean_mark_persistent(l_Lean_IR_StructRC_Context_addProjInfo___closed__3); +l_Lean_IR_StructRC_Context_addProjInfo___closed__4 = _init_l_Lean_IR_StructRC_Context_addProjInfo___closed__4(); +lean_mark_persistent(l_Lean_IR_StructRC_Context_addProjInfo___closed__4); +l_Lean_IR_StructRC_Context_addProjInfo___closed__5 = _init_l_Lean_IR_StructRC_Context_addProjInfo___closed__5(); +lean_mark_persistent(l_Lean_IR_StructRC_Context_addProjInfo___closed__5); +l_Lean_IR_StructRC_Context_addProjInfo___closed__6 = _init_l_Lean_IR_StructRC_Context_addProjInfo___closed__6(); +lean_mark_persistent(l_Lean_IR_StructRC_Context_addProjInfo___closed__6); +l_Lean_IR_StructRC_Context_addProjInfo___closed__7 = _init_l_Lean_IR_StructRC_Context_addProjInfo___closed__7(); +lean_mark_persistent(l_Lean_IR_StructRC_Context_addProjInfo___closed__7); +l_Lean_IR_StructRC_Context_addProjInfo___closed__8 = _init_l_Lean_IR_StructRC_Context_addProjInfo___closed__8(); +lean_mark_persistent(l_Lean_IR_StructRC_Context_addProjInfo___closed__8); +l_Lean_IR_StructRC_Context_addProjInfo___closed__9 = _init_l_Lean_IR_StructRC_Context_addProjInfo___closed__9(); +lean_mark_persistent(l_Lean_IR_StructRC_Context_addProjInfo___closed__9); +l_Lean_IR_StructRC_Context_addProjInfo___closed__10 = _init_l_Lean_IR_StructRC_Context_addProjInfo___closed__10(); +lean_mark_persistent(l_Lean_IR_StructRC_Context_addProjInfo___closed__10); +l_Lean_IR_StructRC_Context_addProjInfo___closed__11 = _init_l_Lean_IR_StructRC_Context_addProjInfo___closed__11(); +lean_mark_persistent(l_Lean_IR_StructRC_Context_addProjInfo___closed__11); +l_Lean_IR_StructRC_Context_accumulateRCDiff___closed__0 = _init_l_Lean_IR_StructRC_Context_accumulateRCDiff___closed__0(); +lean_mark_persistent(l_Lean_IR_StructRC_Context_accumulateRCDiff___closed__0); +l_Lean_IR_StructRC_Context_resetRC___closed__0 = _init_l_Lean_IR_StructRC_Context_resetRC___closed__0(); +lean_mark_persistent(l_Lean_IR_StructRC_Context_resetRC___closed__0); +l_panic___at___00Lean_IR_StructRC_atConstructorIndex_spec__0___closed__0 = _init_l_panic___at___00Lean_IR_StructRC_atConstructorIndex_spec__0___closed__0(); +lean_mark_persistent(l_panic___at___00Lean_IR_StructRC_atConstructorIndex_spec__0___closed__0); +l_Lean_IR_StructRC_atConstructorIndex___closed__0 = _init_l_Lean_IR_StructRC_atConstructorIndex___closed__0(); +lean_mark_persistent(l_Lean_IR_StructRC_atConstructorIndex___closed__0); +l_Lean_IR_StructRC_atConstructorIndex___closed__1 = _init_l_Lean_IR_StructRC_atConstructorIndex___closed__1(); +lean_mark_persistent(l_Lean_IR_StructRC_atConstructorIndex___closed__1); +l_Lean_IR_StructRC_atConstructorIndex___closed__2 = _init_l_Lean_IR_StructRC_atConstructorIndex___closed__2(); +lean_mark_persistent(l_Lean_IR_StructRC_atConstructorIndex___closed__2); +l_Lean_IR_StructRC_visitExpr___lam__0___closed__0 = _init_l_Lean_IR_StructRC_visitExpr___lam__0___closed__0(); +lean_mark_persistent(l_Lean_IR_StructRC_visitExpr___lam__0___closed__0); +l_Lean_IR_StructRC_visitExpr___closed__0 = _init_l_Lean_IR_StructRC_visitExpr___closed__0(); +lean_mark_persistent(l_Lean_IR_StructRC_visitExpr___closed__0); +l_Lean_IR_StructRC_visitExpr___closed__1 = _init_l_Lean_IR_StructRC_visitExpr___closed__1(); +lean_mark_persistent(l_Lean_IR_StructRC_visitExpr___closed__1); +l_Lean_IR_StructRC_visitFnBody___closed__0 = _init_l_Lean_IR_StructRC_visitFnBody___closed__0(); +lean_mark_persistent(l_Lean_IR_StructRC_visitFnBody___closed__0); +l_Lean_IR_StructRC_visitDecl___closed__0 = _init_l_Lean_IR_StructRC_visitDecl___closed__0(); +lean_mark_persistent(l_Lean_IR_StructRC_visitDecl___closed__0); +l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__0_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__1_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__2_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__2_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__2_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__3_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__3_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__3_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__4_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__4_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__4_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__5_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__5_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__5_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__6_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__6_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__6_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__7_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__7_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__7_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__8_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__8_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__8_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__9_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__9_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__9_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__10_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__10_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__10_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__11_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__11_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__11_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__12_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__12_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__12_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__13_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__13_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__13_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__14_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__14_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__14_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__15_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__15_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__15_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__16_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__16_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__16_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__17_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__17_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__17_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__18_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__18_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__18_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__19_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__19_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__19_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__20_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__20_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__20_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__21_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__21_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__21_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__22_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__22_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__22_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__23_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__23_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__23_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__24_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__24_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__24_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__25_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__25_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__25_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__26_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__26_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__26_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__27_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__27_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__27_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__28_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__28_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__28_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__29_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__29_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__29_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_); +l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__30_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_ = _init_l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__30_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn___closed__30_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_); +if (builtin) {res = l___private_Lean_Compiler_IR_StructRC_0__Lean_IR_initFn_00___x40_Lean_Compiler_IR_StructRC_2818627513____hygCtx___hyg_2_(); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Compiler/IR/ToIR.c b/stage0/stdlib/Lean/Compiler/IR/ToIR.c index 629970313875..722bb3e958f5 100644 --- a/stage0/stdlib/Lean/Compiler/IR/ToIR.c +++ b/stage0/stdlib/Lean/Compiler/IR/ToIR.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Compiler.IR.ToIR -// Imports: public import Lean.Compiler.IR.CompilerM public import Lean.Compiler.IR.ToIRType +// Imports: public import Lean.Compiler.IR.CompilerM public import Lean.Compiler.IR.ToIRType import Lean.Compiler.IR.UnboxResult import Lean.Compiler.IR.Boxing #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -47,6 +47,7 @@ static lean_object* l_Lean_IR_ToIR_instMonadFVarSubstMFalse___closed__6; static lean_object* l_Lean_IR_ToIR_lowerLet___closed__2; LEAN_EXPORT lean_object* l_Lean_IR_ToIR_getFVarValue(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_IR_findEnvDecl(lean_object*, lean_object*, uint8_t); +uint8_t l_Array_isEmpty___redArg(lean_object*); static lean_object* l_Lean_IR_ToIR_taggedReturnAttr___regBuiltin_Lean_IR_ToIR_taggedReturnAttr_declRange__3___closed__5; static lean_object* l_Lean_addMessageContextPartial___at___00Lean_throwNamedError___at___00Lean_IR_ToIR_lowerLet_spec__8_spec__16___closed__4; static lean_object* l_Lean_IR_ToIR_initFn___closed__4_00___x40_Lean_Compiler_IR_ToIR_1721792695____hygCtx___hyg_2_; @@ -317,6 +318,7 @@ lean_object* l_instMonadLiftT___lam__0___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ToIR_TranslatedProj_ctorIdx___boxed(lean_object*); size_t lean_usize_sub(size_t, size_t); static lean_object* l_Lean_IR_ToIR_lowerCode___closed__5; +uint8_t l_Lean_IR_IRType_isStruct(lean_object*); lean_object* l_Lean_Core_liftIOCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_ToIR_lowerLet___closed__12; static lean_object* l_Lean_IR_ToIR_lowerDecl___closed__4; @@ -332,6 +334,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lower LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_mkPap(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ToIR_TranslatedProj_ctorElim(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_ToIR_lowerCode___closed__4; +lean_object* l_Lean_IR_ExplicitBoxing_boxedConstDecl(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_toIR(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x21___at___00Lean_IR_ToIR_getFVarValue_spec__0_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_toIR___closed__0; @@ -369,6 +372,7 @@ lean_object* l_instMonadLiftTOfMonadLift___redArg___lam__0(lean_object*, lean_ob LEAN_EXPORT lean_object* l_Lean_IR_ToIR_addDecl___boxed(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*); +uint8_t l_Lean_IR_UnboxResult_hasUnboxAttr(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ToIR_getFVarValue___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Std_Data_DHashMap_Internal_Defs_0__Std_DHashMap_Internal_Raw_u2080_expand_go___at___00Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insertIfNew___at___00Lean_IR_ToIR_bindVar_spec__0_spec__1_spec__2(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_IR_IRType_isScalar(lean_object*); @@ -530,7 +534,7 @@ static lean_object* _init_l_Lean_IR_ToIR_taggedReturnAttr___regBuiltin_Lean_IR_T { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(0u); -x_2 = lean_unsigned_to_nat(22u); +x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); @@ -542,7 +546,7 @@ static lean_object* _init_l_Lean_IR_ToIR_taggedReturnAttr___regBuiltin_Lean_IR_T { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(93u); -x_2 = lean_unsigned_to_nat(28u); +x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); @@ -570,7 +574,7 @@ static lean_object* _init_l_Lean_IR_ToIR_taggedReturnAttr___regBuiltin_Lean_IR_T { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(19u); -x_2 = lean_unsigned_to_nat(27u); +x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); @@ -582,7 +586,7 @@ static lean_object* _init_l_Lean_IR_ToIR_taggedReturnAttr___regBuiltin_Lean_IR_T { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(35u); -x_2 = lean_unsigned_to_nat(27u); +x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); @@ -2666,125 +2670,116 @@ 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; -x_6 = lean_ctor_get(x_3, 1); -lean_ctor_set_tag(x_3, 3); -lean_ctor_set(x_3, 1, x_1); -x_7 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_7, 0, x_3); -x_8 = lean_alloc_ctor(0, 2, 0); +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_6 = lean_ctor_get(x_3, 0); +x_7 = lean_ctor_get(x_2, 1); +lean_inc(x_7); +x_8 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_8, 0, x_7); lean_ctor_set(x_8, 1, x_6); -return x_8; +lean_ctor_set(x_8, 2, x_1); +x_9 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set_tag(x_3, 0); +lean_ctor_set(x_3, 0, x_9); +return x_3; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_9 = lean_ctor_get(x_3, 0); -x_10 = lean_ctor_get(x_3, 1); +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_10 = lean_ctor_get(x_3, 0); +x_11 = lean_ctor_get(x_3, 1); +lean_inc(x_11); lean_inc(x_10); -lean_inc(x_9); lean_dec(x_3); -x_11 = lean_alloc_ctor(3, 2, 0); -lean_ctor_set(x_11, 0, x_9); -lean_ctor_set(x_11, 1, x_1); -x_12 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_12, 0, x_11); -x_13 = lean_alloc_ctor(0, 2, 0); +x_12 = lean_ctor_get(x_2, 1); +lean_inc(x_12); +x_13 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_13, 0, x_12); lean_ctor_set(x_13, 1, x_10); -return x_13; +lean_ctor_set(x_13, 2, x_1); +x_14 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_14, 0, x_13); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_11); +return x_15; } } case 2: { -uint8_t x_14; -x_14 = !lean_is_exclusive(x_3); -if (x_14 == 0) +uint8_t x_16; +x_16 = !lean_is_exclusive(x_3); +if (x_16 == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_15 = lean_ctor_get(x_3, 0); -x_16 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_1); +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_17 = lean_ctor_get(x_3, 0); +x_18 = lean_ctor_get(x_2, 1); +lean_inc(x_18); +x_19 = lean_alloc_ctor(4, 3, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +lean_ctor_set(x_19, 2, x_1); lean_ctor_set_tag(x_3, 0); -lean_ctor_set(x_3, 0, x_16); -x_17 = lean_box(5); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_3); -lean_ctor_set(x_18, 1, x_17); -return x_18; +lean_ctor_set(x_3, 0, x_19); +x_20 = lean_box(5); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_3); +lean_ctor_set(x_21, 1, x_20); +return x_21; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_19 = lean_ctor_get(x_3, 0); -lean_inc(x_19); +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_22 = lean_ctor_get(x_3, 0); +lean_inc(x_22); lean_dec(x_3); -x_20 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_1); -x_21 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_21, 0, x_20); -x_22 = lean_box(5); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -return x_23; +x_23 = lean_ctor_get(x_2, 1); +lean_inc(x_23); +x_24 = lean_alloc_ctor(4, 3, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_22); +lean_ctor_set(x_24, 2, x_1); +x_25 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_25, 0, x_24); +x_26 = lean_box(5); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; } } case 3: { -uint8_t x_24; -x_24 = !lean_is_exclusive(x_3); -if (x_24 == 0) -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_25 = lean_ctor_get(x_3, 2); -x_26 = lean_ctor_get(x_3, 0); -lean_dec(x_26); -x_27 = lean_ctor_get(x_2, 2); -x_28 = lean_ctor_get(x_2, 3); -x_29 = lean_nat_add(x_27, x_28); -lean_ctor_set_tag(x_3, 5); -lean_ctor_set(x_3, 2, x_1); -lean_ctor_set(x_3, 0, x_29); -x_30 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_30, 0, x_3); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_25); -return x_31; -} -else -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_32 = lean_ctor_get(x_3, 1); -x_33 = lean_ctor_get(x_3, 2); -lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_3); -x_34 = lean_ctor_get(x_2, 2); -x_35 = lean_ctor_get(x_2, 3); -x_36 = lean_nat_add(x_34, x_35); -x_37 = lean_alloc_ctor(5, 3, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_32); -lean_ctor_set(x_37, 2, x_1); -x_38 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_38, 0, x_37); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_33); -return x_39; -} +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_28 = lean_ctor_get(x_3, 1); +lean_inc(x_28); +x_29 = lean_ctor_get(x_3, 2); +lean_inc(x_29); +lean_dec_ref(x_3); +x_30 = lean_ctor_get(x_2, 1); +x_31 = lean_ctor_get(x_2, 2); +x_32 = lean_ctor_get(x_2, 3); +x_33 = lean_nat_add(x_31, x_32); +lean_inc(x_30); +x_34 = lean_alloc_ctor(5, 4, 0); +lean_ctor_set(x_34, 0, x_30); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_28); +lean_ctor_set(x_34, 3, x_1); +x_35 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_35, 0, x_34); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_29); +return x_36; } default: { -lean_object* x_40; +lean_object* x_37; lean_dec(x_1); -x_40 = l_Lean_IR_ToIR_lowerProj___closed__1; -return x_40; +x_37 = l_Lean_IR_ToIR_lowerProj___closed__1; +return x_37; } } } @@ -3650,7 +3645,7 @@ static lean_object* _init_l_Lean_IR_ToIR_lowerLet___closed__2() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_IR_ToIR_lowerLet___closed__1; x_2 = lean_unsigned_to_nat(12u); -x_3 = lean_unsigned_to_nat(186u); +x_3 = lean_unsigned_to_nat(188u); x_4 = l_Lean_IR_ToIR_lowerLet___closed__0; x_5 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerAlt_loop___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -4719,7 +4714,7 @@ static lean_object* _init_l_Lean_IR_ToIR_lowerLet___closed__4() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_IR_ToIR_lowerLet___closed__3; x_2 = lean_unsigned_to_nat(14u); -x_3 = lean_unsigned_to_nat(256u); +x_3 = lean_unsigned_to_nat(263u); x_4 = l_Lean_IR_ToIR_lowerLet___closed__0; x_5 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerAlt_loop___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -5095,115 +5090,127 @@ return x_51; LEAN_EXPORT lean_object* l_Lean_IR_ToIR_lowerLet(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_14; uint8_t x_15; +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; x_14 = lean_st_ref_get(x_3); -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; -x_16 = lean_ctor_get(x_1, 0); +x_15 = lean_ctor_get(x_1, 0); +x_16 = lean_ctor_get(x_1, 2); x_17 = lean_ctor_get(x_1, 3); x_18 = lean_ctor_get(x_14, 3); -x_19 = lean_ctor_get(x_14, 2); -lean_dec(x_19); -x_20 = lean_ctor_get(x_14, 1); -lean_dec(x_20); -x_21 = lean_ctor_get(x_14, 0); -lean_dec(x_21); -x_22 = 0; +lean_inc_ref(x_18); +if (lean_is_exclusive(x_14)) { + lean_ctor_release(x_14, 0); + lean_ctor_release(x_14, 1); + lean_ctor_release(x_14, 2); + lean_ctor_release(x_14, 3); + x_19 = x_14; +} else { + lean_dec_ref(x_14); + x_19 = lean_box(0); +} +x_20 = 0; lean_inc(x_17); -x_23 = l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_normLetValueImp(x_18, x_17, x_22); +x_21 = l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_normLetValueImp(x_18, x_17, x_20); lean_dec_ref(x_18); -switch (lean_obj_tag(x_23)) { +switch (lean_obj_tag(x_21)) { case 0: { -uint8_t x_24; -lean_inc(x_16); +uint8_t x_22; +lean_inc(x_15); lean_dec_ref(x_1); -x_24 = !lean_is_exclusive(x_23); -if (x_24 == 0) +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) { -lean_object* x_25; lean_object* x_26; -x_25 = lean_ctor_get(x_23, 0); -x_26 = l_Lean_IR_ToIR_bindVar___redArg(x_16, x_3); -if (lean_obj_tag(x_26) == 0) +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_21, 0); +x_24 = l_Lean_IR_ToIR_bindVar___redArg(x_15, x_3); +if (lean_obj_tag(x_24) == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +lean_dec_ref(x_24); +x_26 = l_Lean_IR_ToIR_lowerLitValue(x_23); 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_ref(x_26); -x_28 = l_Lean_IR_ToIR_lowerLitValue(x_25); -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_28, 1); -lean_inc(x_30); -lean_dec_ref(x_28); -x_31 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); -if (lean_obj_tag(x_31) == 0) +x_29 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); +if (lean_obj_tag(x_29) == 0) { -uint8_t x_32; -x_32 = !lean_is_exclusive(x_31); -if (x_32 == 0) +uint8_t x_30; +x_30 = !lean_is_exclusive(x_29); +if (x_30 == 0) { -lean_object* x_33; -x_33 = lean_ctor_get(x_31, 0); -lean_ctor_set_tag(x_23, 11); -lean_ctor_set(x_23, 0, x_29); -lean_ctor_set(x_14, 3, x_33); -lean_ctor_set(x_14, 2, x_23); -lean_ctor_set(x_14, 1, x_30); -lean_ctor_set(x_14, 0, x_27); -lean_ctor_set(x_31, 0, x_14); -return x_31; +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_29, 0); +lean_ctor_set_tag(x_21, 11); +lean_ctor_set(x_21, 0, x_27); +if (lean_is_scalar(x_19)) { + x_32 = lean_alloc_ctor(0, 4, 0); +} else { + x_32 = x_19; +} +lean_ctor_set(x_32, 0, x_25); +lean_ctor_set(x_32, 1, x_28); +lean_ctor_set(x_32, 2, x_21); +lean_ctor_set(x_32, 3, x_31); +lean_ctor_set(x_29, 0, x_32); +return x_29; } else { -lean_object* x_34; lean_object* x_35; -x_34 = lean_ctor_get(x_31, 0); -lean_inc(x_34); -lean_dec(x_31); -lean_ctor_set_tag(x_23, 11); -lean_ctor_set(x_23, 0, x_29); -lean_ctor_set(x_14, 3, x_34); -lean_ctor_set(x_14, 2, x_23); -lean_ctor_set(x_14, 1, x_30); -lean_ctor_set(x_14, 0, x_27); +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_29, 0); +lean_inc(x_33); +lean_dec(x_29); +lean_ctor_set_tag(x_21, 11); +lean_ctor_set(x_21, 0, x_27); +if (lean_is_scalar(x_19)) { + x_34 = lean_alloc_ctor(0, 4, 0); +} else { + x_34 = x_19; +} +lean_ctor_set(x_34, 0, x_25); +lean_ctor_set(x_34, 1, x_28); +lean_ctor_set(x_34, 2, x_21); +lean_ctor_set(x_34, 3, x_33); x_35 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_35, 0, x_14); +lean_ctor_set(x_35, 0, x_34); return x_35; } } else { -lean_dec(x_30); -lean_dec(x_29); +lean_dec(x_28); lean_dec(x_27); -lean_free_object(x_23); -lean_free_object(x_14); -return x_31; +lean_dec(x_25); +lean_free_object(x_21); +lean_dec(x_19); +return x_29; } } else { uint8_t x_36; -lean_free_object(x_23); -lean_dec_ref(x_25); -lean_free_object(x_14); +lean_free_object(x_21); +lean_dec_ref(x_23); +lean_dec(x_19); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -x_36 = !lean_is_exclusive(x_26); +x_36 = !lean_is_exclusive(x_24); if (x_36 == 0) { -return x_26; +return x_24; } else { lean_object* x_37; lean_object* x_38; -x_37 = lean_ctor_get(x_26, 0); +x_37 = lean_ctor_get(x_24, 0); lean_inc(x_37); -lean_dec(x_26); +lean_dec(x_24); x_38 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_38, 0, x_37); return x_38; @@ -5213,10 +5220,10 @@ return x_38; else { lean_object* x_39; lean_object* x_40; -x_39 = lean_ctor_get(x_23, 0); +x_39 = lean_ctor_get(x_21, 0); lean_inc(x_39); -lean_dec(x_23); -x_40 = l_Lean_IR_ToIR_bindVar___redArg(x_16, x_3); +lean_dec(x_21); +x_40 = l_Lean_IR_ToIR_bindVar___redArg(x_15, x_3); if (lean_obj_tag(x_40) == 0) { lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; @@ -5232,7 +5239,7 @@ lean_dec_ref(x_42); x_45 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); if (lean_obj_tag(x_45) == 0) { -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; x_46 = lean_ctor_get(x_45, 0); lean_inc(x_46); if (lean_is_exclusive(x_45)) { @@ -5244,492 +5251,507 @@ if (lean_is_exclusive(x_45)) { } x_48 = lean_alloc_ctor(11, 1, 0); lean_ctor_set(x_48, 0, x_43); -lean_ctor_set(x_14, 3, x_46); -lean_ctor_set(x_14, 2, x_48); -lean_ctor_set(x_14, 1, x_44); -lean_ctor_set(x_14, 0, x_41); +if (lean_is_scalar(x_19)) { + x_49 = lean_alloc_ctor(0, 4, 0); +} else { + x_49 = x_19; +} +lean_ctor_set(x_49, 0, x_41); +lean_ctor_set(x_49, 1, x_44); +lean_ctor_set(x_49, 2, x_48); +lean_ctor_set(x_49, 3, x_46); if (lean_is_scalar(x_47)) { - x_49 = lean_alloc_ctor(0, 1, 0); + x_50 = lean_alloc_ctor(0, 1, 0); } else { - x_49 = x_47; + x_50 = x_47; } -lean_ctor_set(x_49, 0, x_14); -return x_49; +lean_ctor_set(x_50, 0, x_49); +return x_50; } else { lean_dec(x_44); lean_dec(x_43); lean_dec(x_41); -lean_free_object(x_14); +lean_dec(x_19); return x_45; } } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_dec_ref(x_39); -lean_free_object(x_14); +lean_dec(x_19); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -x_50 = lean_ctor_get(x_40, 0); -lean_inc(x_50); +x_51 = lean_ctor_get(x_40, 0); +lean_inc(x_51); if (lean_is_exclusive(x_40)) { lean_ctor_release(x_40, 0); - x_51 = x_40; + x_52 = x_40; } else { lean_dec_ref(x_40); - x_51 = lean_box(0); + x_52 = lean_box(0); } -if (lean_is_scalar(x_51)) { - x_52 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_52)) { + x_53 = lean_alloc_ctor(1, 1, 0); } else { - x_52 = x_51; + x_53 = x_52; } -lean_ctor_set(x_52, 0, x_50); -return x_52; +lean_ctor_set(x_53, 0, x_51); +return x_53; } } } case 1: { -lean_object* x_53; -lean_free_object(x_14); -x_53 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_mkErased___redArg(x_1, x_2, x_3, x_4, x_5); -return x_53; +lean_object* x_54; +lean_dec(x_19); +x_54 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_mkErased___redArg(x_1, x_2, x_3, x_4, x_5); +return x_54; } case 2: { -lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -lean_inc(x_16); +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +lean_inc(x_15); lean_dec_ref(x_1); -x_54 = lean_ctor_get(x_23, 0); -lean_inc(x_54); -x_55 = lean_ctor_get(x_23, 1); +x_55 = lean_ctor_get(x_21, 0); lean_inc(x_55); -x_56 = lean_ctor_get(x_23, 2); +x_56 = lean_ctor_get(x_21, 1); lean_inc(x_56); -lean_dec_ref(x_23); +x_57 = lean_ctor_get(x_21, 2); +lean_inc(x_57); +lean_dec_ref(x_21); lean_inc(x_5); lean_inc_ref(x_4); -lean_inc(x_54); -x_57 = l_Lean_IR_hasTrivialStructure_x3f(x_54, x_4, x_5); -if (lean_obj_tag(x_57) == 0) -{ -lean_object* x_58; -x_58 = lean_ctor_get(x_57, 0); -lean_inc(x_58); -lean_dec_ref(x_57); -if (lean_obj_tag(x_58) == 1) +lean_inc(x_55); +x_58 = l_Lean_IR_hasTrivialStructure_x3f(x_55, x_4, x_5); +if (lean_obj_tag(x_58) == 0) { -uint8_t x_59; -lean_dec(x_54); -lean_free_object(x_14); -x_59 = !lean_is_exclusive(x_58); -if (x_59 == 0) +lean_object* x_59; +x_59 = lean_ctor_get(x_58, 0); +lean_inc(x_59); +lean_dec_ref(x_58); +if (lean_obj_tag(x_59) == 1) { -lean_object* x_60; lean_object* x_61; uint8_t x_62; -x_60 = lean_ctor_get(x_58, 0); -x_61 = lean_ctor_get(x_60, 2); -lean_inc(x_61); -lean_dec(x_60); -x_62 = lean_nat_dec_eq(x_61, x_55); +uint8_t x_60; lean_dec(x_55); -lean_dec(x_61); -if (x_62 == 0) +lean_dec(x_19); +x_60 = !lean_is_exclusive(x_59); +if (x_60 == 0) { -lean_object* x_63; -lean_free_object(x_58); +lean_object* x_61; lean_object* x_62; uint8_t x_63; +x_61 = lean_ctor_get(x_59, 0); +x_62 = lean_ctor_get(x_61, 2); +lean_inc(x_62); +lean_dec(x_61); +x_63 = lean_nat_dec_eq(x_62, x_56); lean_dec(x_56); -x_63 = l_Lean_IR_ToIR_bindErased___redArg(x_16, x_3); -if (lean_obj_tag(x_63) == 0) +lean_dec(x_62); +if (x_63 == 0) { lean_object* x_64; -lean_dec_ref(x_63); -x_64 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); -return x_64; +lean_free_object(x_59); +lean_dec(x_57); +x_64 = l_Lean_IR_ToIR_bindErased___redArg(x_15, x_3); +if (lean_obj_tag(x_64) == 0) +{ +lean_object* x_65; +lean_dec_ref(x_64); +x_65 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); +return x_65; } else { -uint8_t x_65; +uint8_t x_66; lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -x_65 = !lean_is_exclusive(x_63); -if (x_65 == 0) +x_66 = !lean_is_exclusive(x_64); +if (x_66 == 0) { -return x_63; +return x_64; } else { -lean_object* x_66; lean_object* x_67; -x_66 = lean_ctor_get(x_63, 0); -lean_inc(x_66); -lean_dec(x_63); -x_67 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_67, 0, x_66); -return x_67; +lean_object* x_67; lean_object* x_68; +x_67 = lean_ctor_get(x_64, 0); +lean_inc(x_67); +lean_dec(x_64); +x_68 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_68, 0, x_67); +return x_68; } } } else { -lean_object* x_68; uint8_t x_69; -x_68 = lean_st_ref_take(x_3); -x_69 = !lean_is_exclusive(x_68); -if (x_69 == 0) +lean_object* x_69; uint8_t x_70; +x_69 = lean_st_ref_take(x_3); +x_70 = !lean_is_exclusive(x_69); +if (x_70 == 0) { -lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_70 = lean_ctor_get(x_68, 3); -lean_ctor_set(x_58, 0, x_56); -x_71 = l_Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_ToIR_lowerParam_spec__0___redArg(x_70, x_16, x_58); -lean_ctor_set(x_68, 3, x_71); -x_72 = lean_st_ref_set(x_3, x_68); -x_73 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); -return x_73; +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_71 = lean_ctor_get(x_69, 3); +lean_ctor_set(x_59, 0, x_57); +x_72 = l_Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_ToIR_lowerParam_spec__0___redArg(x_71, x_15, x_59); +lean_ctor_set(x_69, 3, x_72); +x_73 = lean_st_ref_set(x_3, x_69); +x_74 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); +return x_74; } else { -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_74 = lean_ctor_get(x_68, 0); -x_75 = lean_ctor_get(x_68, 1); -x_76 = lean_ctor_get(x_68, 2); -x_77 = lean_ctor_get(x_68, 3); +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_75 = lean_ctor_get(x_69, 0); +x_76 = lean_ctor_get(x_69, 1); +x_77 = lean_ctor_get(x_69, 2); +x_78 = lean_ctor_get(x_69, 3); +lean_inc(x_78); lean_inc(x_77); lean_inc(x_76); lean_inc(x_75); -lean_inc(x_74); -lean_dec(x_68); -lean_ctor_set(x_58, 0, x_56); -x_78 = l_Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_ToIR_lowerParam_spec__0___redArg(x_77, x_16, x_58); -x_79 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_79, 0, x_74); -lean_ctor_set(x_79, 1, x_75); -lean_ctor_set(x_79, 2, x_76); -lean_ctor_set(x_79, 3, x_78); -x_80 = lean_st_ref_set(x_3, x_79); -x_81 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); -return x_81; +lean_dec(x_69); +lean_ctor_set(x_59, 0, x_57); +x_79 = l_Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_ToIR_lowerParam_spec__0___redArg(x_78, x_15, x_59); +x_80 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_80, 0, x_75); +lean_ctor_set(x_80, 1, x_76); +lean_ctor_set(x_80, 2, x_77); +lean_ctor_set(x_80, 3, x_79); +x_81 = lean_st_ref_set(x_3, x_80); +x_82 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); +return x_82; } } } else { -lean_object* x_82; lean_object* x_83; uint8_t x_84; -x_82 = lean_ctor_get(x_58, 0); -lean_inc(x_82); -lean_dec(x_58); -x_83 = lean_ctor_get(x_82, 2); +lean_object* x_83; lean_object* x_84; uint8_t x_85; +x_83 = lean_ctor_get(x_59, 0); lean_inc(x_83); -lean_dec(x_82); -x_84 = lean_nat_dec_eq(x_83, x_55); -lean_dec(x_55); +lean_dec(x_59); +x_84 = lean_ctor_get(x_83, 2); +lean_inc(x_84); lean_dec(x_83); -if (x_84 == 0) -{ -lean_object* x_85; +x_85 = lean_nat_dec_eq(x_84, x_56); lean_dec(x_56); -x_85 = l_Lean_IR_ToIR_bindErased___redArg(x_16, x_3); -if (lean_obj_tag(x_85) == 0) +lean_dec(x_84); +if (x_85 == 0) { lean_object* x_86; -lean_dec_ref(x_85); -x_86 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); -return x_86; +lean_dec(x_57); +x_86 = l_Lean_IR_ToIR_bindErased___redArg(x_15, x_3); +if (lean_obj_tag(x_86) == 0) +{ +lean_object* x_87; +lean_dec_ref(x_86); +x_87 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); +return x_87; } else { -lean_object* x_87; lean_object* x_88; lean_object* x_89; +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -x_87 = lean_ctor_get(x_85, 0); -lean_inc(x_87); -if (lean_is_exclusive(x_85)) { - lean_ctor_release(x_85, 0); - x_88 = x_85; +x_88 = lean_ctor_get(x_86, 0); +lean_inc(x_88); +if (lean_is_exclusive(x_86)) { + lean_ctor_release(x_86, 0); + x_89 = x_86; } else { - lean_dec_ref(x_85); - x_88 = lean_box(0); + lean_dec_ref(x_86); + x_89 = lean_box(0); } -if (lean_is_scalar(x_88)) { - x_89 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_89)) { + x_90 = lean_alloc_ctor(1, 1, 0); } else { - x_89 = x_88; + x_90 = x_89; } -lean_ctor_set(x_89, 0, x_87); -return x_89; +lean_ctor_set(x_90, 0, x_88); +return x_90; } } else { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_90 = lean_st_ref_take(x_3); -x_91 = lean_ctor_get(x_90, 0); -lean_inc_ref(x_91); -x_92 = lean_ctor_get(x_90, 1); +lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_91 = lean_st_ref_take(x_3); +x_92 = lean_ctor_get(x_91, 0); lean_inc_ref(x_92); -x_93 = lean_ctor_get(x_90, 2); -lean_inc(x_93); -x_94 = lean_ctor_get(x_90, 3); -lean_inc_ref(x_94); -if (lean_is_exclusive(x_90)) { - lean_ctor_release(x_90, 0); - lean_ctor_release(x_90, 1); - lean_ctor_release(x_90, 2); - lean_ctor_release(x_90, 3); - x_95 = x_90; -} else { - lean_dec_ref(x_90); - x_95 = lean_box(0); -} -x_96 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_96, 0, x_56); -x_97 = l_Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_ToIR_lowerParam_spec__0___redArg(x_94, x_16, x_96); -if (lean_is_scalar(x_95)) { - x_98 = lean_alloc_ctor(0, 4, 0); -} else { - x_98 = x_95; -} -lean_ctor_set(x_98, 0, x_91); -lean_ctor_set(x_98, 1, x_92); -lean_ctor_set(x_98, 2, x_93); -lean_ctor_set(x_98, 3, x_97); -x_99 = lean_st_ref_set(x_3, x_98); -x_100 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); -return x_100; +x_93 = lean_ctor_get(x_91, 1); +lean_inc_ref(x_93); +x_94 = lean_ctor_get(x_91, 2); +lean_inc(x_94); +x_95 = lean_ctor_get(x_91, 3); +lean_inc_ref(x_95); +if (lean_is_exclusive(x_91)) { + lean_ctor_release(x_91, 0); + lean_ctor_release(x_91, 1); + lean_ctor_release(x_91, 2); + lean_ctor_release(x_91, 3); + x_96 = x_91; +} else { + lean_dec_ref(x_91); + x_96 = lean_box(0); +} +x_97 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_97, 0, x_57); +x_98 = l_Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_ToIR_lowerParam_spec__0___redArg(x_95, x_15, x_97); +if (lean_is_scalar(x_96)) { + x_99 = lean_alloc_ctor(0, 4, 0); +} else { + x_99 = x_96; +} +lean_ctor_set(x_99, 0, x_92); +lean_ctor_set(x_99, 1, x_93); +lean_ctor_set(x_99, 2, x_94); +lean_ctor_set(x_99, 3, x_98); +x_100 = lean_st_ref_set(x_3, x_99); +x_101 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); +return x_101; } } } else { -lean_object* x_101; -lean_dec(x_58); -x_101 = l_Lean_IR_ToIR_getFVarValue___redArg(x_56, x_3); -if (lean_obj_tag(x_101) == 0) -{ lean_object* x_102; -x_102 = lean_ctor_get(x_101, 0); -lean_inc(x_102); -lean_dec_ref(x_101); +lean_dec(x_59); +x_102 = l_Lean_IR_ToIR_getFVarValue___redArg(x_57, x_3); if (lean_obj_tag(x_102) == 0) { -lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +lean_object* x_103; x_103 = lean_ctor_get(x_102, 0); lean_inc(x_103); lean_dec_ref(x_102); -x_104 = lean_st_ref_get(x_5); -x_105 = lean_ctor_get(x_104, 0); -lean_inc_ref(x_105); -lean_dec(x_104); -x_106 = l_Lean_Environment_find_x3f(x_105, x_54, x_22); -if (lean_obj_tag(x_106) == 1) -{ -lean_object* x_107; -x_107 = lean_ctor_get(x_106, 0); -lean_inc(x_107); -lean_dec_ref(x_106); -if (lean_obj_tag(x_107) == 5) -{ -lean_object* x_108; lean_object* x_109; +if (lean_obj_tag(x_103) == 0) +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_104 = lean_ctor_get(x_103, 0); +lean_inc(x_104); +lean_dec_ref(x_103); +x_105 = lean_st_ref_get(x_5); +x_106 = lean_ctor_get(x_105, 0); +lean_inc_ref(x_106); +lean_dec(x_105); +x_107 = l_Lean_Environment_find_x3f(x_106, x_55, x_20); +if (lean_obj_tag(x_107) == 1) +{ +lean_object* x_108; x_108 = lean_ctor_get(x_107, 0); -lean_inc_ref(x_108); +lean_inc(x_108); lean_dec_ref(x_107); -x_109 = lean_ctor_get(x_108, 4); -lean_inc(x_109); +if (lean_obj_tag(x_108) == 5) +{ +lean_object* x_109; lean_object* x_110; +x_109 = lean_ctor_get(x_108, 0); +lean_inc_ref(x_109); lean_dec_ref(x_108); -if (lean_obj_tag(x_109) == 1) +x_110 = lean_ctor_get(x_109, 4); +lean_inc(x_110); +lean_dec_ref(x_109); +if (lean_obj_tag(x_110) == 1) { -lean_object* x_110; -x_110 = lean_ctor_get(x_109, 1); -if (lean_obj_tag(x_110) == 0) +lean_object* x_111; +x_111 = lean_ctor_get(x_110, 1); +if (lean_obj_tag(x_111) == 0) { -lean_object* x_111; lean_object* x_112; -x_111 = lean_ctor_get(x_109, 0); -lean_inc(x_111); -lean_dec_ref(x_109); +lean_object* x_112; lean_object* x_113; +x_112 = lean_ctor_get(x_110, 0); +lean_inc(x_112); +lean_dec_ref(x_110); lean_inc(x_5); lean_inc_ref(x_4); -x_112 = l_Lean_IR_getCtorLayout(x_111, x_4, x_5); -if (lean_obj_tag(x_112) == 0) +x_113 = l_Lean_IR_getCtorLayout(x_112, x_4, x_5); +if (lean_obj_tag(x_113) == 0) { -lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; -x_113 = lean_ctor_get(x_112, 0); -lean_inc(x_113); -lean_dec_ref(x_112); +lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; x_114 = lean_ctor_get(x_113, 0); -lean_inc_ref(x_114); -x_115 = lean_ctor_get(x_113, 1); +lean_inc(x_114); +lean_dec_ref(x_113); +x_115 = lean_ctor_get(x_114, 0); lean_inc_ref(x_115); -lean_dec(x_113); -x_116 = lean_box(0); -x_117 = lean_array_get(x_116, x_115, x_55); -lean_dec(x_55); +x_116 = lean_ctor_get(x_114, 1); +lean_inc_ref(x_116); +lean_dec(x_114); +x_117 = lean_box(0); +x_118 = lean_array_get(x_117, x_116, x_56); +lean_dec(x_56); +lean_dec_ref(x_116); +x_119 = l_Lean_IR_ToIR_lowerProj(x_104, x_115, x_118); lean_dec_ref(x_115); -x_118 = l_Lean_IR_ToIR_lowerProj(x_103, x_114, x_117); -lean_dec_ref(x_114); -x_119 = lean_ctor_get(x_118, 0); -lean_inc(x_119); -if (lean_obj_tag(x_119) == 0) -{ -lean_object* x_120; lean_object* x_121; lean_object* x_122; -x_120 = lean_ctor_get(x_118, 1); +x_120 = lean_ctor_get(x_119, 0); lean_inc(x_120); -lean_dec_ref(x_118); -x_121 = lean_ctor_get(x_119, 0); -lean_inc_ref(x_121); -lean_dec_ref(x_119); -x_122 = l_Lean_IR_ToIR_bindVar___redArg(x_16, x_3); -if (lean_obj_tag(x_122) == 0) -{ -lean_object* x_123; lean_object* x_124; -x_123 = lean_ctor_get(x_122, 0); -lean_inc(x_123); -lean_dec_ref(x_122); -x_124 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); -if (lean_obj_tag(x_124) == 0) +if (lean_obj_tag(x_120) == 0) { -uint8_t x_125; -x_125 = !lean_is_exclusive(x_124); -if (x_125 == 0) +lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_121 = lean_ctor_get(x_119, 1); +lean_inc(x_121); +lean_dec_ref(x_119); +x_122 = lean_ctor_get(x_120, 0); +lean_inc_ref(x_122); +lean_dec_ref(x_120); +x_123 = l_Lean_IR_ToIR_bindVar___redArg(x_15, x_3); +if (lean_obj_tag(x_123) == 0) +{ +lean_object* x_124; lean_object* x_125; +x_124 = lean_ctor_get(x_123, 0); +lean_inc(x_124); +lean_dec_ref(x_123); +x_125 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); +if (lean_obj_tag(x_125) == 0) +{ +uint8_t x_126; +x_126 = !lean_is_exclusive(x_125); +if (x_126 == 0) { -lean_object* x_126; -x_126 = lean_ctor_get(x_124, 0); -lean_ctor_set(x_14, 3, x_126); -lean_ctor_set(x_14, 2, x_121); -lean_ctor_set(x_14, 1, x_120); -lean_ctor_set(x_14, 0, x_123); -lean_ctor_set(x_124, 0, x_14); -return x_124; +lean_object* x_127; lean_object* x_128; +x_127 = lean_ctor_get(x_125, 0); +if (lean_is_scalar(x_19)) { + x_128 = lean_alloc_ctor(0, 4, 0); +} else { + x_128 = x_19; +} +lean_ctor_set(x_128, 0, x_124); +lean_ctor_set(x_128, 1, x_121); +lean_ctor_set(x_128, 2, x_122); +lean_ctor_set(x_128, 3, x_127); +lean_ctor_set(x_125, 0, x_128); +return x_125; } else { -lean_object* x_127; lean_object* x_128; -x_127 = lean_ctor_get(x_124, 0); -lean_inc(x_127); -lean_dec(x_124); -lean_ctor_set(x_14, 3, x_127); -lean_ctor_set(x_14, 2, x_121); -lean_ctor_set(x_14, 1, x_120); -lean_ctor_set(x_14, 0, x_123); -x_128 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_128, 0, x_14); -return x_128; +lean_object* x_129; lean_object* x_130; lean_object* x_131; +x_129 = lean_ctor_get(x_125, 0); +lean_inc(x_129); +lean_dec(x_125); +if (lean_is_scalar(x_19)) { + x_130 = lean_alloc_ctor(0, 4, 0); +} else { + x_130 = x_19; +} +lean_ctor_set(x_130, 0, x_124); +lean_ctor_set(x_130, 1, x_121); +lean_ctor_set(x_130, 2, x_122); +lean_ctor_set(x_130, 3, x_129); +x_131 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_131, 0, x_130); +return x_131; } } else { -lean_dec(x_123); -lean_dec_ref(x_121); -lean_dec(x_120); -lean_free_object(x_14); -return x_124; +lean_dec(x_124); +lean_dec_ref(x_122); +lean_dec(x_121); +lean_dec(x_19); +return x_125; } } else { -uint8_t x_129; -lean_dec_ref(x_121); -lean_dec(x_120); -lean_free_object(x_14); +uint8_t x_132; +lean_dec_ref(x_122); +lean_dec(x_121); +lean_dec(x_19); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -x_129 = !lean_is_exclusive(x_122); -if (x_129 == 0) +x_132 = !lean_is_exclusive(x_123); +if (x_132 == 0) { -return x_122; +return x_123; } else { -lean_object* x_130; lean_object* x_131; -x_130 = lean_ctor_get(x_122, 0); -lean_inc(x_130); -lean_dec(x_122); -x_131 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_131, 0, x_130); -return x_131; +lean_object* x_133; lean_object* x_134; +x_133 = lean_ctor_get(x_123, 0); +lean_inc(x_133); +lean_dec(x_123); +x_134 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_134, 0, x_133); +return x_134; } } } else { -lean_object* x_132; -lean_dec_ref(x_118); -lean_free_object(x_14); -x_132 = l_Lean_IR_ToIR_bindErased___redArg(x_16, x_3); -if (lean_obj_tag(x_132) == 0) +lean_object* x_135; +lean_dec_ref(x_119); +lean_dec(x_19); +x_135 = l_Lean_IR_ToIR_bindErased___redArg(x_15, x_3); +if (lean_obj_tag(x_135) == 0) { -lean_object* x_133; -lean_dec_ref(x_132); -x_133 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); -return x_133; +lean_object* x_136; +lean_dec_ref(x_135); +x_136 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); +return x_136; } else { -uint8_t x_134; +uint8_t x_137; lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -x_134 = !lean_is_exclusive(x_132); -if (x_134 == 0) +x_137 = !lean_is_exclusive(x_135); +if (x_137 == 0) { -return x_132; +return x_135; } else { -lean_object* x_135; lean_object* x_136; -x_135 = lean_ctor_get(x_132, 0); -lean_inc(x_135); -lean_dec(x_132); -x_136 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_136, 0, x_135); -return x_136; +lean_object* x_138; lean_object* x_139; +x_138 = lean_ctor_get(x_135, 0); +lean_inc(x_138); +lean_dec(x_135); +x_139 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_139, 0, x_138); +return x_139; } } } } else { -uint8_t x_137; -lean_dec(x_103); -lean_dec(x_55); -lean_free_object(x_14); -lean_dec(x_16); +uint8_t x_140; +lean_dec(x_104); +lean_dec(x_56); +lean_dec(x_19); +lean_dec(x_15); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -x_137 = !lean_is_exclusive(x_112); -if (x_137 == 0) +x_140 = !lean_is_exclusive(x_113); +if (x_140 == 0) { -return x_112; +return x_113; } else { -lean_object* x_138; lean_object* x_139; -x_138 = lean_ctor_get(x_112, 0); -lean_inc(x_138); -lean_dec(x_112); -x_139 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_139, 0, x_138); -return x_139; +lean_object* x_141; lean_object* x_142; +x_141 = lean_ctor_get(x_113, 0); +lean_inc(x_141); +lean_dec(x_113); +x_142 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_142, 0, x_141); +return x_142; } } } else { -lean_dec_ref(x_109); -lean_dec(x_103); -lean_dec(x_55); -lean_free_object(x_14); -lean_dec(x_16); +lean_dec_ref(x_110); +lean_dec(x_104); +lean_dec(x_56); +lean_dec(x_19); +lean_dec(x_15); lean_dec_ref(x_2); x_7 = x_3; x_8 = x_4; @@ -5740,11 +5762,11 @@ goto block_13; } else { -lean_dec(x_109); -lean_dec(x_103); -lean_dec(x_55); -lean_free_object(x_14); -lean_dec(x_16); +lean_dec(x_110); +lean_dec(x_104); +lean_dec(x_56); +lean_dec(x_19); +lean_dec(x_15); lean_dec_ref(x_2); x_7 = x_3; x_8 = x_4; @@ -5755,11 +5777,11 @@ goto block_13; } else { -lean_dec(x_107); -lean_dec(x_103); -lean_dec(x_55); -lean_free_object(x_14); -lean_dec(x_16); +lean_dec(x_108); +lean_dec(x_104); +lean_dec(x_56); +lean_dec(x_19); +lean_dec(x_15); lean_dec_ref(x_2); x_7 = x_3; x_8 = x_4; @@ -5770,11 +5792,11 @@ goto block_13; } else { -lean_dec(x_106); -lean_dec(x_103); -lean_dec(x_55); -lean_free_object(x_14); -lean_dec(x_16); +lean_dec(x_107); +lean_dec(x_104); +lean_dec(x_56); +lean_dec(x_19); +lean_dec(x_15); lean_dec_ref(x_2); x_7 = x_3; x_8 = x_4; @@ -5785,65 +5807,36 @@ goto block_13; } else { -lean_object* x_140; +lean_object* x_143; +lean_dec(x_56); lean_dec(x_55); -lean_dec(x_54); -lean_free_object(x_14); -x_140 = l_Lean_IR_ToIR_bindErased___redArg(x_16, x_3); -if (lean_obj_tag(x_140) == 0) -{ -lean_object* x_141; -lean_dec_ref(x_140); -x_141 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); -return x_141; -} -else -{ -uint8_t x_142; -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -x_142 = !lean_is_exclusive(x_140); -if (x_142 == 0) -{ -return x_140; -} -else +lean_dec(x_19); +x_143 = l_Lean_IR_ToIR_bindErased___redArg(x_15, x_3); +if (lean_obj_tag(x_143) == 0) { -lean_object* x_143; lean_object* x_144; -x_143 = lean_ctor_get(x_140, 0); -lean_inc(x_143); -lean_dec(x_140); -x_144 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_144, 0, x_143); +lean_object* x_144; +lean_dec_ref(x_143); +x_144 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); return x_144; } -} -} -} else { uint8_t x_145; -lean_dec(x_55); -lean_dec(x_54); -lean_free_object(x_14); -lean_dec(x_16); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -x_145 = !lean_is_exclusive(x_101); +x_145 = !lean_is_exclusive(x_143); if (x_145 == 0) { -return x_101; +return x_143; } else { lean_object* x_146; lean_object* x_147; -x_146 = lean_ctor_get(x_101, 0); +x_146 = lean_ctor_get(x_143, 0); lean_inc(x_146); -lean_dec(x_101); +lean_dec(x_143); x_147 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_147, 0, x_146); return x_147; @@ -5856,3907 +5849,2497 @@ else uint8_t x_148; lean_dec(x_56); lean_dec(x_55); -lean_dec(x_54); -lean_free_object(x_14); -lean_dec(x_16); +lean_dec(x_19); +lean_dec(x_15); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -x_148 = !lean_is_exclusive(x_57); +x_148 = !lean_is_exclusive(x_102); if (x_148 == 0) { -return x_57; +return x_102; } else { lean_object* x_149; lean_object* x_150; -x_149 = lean_ctor_get(x_57, 0); +x_149 = lean_ctor_get(x_102, 0); lean_inc(x_149); -lean_dec(x_57); +lean_dec(x_102); x_150 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_150, 0, x_149); return x_150; } } } +} +else +{ +uint8_t x_151; +lean_dec(x_57); +lean_dec(x_56); +lean_dec(x_55); +lean_dec(x_19); +lean_dec(x_15); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec(x_3); +lean_dec_ref(x_2); +x_151 = !lean_is_exclusive(x_58); +if (x_151 == 0) +{ +return x_58; +} +else +{ +lean_object* x_152; lean_object* x_153; +x_152 = lean_ctor_get(x_58, 0); +lean_inc(x_152); +lean_dec(x_58); +x_153 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_153, 0, x_152); +return x_153; +} +} +} case 3: { -lean_object* x_151; lean_object* x_152; size_t x_153; size_t x_154; lean_object* x_155; -x_151 = lean_ctor_get(x_23, 0); -lean_inc(x_151); -x_152 = lean_ctor_get(x_23, 2); -lean_inc_ref(x_152); -lean_dec_ref(x_23); -x_153 = lean_array_size(x_152); -x_154 = 0; -lean_inc_ref(x_152); -x_155 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ToIR_lowerCode_spec__3___redArg(x_153, x_154, x_152, x_3); -if (lean_obj_tag(x_155) == 0) -{ -lean_object* x_156; lean_object* x_157; -x_156 = lean_ctor_get(x_155, 0); -lean_inc(x_156); -lean_dec_ref(x_155); -lean_inc(x_151); -x_157 = l_Lean_IR_ToIR_findDecl___redArg(x_151, x_5); -if (lean_obj_tag(x_157) == 0) -{ -lean_object* x_158; -x_158 = lean_ctor_get(x_157, 0); -lean_inc(x_158); -lean_dec_ref(x_157); -if (lean_obj_tag(x_158) == 1) -{ -lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; -lean_dec_ref(x_152); -lean_free_object(x_14); +lean_object* x_154; lean_object* x_155; size_t x_156; size_t x_157; lean_object* x_158; +x_154 = lean_ctor_get(x_21, 0); +lean_inc(x_154); +x_155 = lean_ctor_get(x_21, 2); +lean_inc_ref(x_155); +lean_dec_ref(x_21); +x_156 = lean_array_size(x_155); +x_157 = 0; +lean_inc_ref(x_155); +x_158 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ToIR_lowerCode_spec__3___redArg(x_156, x_157, x_155, x_3); +if (lean_obj_tag(x_158) == 0) +{ +lean_object* x_159; lean_object* x_160; x_159 = lean_ctor_get(x_158, 0); lean_inc(x_159); lean_dec_ref(x_158); -x_160 = l_Lean_IR_Decl_params(x_159); -lean_dec(x_159); -x_161 = lean_array_get_size(x_160); +lean_inc(x_154); +x_160 = l_Lean_IR_ToIR_findDecl___redArg(x_154, x_5); +if (lean_obj_tag(x_160) == 0) +{ +lean_object* x_161; +x_161 = lean_ctor_get(x_160, 0); +lean_inc(x_161); lean_dec_ref(x_160); -x_162 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_mkApplication(x_1, x_2, x_151, x_161, x_156, x_3, x_4, x_5); -return x_162; +if (lean_obj_tag(x_161) == 1) +{ +lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; +lean_dec_ref(x_155); +lean_dec(x_19); +x_162 = lean_ctor_get(x_161, 0); +lean_inc(x_162); +lean_dec_ref(x_161); +x_163 = l_Lean_IR_Decl_params(x_162); +lean_dec(x_162); +x_164 = lean_array_get_size(x_163); +lean_dec_ref(x_163); +x_165 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_mkApplication(x_1, x_2, x_154, x_164, x_159, x_3, x_4, x_5); +return x_165; } else { -lean_object* x_163; -lean_dec(x_158); -lean_inc(x_151); -x_163 = l_Lean_Compiler_LCNF_getMonoDecl_x3f___redArg(x_151, x_5); -if (lean_obj_tag(x_163) == 0) +lean_object* x_166; +lean_dec(x_161); +lean_inc(x_154); +x_166 = l_Lean_Compiler_LCNF_getMonoDecl_x3f___redArg(x_154, x_5); +if (lean_obj_tag(x_166) == 0) { -lean_object* x_164; -x_164 = lean_ctor_get(x_163, 0); -lean_inc(x_164); -lean_dec_ref(x_163); -if (lean_obj_tag(x_164) == 1) -{ -lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; -lean_dec_ref(x_152); -lean_free_object(x_14); -x_165 = lean_ctor_get(x_164, 0); -lean_inc(x_165); -lean_dec_ref(x_164); -x_166 = lean_ctor_get(x_165, 3); -lean_inc_ref(x_166); -lean_dec(x_165); -x_167 = lean_array_get_size(x_166); +lean_object* x_167; +x_167 = lean_ctor_get(x_166, 0); +lean_inc(x_167); lean_dec_ref(x_166); -x_168 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_mkApplication(x_1, x_2, x_151, x_167, x_156, x_3, x_4, x_5); -return x_168; +if (lean_obj_tag(x_167) == 1) +{ +lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; +lean_dec_ref(x_155); +lean_dec(x_19); +x_168 = lean_ctor_get(x_167, 0); +lean_inc(x_168); +lean_dec_ref(x_167); +x_169 = lean_ctor_get(x_168, 3); +lean_inc_ref(x_169); +lean_dec(x_168); +x_170 = lean_array_get_size(x_169); +lean_dec_ref(x_169); +x_171 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_mkApplication(x_1, x_2, x_154, x_170, x_159, x_3, x_4, x_5); +return x_171; } else { -lean_object* x_169; lean_object* x_170; lean_object* x_171; -lean_dec(x_164); -x_169 = lean_st_ref_get(x_5); -x_170 = lean_ctor_get(x_169, 0); -lean_inc_ref(x_170); -lean_dec(x_169); -lean_inc(x_151); -x_171 = l_Lean_Environment_find_x3f(x_170, x_151, x_22); -if (lean_obj_tag(x_171) == 0) -{ -lean_object* x_172; lean_object* x_173; -lean_dec(x_156); -lean_dec_ref(x_152); -lean_dec(x_151); -lean_free_object(x_14); +lean_object* x_172; lean_object* x_173; lean_object* x_174; +lean_dec(x_167); +x_172 = lean_st_ref_get(x_5); +x_173 = lean_ctor_get(x_172, 0); +lean_inc_ref(x_173); +lean_dec(x_172); +lean_inc(x_154); +x_174 = l_Lean_Environment_find_x3f(x_173, x_154, x_20); +if (lean_obj_tag(x_174) == 0) +{ +lean_object* x_175; lean_object* x_176; +lean_dec(x_159); +lean_dec_ref(x_155); +lean_dec(x_154); +lean_dec(x_19); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_172 = l_Lean_IR_ToIR_lowerLet___closed__4; -x_173 = l_panic___at___00__private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerAlt_loop_spec__0(x_172, x_3, x_4, x_5); -return x_173; +x_175 = l_Lean_IR_ToIR_lowerLet___closed__4; +x_176 = l_panic___at___00__private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerAlt_loop_spec__0(x_175, x_3, x_4, x_5); +return x_176; } else { -uint8_t x_174; -x_174 = !lean_is_exclusive(x_171); -if (x_174 == 0) +uint8_t x_177; +x_177 = !lean_is_exclusive(x_174); +if (x_177 == 0) { -lean_object* x_175; -x_175 = lean_ctor_get(x_171, 0); -switch (lean_obj_tag(x_175)) { +lean_object* x_178; +x_178 = lean_ctor_get(x_174, 0); +switch (lean_obj_tag(x_178)) { case 0: { -uint8_t x_176; -lean_free_object(x_171); -lean_dec(x_156); -lean_dec_ref(x_152); -lean_free_object(x_14); +uint8_t x_179; +lean_free_object(x_174); +lean_dec(x_159); +lean_dec_ref(x_155); +lean_dec(x_19); lean_dec(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_176 = !lean_is_exclusive(x_175); -if (x_176 == 0) -{ -lean_object* x_177; lean_object* x_178; lean_object* x_179; uint8_t x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; -x_177 = lean_ctor_get(x_175, 0); -lean_dec(x_177); -x_178 = l_Lean_IR_ToIR_lowerLet___closed__7; -x_179 = l_Lean_IR_ToIR_lowerLet___closed__9; -x_180 = 1; -x_181 = l_Lean_Name_toString(x_151, x_180); -lean_ctor_set_tag(x_175, 3); -lean_ctor_set(x_175, 0, x_181); -x_182 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_182, 0, x_179); -lean_ctor_set(x_182, 1, x_175); -x_183 = l_Lean_IR_ToIR_lowerLet___closed__11; -x_184 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_184, 0, x_182); -lean_ctor_set(x_184, 1, x_183); -x_185 = l_Lean_MessageData_ofFormat(x_184); -x_186 = l_Lean_throwNamedError___at___00Lean_IR_ToIR_lowerLet_spec__8___redArg(x_178, x_185, x_4, x_5); +x_179 = !lean_is_exclusive(x_178); +if (x_179 == 0) +{ +lean_object* x_180; lean_object* x_181; lean_object* x_182; uint8_t x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; +x_180 = lean_ctor_get(x_178, 0); +lean_dec(x_180); +x_181 = l_Lean_IR_ToIR_lowerLet___closed__7; +x_182 = l_Lean_IR_ToIR_lowerLet___closed__9; +x_183 = 1; +x_184 = l_Lean_Name_toString(x_154, x_183); +lean_ctor_set_tag(x_178, 3); +lean_ctor_set(x_178, 0, x_184); +x_185 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_185, 0, x_182); +lean_ctor_set(x_185, 1, x_178); +x_186 = l_Lean_IR_ToIR_lowerLet___closed__11; +x_187 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_187, 0, x_185); +lean_ctor_set(x_187, 1, x_186); +x_188 = l_Lean_MessageData_ofFormat(x_187); +x_189 = l_Lean_throwNamedError___at___00Lean_IR_ToIR_lowerLet_spec__8___redArg(x_181, x_188, x_4, x_5); lean_dec(x_5); lean_dec_ref(x_4); -return x_186; -} -else -{ -lean_object* x_187; lean_object* x_188; uint8_t 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_dec(x_175); -x_187 = l_Lean_IR_ToIR_lowerLet___closed__7; -x_188 = l_Lean_IR_ToIR_lowerLet___closed__9; -x_189 = 1; -x_190 = l_Lean_Name_toString(x_151, x_189); -x_191 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_191, 0, x_190); -x_192 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_192, 0, x_188); -lean_ctor_set(x_192, 1, x_191); -x_193 = l_Lean_IR_ToIR_lowerLet___closed__11; -x_194 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_194, 0, x_192); -lean_ctor_set(x_194, 1, x_193); -x_195 = l_Lean_MessageData_ofFormat(x_194); -x_196 = l_Lean_throwNamedError___at___00Lean_IR_ToIR_lowerLet_spec__8___redArg(x_187, x_195, x_4, x_5); +return x_189; +} +else +{ +lean_object* x_190; lean_object* x_191; uint8_t x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; +lean_dec(x_178); +x_190 = l_Lean_IR_ToIR_lowerLet___closed__7; +x_191 = l_Lean_IR_ToIR_lowerLet___closed__9; +x_192 = 1; +x_193 = l_Lean_Name_toString(x_154, x_192); +x_194 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_194, 0, x_193); +x_195 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_195, 0, x_191); +lean_ctor_set(x_195, 1, x_194); +x_196 = l_Lean_IR_ToIR_lowerLet___closed__11; +x_197 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_197, 0, x_195); +lean_ctor_set(x_197, 1, x_196); +x_198 = l_Lean_MessageData_ofFormat(x_197); +x_199 = l_Lean_throwNamedError___at___00Lean_IR_ToIR_lowerLet_spec__8___redArg(x_190, x_198, x_4, x_5); lean_dec(x_5); lean_dec_ref(x_4); -return x_196; +return x_199; } } case 2: { -uint8_t x_197; -lean_free_object(x_171); -lean_dec(x_156); -lean_dec_ref(x_152); -lean_free_object(x_14); +uint8_t x_200; +lean_free_object(x_174); +lean_dec(x_159); +lean_dec_ref(x_155); +lean_dec(x_19); lean_dec(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_197 = !lean_is_exclusive(x_175); -if (x_197 == 0) -{ -lean_object* x_198; lean_object* x_199; lean_object* x_200; uint8_t x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; -x_198 = lean_ctor_get(x_175, 0); -lean_dec(x_198); -x_199 = l_Lean_IR_ToIR_lowerLet___closed__7; -x_200 = l_Lean_IR_ToIR_lowerLet___closed__9; -x_201 = 1; -x_202 = l_Lean_Name_toString(x_151, x_201); -lean_ctor_set_tag(x_175, 3); -lean_ctor_set(x_175, 0, x_202); -x_203 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_203, 0, x_200); -lean_ctor_set(x_203, 1, x_175); -x_204 = l_Lean_IR_ToIR_lowerLet___closed__11; -x_205 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_205, 0, x_203); -lean_ctor_set(x_205, 1, x_204); -x_206 = l_Lean_MessageData_ofFormat(x_205); -x_207 = l_Lean_throwNamedError___at___00Lean_IR_ToIR_lowerLet_spec__8___redArg(x_199, x_206, x_4, x_5); +x_200 = !lean_is_exclusive(x_178); +if (x_200 == 0) +{ +lean_object* x_201; lean_object* x_202; lean_object* x_203; uint8_t x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; +x_201 = lean_ctor_get(x_178, 0); +lean_dec(x_201); +x_202 = l_Lean_IR_ToIR_lowerLet___closed__7; +x_203 = l_Lean_IR_ToIR_lowerLet___closed__9; +x_204 = 1; +x_205 = l_Lean_Name_toString(x_154, x_204); +lean_ctor_set_tag(x_178, 3); +lean_ctor_set(x_178, 0, x_205); +x_206 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_206, 0, x_203); +lean_ctor_set(x_206, 1, x_178); +x_207 = l_Lean_IR_ToIR_lowerLet___closed__11; +x_208 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_208, 0, x_206); +lean_ctor_set(x_208, 1, x_207); +x_209 = l_Lean_MessageData_ofFormat(x_208); +x_210 = l_Lean_throwNamedError___at___00Lean_IR_ToIR_lowerLet_spec__8___redArg(x_202, x_209, x_4, x_5); lean_dec(x_5); lean_dec_ref(x_4); -return x_207; -} -else -{ -lean_object* x_208; lean_object* x_209; uint8_t x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; -lean_dec(x_175); -x_208 = l_Lean_IR_ToIR_lowerLet___closed__7; -x_209 = l_Lean_IR_ToIR_lowerLet___closed__9; -x_210 = 1; -x_211 = l_Lean_Name_toString(x_151, x_210); -x_212 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_212, 0, x_211); -x_213 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_213, 0, x_209); -lean_ctor_set(x_213, 1, x_212); -x_214 = l_Lean_IR_ToIR_lowerLet___closed__11; -x_215 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_215, 0, x_213); -lean_ctor_set(x_215, 1, x_214); -x_216 = l_Lean_MessageData_ofFormat(x_215); -x_217 = l_Lean_throwNamedError___at___00Lean_IR_ToIR_lowerLet_spec__8___redArg(x_208, x_216, x_4, x_5); +return x_210; +} +else +{ +lean_object* x_211; lean_object* x_212; uint8_t x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; +lean_dec(x_178); +x_211 = l_Lean_IR_ToIR_lowerLet___closed__7; +x_212 = l_Lean_IR_ToIR_lowerLet___closed__9; +x_213 = 1; +x_214 = l_Lean_Name_toString(x_154, x_213); +x_215 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_215, 0, x_214); +x_216 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_216, 0, x_212); +lean_ctor_set(x_216, 1, x_215); +x_217 = l_Lean_IR_ToIR_lowerLet___closed__11; +x_218 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_218, 0, x_216); +lean_ctor_set(x_218, 1, x_217); +x_219 = l_Lean_MessageData_ofFormat(x_218); +x_220 = l_Lean_throwNamedError___at___00Lean_IR_ToIR_lowerLet_spec__8___redArg(x_211, x_219, x_4, x_5); lean_dec(x_5); lean_dec_ref(x_4); -return x_217; +return x_220; } } case 4: { -uint8_t x_218; -lean_free_object(x_171); -lean_dec(x_156); -lean_dec_ref(x_152); -lean_free_object(x_14); +uint8_t x_221; +lean_free_object(x_174); +lean_dec(x_159); +lean_dec_ref(x_155); +lean_dec(x_19); lean_dec(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_218 = !lean_is_exclusive(x_175); -if (x_218 == 0) -{ -lean_object* x_219; lean_object* x_220; lean_object* x_221; uint8_t x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; -x_219 = lean_ctor_get(x_175, 0); -lean_dec(x_219); -x_220 = l_Lean_IR_ToIR_lowerLet___closed__7; -x_221 = l_Lean_IR_ToIR_lowerLet___closed__9; -x_222 = 1; -x_223 = l_Lean_Name_toString(x_151, x_222); -lean_ctor_set_tag(x_175, 3); -lean_ctor_set(x_175, 0, x_223); -x_224 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_224, 0, x_221); -lean_ctor_set(x_224, 1, x_175); -x_225 = l_Lean_IR_ToIR_lowerLet___closed__11; -x_226 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_226, 0, x_224); -lean_ctor_set(x_226, 1, x_225); -x_227 = l_Lean_MessageData_ofFormat(x_226); -x_228 = l_Lean_throwNamedError___at___00Lean_IR_ToIR_lowerLet_spec__8___redArg(x_220, x_227, x_4, x_5); +x_221 = !lean_is_exclusive(x_178); +if (x_221 == 0) +{ +lean_object* x_222; lean_object* x_223; lean_object* x_224; uint8_t x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; +x_222 = lean_ctor_get(x_178, 0); +lean_dec(x_222); +x_223 = l_Lean_IR_ToIR_lowerLet___closed__7; +x_224 = l_Lean_IR_ToIR_lowerLet___closed__9; +x_225 = 1; +x_226 = l_Lean_Name_toString(x_154, x_225); +lean_ctor_set_tag(x_178, 3); +lean_ctor_set(x_178, 0, x_226); +x_227 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_227, 0, x_224); +lean_ctor_set(x_227, 1, x_178); +x_228 = l_Lean_IR_ToIR_lowerLet___closed__11; +x_229 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_229, 0, x_227); +lean_ctor_set(x_229, 1, x_228); +x_230 = l_Lean_MessageData_ofFormat(x_229); +x_231 = l_Lean_throwNamedError___at___00Lean_IR_ToIR_lowerLet_spec__8___redArg(x_223, x_230, x_4, x_5); lean_dec(x_5); lean_dec_ref(x_4); -return x_228; -} -else -{ -lean_object* x_229; lean_object* x_230; uint8_t x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; -lean_dec(x_175); -x_229 = l_Lean_IR_ToIR_lowerLet___closed__7; -x_230 = l_Lean_IR_ToIR_lowerLet___closed__9; -x_231 = 1; -x_232 = l_Lean_Name_toString(x_151, x_231); -x_233 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_233, 0, x_232); -x_234 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_234, 0, x_230); -lean_ctor_set(x_234, 1, x_233); -x_235 = l_Lean_IR_ToIR_lowerLet___closed__11; -x_236 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_236, 0, x_234); -lean_ctor_set(x_236, 1, x_235); -x_237 = l_Lean_MessageData_ofFormat(x_236); -x_238 = l_Lean_throwNamedError___at___00Lean_IR_ToIR_lowerLet_spec__8___redArg(x_229, x_237, x_4, x_5); +return x_231; +} +else +{ +lean_object* x_232; lean_object* x_233; uint8_t x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; +lean_dec(x_178); +x_232 = l_Lean_IR_ToIR_lowerLet___closed__7; +x_233 = l_Lean_IR_ToIR_lowerLet___closed__9; +x_234 = 1; +x_235 = l_Lean_Name_toString(x_154, x_234); +x_236 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_236, 0, x_235); +x_237 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_237, 0, x_233); +lean_ctor_set(x_237, 1, x_236); +x_238 = l_Lean_IR_ToIR_lowerLet___closed__11; +x_239 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_239, 0, x_237); +lean_ctor_set(x_239, 1, x_238); +x_240 = l_Lean_MessageData_ofFormat(x_239); +x_241 = l_Lean_throwNamedError___at___00Lean_IR_ToIR_lowerLet_spec__8___redArg(x_232, x_240, x_4, x_5); lean_dec(x_5); lean_dec_ref(x_4); -return x_238; +return x_241; } } case 5: { -uint8_t x_239; -lean_free_object(x_171); -lean_dec(x_156); -lean_dec_ref(x_152); -lean_free_object(x_14); +uint8_t x_242; +lean_free_object(x_174); +lean_dec(x_159); +lean_dec_ref(x_155); +lean_dec(x_19); lean_dec(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_239 = !lean_is_exclusive(x_175); -if (x_239 == 0) -{ -lean_object* x_240; lean_object* x_241; lean_object* x_242; uint8_t x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; -x_240 = lean_ctor_get(x_175, 0); -lean_dec(x_240); -x_241 = l_Lean_IR_ToIR_lowerLet___closed__7; -x_242 = l_Lean_IR_ToIR_lowerLet___closed__9; -x_243 = 1; -x_244 = l_Lean_Name_toString(x_151, x_243); -lean_ctor_set_tag(x_175, 3); -lean_ctor_set(x_175, 0, x_244); -x_245 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_245, 0, x_242); -lean_ctor_set(x_245, 1, x_175); -x_246 = l_Lean_IR_ToIR_lowerLet___closed__11; -x_247 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_247, 0, x_245); -lean_ctor_set(x_247, 1, x_246); -x_248 = l_Lean_MessageData_ofFormat(x_247); -x_249 = l_Lean_throwNamedError___at___00Lean_IR_ToIR_lowerLet_spec__8___redArg(x_241, x_248, x_4, x_5); +x_242 = !lean_is_exclusive(x_178); +if (x_242 == 0) +{ +lean_object* x_243; lean_object* x_244; lean_object* x_245; uint8_t x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; +x_243 = lean_ctor_get(x_178, 0); +lean_dec(x_243); +x_244 = l_Lean_IR_ToIR_lowerLet___closed__7; +x_245 = l_Lean_IR_ToIR_lowerLet___closed__9; +x_246 = 1; +x_247 = l_Lean_Name_toString(x_154, x_246); +lean_ctor_set_tag(x_178, 3); +lean_ctor_set(x_178, 0, x_247); +x_248 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_248, 0, x_245); +lean_ctor_set(x_248, 1, x_178); +x_249 = l_Lean_IR_ToIR_lowerLet___closed__11; +x_250 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_250, 0, x_248); +lean_ctor_set(x_250, 1, x_249); +x_251 = l_Lean_MessageData_ofFormat(x_250); +x_252 = l_Lean_throwNamedError___at___00Lean_IR_ToIR_lowerLet_spec__8___redArg(x_244, x_251, x_4, x_5); lean_dec(x_5); lean_dec_ref(x_4); -return x_249; -} -else -{ -lean_object* x_250; lean_object* x_251; uint8_t x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; -lean_dec(x_175); -x_250 = l_Lean_IR_ToIR_lowerLet___closed__7; -x_251 = l_Lean_IR_ToIR_lowerLet___closed__9; -x_252 = 1; -x_253 = l_Lean_Name_toString(x_151, x_252); -x_254 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_254, 0, x_253); -x_255 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_255, 0, x_251); -lean_ctor_set(x_255, 1, x_254); -x_256 = l_Lean_IR_ToIR_lowerLet___closed__11; -x_257 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_257, 0, x_255); -lean_ctor_set(x_257, 1, x_256); -x_258 = l_Lean_MessageData_ofFormat(x_257); -x_259 = l_Lean_throwNamedError___at___00Lean_IR_ToIR_lowerLet_spec__8___redArg(x_250, x_258, x_4, x_5); +return x_252; +} +else +{ +lean_object* x_253; lean_object* x_254; uint8_t x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; +lean_dec(x_178); +x_253 = l_Lean_IR_ToIR_lowerLet___closed__7; +x_254 = l_Lean_IR_ToIR_lowerLet___closed__9; +x_255 = 1; +x_256 = l_Lean_Name_toString(x_154, x_255); +x_257 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_257, 0, x_256); +x_258 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_258, 0, x_254); +lean_ctor_set(x_258, 1, x_257); +x_259 = l_Lean_IR_ToIR_lowerLet___closed__11; +x_260 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_260, 0, x_258); +lean_ctor_set(x_260, 1, x_259); +x_261 = l_Lean_MessageData_ofFormat(x_260); +x_262 = l_Lean_throwNamedError___at___00Lean_IR_ToIR_lowerLet_spec__8___redArg(x_253, x_261, x_4, x_5); lean_dec(x_5); lean_dec_ref(x_4); -return x_259; +return x_262; } } case 6: { -uint8_t x_260; -lean_inc(x_16); +uint8_t x_263; +lean_inc_ref(x_16); +lean_inc(x_15); lean_dec_ref(x_1); -x_260 = !lean_is_exclusive(x_175); -if (x_260 == 0) -{ -lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; -x_261 = lean_ctor_get(x_175, 0); -x_262 = lean_ctor_get(x_261, 1); -lean_inc(x_262); -x_263 = lean_ctor_get(x_261, 2); -lean_inc(x_263); -x_264 = lean_ctor_get(x_261, 3); -lean_inc(x_264); -lean_dec_ref(x_261); +x_263 = !lean_is_exclusive(x_178); +if (x_263 == 0) +{ +lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; +x_264 = lean_ctor_get(x_178, 0); +x_265 = lean_ctor_get(x_264, 1); +lean_inc(x_265); +x_266 = lean_ctor_get(x_264, 2); +lean_inc(x_266); +x_267 = lean_ctor_get(x_264, 3); +lean_inc(x_267); +lean_dec_ref(x_264); lean_inc(x_5); lean_inc_ref(x_4); -lean_inc(x_262); -x_265 = l_Lean_IR_hasTrivialStructure_x3f(x_262, x_4, x_5); -if (lean_obj_tag(x_265) == 0) +lean_inc(x_265); +x_268 = l_Lean_IR_hasTrivialStructure_x3f(x_265, x_4, x_5); +if (lean_obj_tag(x_268) == 0) { -lean_object* x_266; -x_266 = lean_ctor_get(x_265, 0); -lean_inc(x_266); -lean_dec_ref(x_265); -if (lean_obj_tag(x_266) == 1) -{ -lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; uint8_t x_271; -lean_dec(x_264); -lean_dec(x_263); -lean_dec(x_262); -lean_free_object(x_175); -lean_free_object(x_171); -lean_dec(x_156); -lean_dec(x_151); -lean_free_object(x_14); -x_267 = lean_ctor_get(x_266, 0); -lean_inc(x_267); -lean_dec_ref(x_266); -x_268 = lean_st_ref_take(x_3); -x_269 = lean_ctor_get(x_267, 1); +lean_object* x_269; +x_269 = lean_ctor_get(x_268, 0); lean_inc(x_269); -x_270 = lean_ctor_get(x_267, 2); -lean_inc(x_270); -lean_dec(x_267); -x_271 = !lean_is_exclusive(x_268); -if (x_271 == 0) +lean_dec_ref(x_268); +if (lean_obj_tag(x_269) == 1) { -lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; -x_272 = lean_ctor_get(x_268, 3); -x_273 = lean_box(0); -x_274 = lean_nat_add(x_269, x_270); +lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; uint8_t x_274; +lean_dec(x_267); +lean_dec(x_266); +lean_dec(x_265); +lean_free_object(x_178); +lean_free_object(x_174); +lean_dec(x_159); +lean_dec(x_154); +lean_dec(x_19); +lean_dec_ref(x_16); +x_270 = lean_ctor_get(x_269, 0); +lean_inc(x_270); +lean_dec_ref(x_269); +x_271 = lean_st_ref_take(x_3); +x_272 = lean_ctor_get(x_270, 1); +lean_inc(x_272); +x_273 = lean_ctor_get(x_270, 2); +lean_inc(x_273); lean_dec(x_270); -lean_dec(x_269); -x_275 = lean_array_get(x_273, x_152, x_274); -lean_dec(x_274); -lean_dec_ref(x_152); -x_276 = l_Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_ToIR_lowerParam_spec__0___redArg(x_272, x_16, x_275); -lean_ctor_set(x_268, 3, x_276); -x_277 = lean_st_ref_set(x_3, x_268); -x_278 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); -return x_278; -} -else -{ -lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; -x_279 = lean_ctor_get(x_268, 0); -x_280 = lean_ctor_get(x_268, 1); -x_281 = lean_ctor_get(x_268, 2); -x_282 = lean_ctor_get(x_268, 3); +x_274 = !lean_is_exclusive(x_271); +if (x_274 == 0) +{ +lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; +x_275 = lean_ctor_get(x_271, 3); +x_276 = lean_box(0); +x_277 = lean_nat_add(x_272, x_273); +lean_dec(x_273); +lean_dec(x_272); +x_278 = lean_array_get(x_276, x_155, x_277); +lean_dec(x_277); +lean_dec_ref(x_155); +x_279 = l_Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_ToIR_lowerParam_spec__0___redArg(x_275, x_15, x_278); +lean_ctor_set(x_271, 3, x_279); +x_280 = lean_st_ref_set(x_3, x_271); +x_281 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); +return x_281; +} +else +{ +lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; +x_282 = lean_ctor_get(x_271, 0); +x_283 = lean_ctor_get(x_271, 1); +x_284 = lean_ctor_get(x_271, 2); +x_285 = lean_ctor_get(x_271, 3); +lean_inc(x_285); +lean_inc(x_284); +lean_inc(x_283); lean_inc(x_282); -lean_inc(x_281); -lean_inc(x_280); -lean_inc(x_279); -lean_dec(x_268); -x_283 = lean_box(0); -x_284 = lean_nat_add(x_269, x_270); -lean_dec(x_270); -lean_dec(x_269); -x_285 = lean_array_get(x_283, x_152, x_284); -lean_dec(x_284); -lean_dec_ref(x_152); -x_286 = l_Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_ToIR_lowerParam_spec__0___redArg(x_282, x_16, x_285); -x_287 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_287, 0, x_279); -lean_ctor_set(x_287, 1, x_280); -lean_ctor_set(x_287, 2, x_281); -lean_ctor_set(x_287, 3, x_286); -x_288 = lean_st_ref_set(x_3, x_287); -x_289 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); -return x_289; +lean_dec(x_271); +x_286 = lean_box(0); +x_287 = lean_nat_add(x_272, x_273); +lean_dec(x_273); +lean_dec(x_272); +x_288 = lean_array_get(x_286, x_155, x_287); +lean_dec(x_287); +lean_dec_ref(x_155); +x_289 = l_Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_ToIR_lowerParam_spec__0___redArg(x_285, x_15, x_288); +x_290 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_290, 0, x_282); +lean_ctor_set(x_290, 1, x_283); +lean_ctor_set(x_290, 2, x_284); +lean_ctor_set(x_290, 3, x_289); +x_291 = lean_st_ref_set(x_3, x_290); +x_292 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); +return x_292; } } else { -lean_object* x_290; -lean_dec(x_266); -lean_dec_ref(x_152); -lean_inc(x_5); -lean_inc_ref(x_4); -x_290 = l_Lean_IR_nameToIRType(x_262, x_4, x_5); -if (lean_obj_tag(x_290) == 0) -{ -lean_object* x_291; uint8_t x_292; -x_291 = lean_ctor_get(x_290, 0); -lean_inc(x_291); -lean_dec_ref(x_290); -x_292 = l_Lean_IR_IRType_isScalar(x_291); -if (x_292 == 0) -{ lean_object* x_293; -lean_dec(x_291); -lean_dec(x_263); -lean_free_object(x_175); -lean_free_object(x_171); +lean_dec(x_269); +lean_dec_ref(x_155); lean_inc(x_5); lean_inc_ref(x_4); -x_293 = l_Lean_IR_getCtorLayout(x_151, x_4, x_5); +lean_inc(x_265); +x_293 = l_Lean_IR_nameToIRType(x_265, x_4, x_5); if (lean_obj_tag(x_293) == 0) { -lean_object* x_294; lean_object* x_295; uint8_t x_299; +lean_object* x_294; uint8_t x_295; x_294 = lean_ctor_get(x_293, 0); lean_inc(x_294); -if (lean_is_exclusive(x_293)) { - lean_ctor_release(x_293, 0); - x_295 = x_293; -} else { - lean_dec_ref(x_293); - x_295 = lean_box(0); -} -x_299 = !lean_is_exclusive(x_294); -if (x_299 == 0) -{ -lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; uint8_t x_306; -x_300 = lean_ctor_get(x_294, 0); -x_301 = lean_ctor_get(x_294, 1); -x_302 = lean_array_get_size(x_156); -x_303 = l_Array_extract___redArg(x_156, x_264, x_302); -lean_dec(x_156); -x_304 = lean_array_get_size(x_303); -x_305 = lean_array_get_size(x_301); -x_306 = lean_nat_dec_eq(x_304, x_305); -if (x_306 == 0) +lean_dec_ref(x_293); +x_295 = l_Lean_IR_IRType_isScalar(x_294); +if (x_295 == 0) +{ +lean_object* x_296; +lean_dec(x_294); +lean_dec(x_266); +lean_free_object(x_178); +lean_free_object(x_174); +lean_inc(x_5); +lean_inc_ref(x_4); +x_296 = l_Lean_IR_getCtorLayout(x_154, x_4, x_5); +if (lean_obj_tag(x_296) == 0) +{ +lean_object* x_297; lean_object* x_298; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; uint8_t x_309; +x_297 = lean_ctor_get(x_296, 0); +lean_inc(x_297); +if (lean_is_exclusive(x_296)) { + lean_ctor_release(x_296, 0); + x_298 = x_296; +} else { + lean_dec_ref(x_296); + x_298 = lean_box(0); +} +x_302 = lean_ctor_get(x_297, 0); +lean_inc_ref(x_302); +x_303 = lean_ctor_get(x_297, 1); +lean_inc_ref(x_303); +if (lean_is_exclusive(x_297)) { + lean_ctor_release(x_297, 0); + lean_ctor_release(x_297, 1); + x_304 = x_297; +} else { + lean_dec_ref(x_297); + x_304 = lean_box(0); +} +x_305 = lean_array_get_size(x_159); +x_306 = l_Array_extract___redArg(x_159, x_267, x_305); +lean_dec(x_159); +x_307 = lean_array_get_size(x_306); +x_308 = lean_array_get_size(x_303); +x_309 = lean_nat_dec_eq(x_307, x_308); +if (x_309 == 0) { +lean_dec_ref(x_306); +lean_dec(x_304); lean_dec_ref(x_303); -lean_free_object(x_294); -lean_dec_ref(x_301); -lean_dec_ref(x_300); -lean_free_object(x_14); -lean_dec(x_16); +lean_dec_ref(x_302); +lean_dec(x_265); +lean_dec(x_19); +lean_dec_ref(x_16); +lean_dec(x_15); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -goto block_298; +goto block_301; +} +else +{ +if (x_295 == 0) +{ +lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; +lean_dec(x_298); +x_310 = l_Lean_IR_ToIR_getFVarValue___redArg___closed__0; +x_311 = lean_unsigned_to_nat(0u); +x_312 = l_Lean_IR_ToIR_lowerLet___closed__12; +x_313 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_ToIR_lowerLet_spec__9___redArg(x_308, x_303, x_310, x_306, x_311, x_312); +if (lean_obj_tag(x_313) == 0) +{ +lean_object* x_314; lean_object* x_315; +x_314 = lean_ctor_get(x_313, 0); +lean_inc(x_314); +lean_dec_ref(x_313); +x_315 = l_Lean_IR_ToIR_bindVar___redArg(x_15, x_3); +if (lean_obj_tag(x_315) == 0) +{ +lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_332; lean_object* x_333; uint8_t x_334; +x_316 = lean_ctor_get(x_315, 0); +lean_inc(x_316); +lean_dec_ref(x_315); +x_332 = lean_st_ref_get(x_5); +x_333 = lean_ctor_get(x_332, 0); +lean_inc_ref(x_333); +lean_dec(x_332); +x_334 = l_Lean_IR_UnboxResult_hasUnboxAttr(x_333, x_265); +if (x_334 == 0) +{ +lean_object* x_335; +lean_dec_ref(x_16); +x_335 = l_Lean_IR_CtorInfo_type(x_302); +x_317 = x_335; +x_318 = x_3; +x_319 = x_4; +x_320 = x_5; +x_321 = lean_box(0); +goto block_331; } else { -if (x_292 == 0) +lean_object* x_336; +lean_inc(x_5); +lean_inc_ref(x_4); +x_336 = l_Lean_IR_toIRType(x_16, x_4, x_5); +if (lean_obj_tag(x_336) == 0) +{ +lean_object* x_337; +x_337 = lean_ctor_get(x_336, 0); +lean_inc(x_337); +lean_dec_ref(x_336); +x_317 = x_337; +x_318 = x_3; +x_319 = x_4; +x_320 = x_5; +x_321 = lean_box(0); +goto block_331; +} +else +{ +uint8_t x_338; +lean_dec(x_316); +lean_dec(x_314); +lean_dec_ref(x_306); +lean_dec(x_304); +lean_dec_ref(x_303); +lean_dec_ref(x_302); +lean_dec(x_19); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec(x_3); +lean_dec_ref(x_2); +x_338 = !lean_is_exclusive(x_336); +if (x_338 == 0) { -lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; -lean_dec(x_295); -x_307 = l_Lean_IR_ToIR_getFVarValue___redArg___closed__0; -x_308 = lean_unsigned_to_nat(0u); -x_309 = l_Lean_IR_ToIR_lowerLet___closed__12; -x_310 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_ToIR_lowerLet_spec__9___redArg(x_305, x_301, x_307, x_303, x_308, x_309); -if (lean_obj_tag(x_310) == 0) +return x_336; +} +else { -lean_object* x_311; lean_object* x_312; -x_311 = lean_ctor_get(x_310, 0); -lean_inc(x_311); -lean_dec_ref(x_310); -x_312 = l_Lean_IR_ToIR_bindVar___redArg(x_16, x_3); -if (lean_obj_tag(x_312) == 0) +lean_object* x_339; lean_object* x_340; +x_339 = lean_ctor_get(x_336, 0); +lean_inc(x_339); +lean_dec(x_336); +x_340 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_340, 0, x_339); +return x_340; +} +} +} +block_331: { -lean_object* x_313; lean_object* x_314; -x_313 = lean_ctor_get(x_312, 0); -lean_inc(x_313); -lean_dec_ref(x_312); -lean_inc(x_313); -x_314 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_lowerNonObjectFields___redArg(x_2, x_300, x_301, x_303, x_313, x_3, x_4, x_5); +lean_object* x_322; +lean_inc(x_316); +lean_inc_ref(x_302); +x_322 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_lowerNonObjectFields___redArg(x_2, x_302, x_303, x_306, x_316, x_318, x_319, x_320); +lean_dec_ref(x_306); lean_dec_ref(x_303); -lean_dec_ref(x_301); -if (lean_obj_tag(x_314) == 0) +if (lean_obj_tag(x_322) == 0) { -uint8_t x_315; -x_315 = !lean_is_exclusive(x_314); -if (x_315 == 0) +uint8_t x_323; +x_323 = !lean_is_exclusive(x_322); +if (x_323 == 0) { -lean_object* x_316; lean_object* x_317; -x_316 = lean_ctor_get(x_314, 0); -x_317 = l_Lean_IR_CtorInfo_type(x_300); -lean_ctor_set(x_294, 1, x_311); -lean_ctor_set(x_14, 3, x_316); -lean_ctor_set(x_14, 2, x_294); -lean_ctor_set(x_14, 1, x_317); -lean_ctor_set(x_14, 0, x_313); -lean_ctor_set(x_314, 0, x_14); -return x_314; +lean_object* x_324; lean_object* x_325; lean_object* x_326; +x_324 = lean_ctor_get(x_322, 0); +if (lean_is_scalar(x_304)) { + x_325 = lean_alloc_ctor(0, 2, 0); +} else { + x_325 = x_304; +} +lean_ctor_set(x_325, 0, x_302); +lean_ctor_set(x_325, 1, x_314); +if (lean_is_scalar(x_19)) { + x_326 = lean_alloc_ctor(0, 4, 0); +} else { + x_326 = x_19; +} +lean_ctor_set(x_326, 0, x_316); +lean_ctor_set(x_326, 1, x_317); +lean_ctor_set(x_326, 2, x_325); +lean_ctor_set(x_326, 3, x_324); +lean_ctor_set(x_322, 0, x_326); +return x_322; } else { -lean_object* x_318; lean_object* x_319; lean_object* x_320; -x_318 = lean_ctor_get(x_314, 0); -lean_inc(x_318); -lean_dec(x_314); -x_319 = l_Lean_IR_CtorInfo_type(x_300); -lean_ctor_set(x_294, 1, x_311); -lean_ctor_set(x_14, 3, x_318); -lean_ctor_set(x_14, 2, x_294); -lean_ctor_set(x_14, 1, x_319); -lean_ctor_set(x_14, 0, x_313); -x_320 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_320, 0, x_14); -return x_320; +lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; +x_327 = lean_ctor_get(x_322, 0); +lean_inc(x_327); +lean_dec(x_322); +if (lean_is_scalar(x_304)) { + x_328 = lean_alloc_ctor(0, 2, 0); +} else { + x_328 = x_304; +} +lean_ctor_set(x_328, 0, x_302); +lean_ctor_set(x_328, 1, x_314); +if (lean_is_scalar(x_19)) { + x_329 = lean_alloc_ctor(0, 4, 0); +} else { + x_329 = x_19; +} +lean_ctor_set(x_329, 0, x_316); +lean_ctor_set(x_329, 1, x_317); +lean_ctor_set(x_329, 2, x_328); +lean_ctor_set(x_329, 3, x_327); +x_330 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_330, 0, x_329); +return x_330; } } else { -lean_dec(x_313); -lean_dec(x_311); -lean_free_object(x_294); -lean_dec_ref(x_300); -lean_free_object(x_14); -return x_314; +lean_dec(x_317); +lean_dec(x_316); +lean_dec(x_314); +lean_dec(x_304); +lean_dec_ref(x_302); +lean_dec(x_19); +return x_322; +} } } else { -uint8_t x_321; -lean_dec(x_311); +uint8_t x_341; +lean_dec(x_314); +lean_dec_ref(x_306); +lean_dec(x_304); lean_dec_ref(x_303); -lean_free_object(x_294); -lean_dec_ref(x_301); -lean_dec_ref(x_300); -lean_free_object(x_14); +lean_dec_ref(x_302); +lean_dec(x_265); +lean_dec(x_19); +lean_dec_ref(x_16); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -x_321 = !lean_is_exclusive(x_312); -if (x_321 == 0) +x_341 = !lean_is_exclusive(x_315); +if (x_341 == 0) { -return x_312; +return x_315; } else { -lean_object* x_322; lean_object* x_323; -x_322 = lean_ctor_get(x_312, 0); -lean_inc(x_322); -lean_dec(x_312); -x_323 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_323, 0, x_322); -return x_323; +lean_object* x_342; lean_object* x_343; +x_342 = lean_ctor_get(x_315, 0); +lean_inc(x_342); +lean_dec(x_315); +x_343 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_343, 0, x_342); +return x_343; } } } else { -uint8_t x_324; +uint8_t x_344; +lean_dec_ref(x_306); +lean_dec(x_304); lean_dec_ref(x_303); -lean_free_object(x_294); -lean_dec_ref(x_301); -lean_dec_ref(x_300); -lean_free_object(x_14); -lean_dec(x_16); +lean_dec_ref(x_302); +lean_dec(x_265); +lean_dec(x_19); +lean_dec_ref(x_16); +lean_dec(x_15); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -x_324 = !lean_is_exclusive(x_310); -if (x_324 == 0) +x_344 = !lean_is_exclusive(x_313); +if (x_344 == 0) { -return x_310; +return x_313; } else { -lean_object* x_325; lean_object* x_326; -x_325 = lean_ctor_get(x_310, 0); -lean_inc(x_325); -lean_dec(x_310); -x_326 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_326, 0, x_325); -return x_326; +lean_object* x_345; lean_object* x_346; +x_345 = lean_ctor_get(x_313, 0); +lean_inc(x_345); +lean_dec(x_313); +x_346 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_346, 0, x_345); +return x_346; } } } else { +lean_dec_ref(x_306); +lean_dec(x_304); lean_dec_ref(x_303); -lean_free_object(x_294); -lean_dec_ref(x_301); -lean_dec_ref(x_300); -lean_free_object(x_14); -lean_dec(x_16); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -goto block_298; -} -} -} -else -{ -lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; uint8_t x_333; -x_327 = lean_ctor_get(x_294, 0); -x_328 = lean_ctor_get(x_294, 1); -lean_inc(x_328); -lean_inc(x_327); -lean_dec(x_294); -x_329 = lean_array_get_size(x_156); -x_330 = l_Array_extract___redArg(x_156, x_264, x_329); -lean_dec(x_156); -x_331 = lean_array_get_size(x_330); -x_332 = lean_array_get_size(x_328); -x_333 = lean_nat_dec_eq(x_331, x_332); -if (x_333 == 0) -{ -lean_dec_ref(x_330); -lean_dec_ref(x_328); -lean_dec_ref(x_327); -lean_free_object(x_14); -lean_dec(x_16); +lean_dec_ref(x_302); +lean_dec(x_265); +lean_dec(x_19); +lean_dec_ref(x_16); +lean_dec(x_15); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -goto block_298; -} -else -{ -if (x_292 == 0) -{ -lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; -lean_dec(x_295); -x_334 = l_Lean_IR_ToIR_getFVarValue___redArg___closed__0; -x_335 = lean_unsigned_to_nat(0u); -x_336 = l_Lean_IR_ToIR_lowerLet___closed__12; -x_337 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_ToIR_lowerLet_spec__9___redArg(x_332, x_328, x_334, x_330, x_335, x_336); -if (lean_obj_tag(x_337) == 0) -{ -lean_object* x_338; lean_object* x_339; -x_338 = lean_ctor_get(x_337, 0); -lean_inc(x_338); -lean_dec_ref(x_337); -x_339 = l_Lean_IR_ToIR_bindVar___redArg(x_16, x_3); -if (lean_obj_tag(x_339) == 0) -{ -lean_object* x_340; lean_object* x_341; -x_340 = lean_ctor_get(x_339, 0); -lean_inc(x_340); -lean_dec_ref(x_339); -lean_inc(x_340); -x_341 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_lowerNonObjectFields___redArg(x_2, x_327, x_328, x_330, x_340, x_3, x_4, x_5); -lean_dec_ref(x_330); -lean_dec_ref(x_328); -if (lean_obj_tag(x_341) == 0) -{ -lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; -x_342 = lean_ctor_get(x_341, 0); -lean_inc(x_342); -if (lean_is_exclusive(x_341)) { - lean_ctor_release(x_341, 0); - x_343 = x_341; -} else { - lean_dec_ref(x_341); - x_343 = lean_box(0); -} -x_344 = l_Lean_IR_CtorInfo_type(x_327); -x_345 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_345, 0, x_327); -lean_ctor_set(x_345, 1, x_338); -lean_ctor_set(x_14, 3, x_342); -lean_ctor_set(x_14, 2, x_345); -lean_ctor_set(x_14, 1, x_344); -lean_ctor_set(x_14, 0, x_340); -if (lean_is_scalar(x_343)) { - x_346 = lean_alloc_ctor(0, 1, 0); -} else { - x_346 = x_343; -} -lean_ctor_set(x_346, 0, x_14); -return x_346; -} -else -{ -lean_dec(x_340); -lean_dec(x_338); -lean_dec_ref(x_327); -lean_free_object(x_14); -return x_341; +goto block_301; } } -else +block_301: { -lean_object* x_347; lean_object* x_348; lean_object* x_349; -lean_dec(x_338); -lean_dec_ref(x_330); -lean_dec_ref(x_328); -lean_dec_ref(x_327); -lean_free_object(x_14); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -x_347 = lean_ctor_get(x_339, 0); -lean_inc(x_347); -if (lean_is_exclusive(x_339)) { - lean_ctor_release(x_339, 0); - x_348 = x_339; -} else { - lean_dec_ref(x_339); - x_348 = lean_box(0); -} -if (lean_is_scalar(x_348)) { - x_349 = lean_alloc_ctor(1, 1, 0); +lean_object* x_299; lean_object* x_300; +x_299 = lean_box(12); +if (lean_is_scalar(x_298)) { + x_300 = lean_alloc_ctor(0, 1, 0); } else { - x_349 = x_348; + x_300 = x_298; } -lean_ctor_set(x_349, 0, x_347); -return x_349; +lean_ctor_set(x_300, 0, x_299); +return x_300; } } else { -lean_object* x_350; lean_object* x_351; lean_object* x_352; -lean_dec_ref(x_330); -lean_dec_ref(x_328); -lean_dec_ref(x_327); -lean_free_object(x_14); -lean_dec(x_16); +uint8_t x_347; +lean_dec(x_267); +lean_dec(x_265); +lean_dec(x_159); +lean_dec(x_19); +lean_dec_ref(x_16); +lean_dec(x_15); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -x_350 = lean_ctor_get(x_337, 0); -lean_inc(x_350); -if (lean_is_exclusive(x_337)) { - lean_ctor_release(x_337, 0); - x_351 = x_337; -} else { - lean_dec_ref(x_337); - x_351 = lean_box(0); -} -if (lean_is_scalar(x_351)) { - x_352 = lean_alloc_ctor(1, 1, 0); -} else { - x_352 = x_351; -} -lean_ctor_set(x_352, 0, x_350); -return x_352; -} -} -else +x_347 = !lean_is_exclusive(x_296); +if (x_347 == 0) { -lean_dec_ref(x_330); -lean_dec_ref(x_328); -lean_dec_ref(x_327); -lean_free_object(x_14); -lean_dec(x_16); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -goto block_298; -} -} +return x_296; } -block_298: +else { -lean_object* x_296; lean_object* x_297; -x_296 = lean_box(12); -if (lean_is_scalar(x_295)) { - x_297 = lean_alloc_ctor(0, 1, 0); -} else { - x_297 = x_295; +lean_object* x_348; lean_object* x_349; +x_348 = lean_ctor_get(x_296, 0); +lean_inc(x_348); +lean_dec(x_296); +x_349 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_349, 0, x_348); +return x_349; } -lean_ctor_set(x_297, 0, x_296); -return x_297; } } else { +lean_object* x_350; +lean_dec(x_267); +lean_dec(x_265); +lean_dec(x_159); +lean_dec(x_154); +lean_dec_ref(x_16); +x_350 = l_Lean_IR_ToIR_bindVar___redArg(x_15, x_3); +if (lean_obj_tag(x_350) == 0) +{ +lean_object* x_351; lean_object* x_352; +x_351 = lean_ctor_get(x_350, 0); +lean_inc(x_351); +lean_dec_ref(x_350); +x_352 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); +if (lean_obj_tag(x_352) == 0) +{ uint8_t x_353; -lean_dec(x_264); -lean_dec(x_156); -lean_free_object(x_14); -lean_dec(x_16); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -x_353 = !lean_is_exclusive(x_293); +x_353 = !lean_is_exclusive(x_352); if (x_353 == 0) { -return x_293; -} -else -{ lean_object* x_354; lean_object* x_355; -x_354 = lean_ctor_get(x_293, 0); -lean_inc(x_354); -lean_dec(x_293); -x_355 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_355, 0, x_354); -return x_355; -} -} +x_354 = lean_ctor_get(x_352, 0); +lean_ctor_set_tag(x_178, 0); +lean_ctor_set(x_178, 0, x_266); +lean_ctor_set_tag(x_174, 11); +if (lean_is_scalar(x_19)) { + x_355 = lean_alloc_ctor(0, 4, 0); +} else { + x_355 = x_19; +} +lean_ctor_set(x_355, 0, x_351); +lean_ctor_set(x_355, 1, x_294); +lean_ctor_set(x_355, 2, x_174); +lean_ctor_set(x_355, 3, x_354); +lean_ctor_set(x_352, 0, x_355); +return x_352; } else { -lean_object* x_356; -lean_dec(x_264); -lean_dec(x_156); -lean_dec(x_151); -x_356 = l_Lean_IR_ToIR_bindVar___redArg(x_16, x_3); -if (lean_obj_tag(x_356) == 0) -{ -lean_object* x_357; lean_object* x_358; -x_357 = lean_ctor_get(x_356, 0); -lean_inc(x_357); -lean_dec_ref(x_356); -x_358 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); -if (lean_obj_tag(x_358) == 0) -{ -uint8_t x_359; -x_359 = !lean_is_exclusive(x_358); -if (x_359 == 0) -{ -lean_object* x_360; -x_360 = lean_ctor_get(x_358, 0); -lean_ctor_set_tag(x_175, 0); -lean_ctor_set(x_175, 0, x_263); -lean_ctor_set_tag(x_171, 11); -lean_ctor_set(x_14, 3, x_360); -lean_ctor_set(x_14, 2, x_171); -lean_ctor_set(x_14, 1, x_291); -lean_ctor_set(x_14, 0, x_357); -lean_ctor_set(x_358, 0, x_14); -return x_358; +lean_object* x_356; lean_object* x_357; lean_object* x_358; +x_356 = lean_ctor_get(x_352, 0); +lean_inc(x_356); +lean_dec(x_352); +lean_ctor_set_tag(x_178, 0); +lean_ctor_set(x_178, 0, x_266); +lean_ctor_set_tag(x_174, 11); +if (lean_is_scalar(x_19)) { + x_357 = lean_alloc_ctor(0, 4, 0); +} else { + x_357 = x_19; } -else -{ -lean_object* x_361; lean_object* x_362; -x_361 = lean_ctor_get(x_358, 0); -lean_inc(x_361); -lean_dec(x_358); -lean_ctor_set_tag(x_175, 0); -lean_ctor_set(x_175, 0, x_263); -lean_ctor_set_tag(x_171, 11); -lean_ctor_set(x_14, 3, x_361); -lean_ctor_set(x_14, 2, x_171); -lean_ctor_set(x_14, 1, x_291); -lean_ctor_set(x_14, 0, x_357); -x_362 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_362, 0, x_14); -return x_362; +lean_ctor_set(x_357, 0, x_351); +lean_ctor_set(x_357, 1, x_294); +lean_ctor_set(x_357, 2, x_174); +lean_ctor_set(x_357, 3, x_356); +x_358 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_358, 0, x_357); +return x_358; } } else { -lean_dec(x_357); -lean_dec(x_291); -lean_dec(x_263); -lean_free_object(x_175); -lean_free_object(x_171); -lean_free_object(x_14); -return x_358; +lean_dec(x_351); +lean_dec(x_294); +lean_dec(x_266); +lean_free_object(x_178); +lean_free_object(x_174); +lean_dec(x_19); +return x_352; } } else { -uint8_t x_363; -lean_dec(x_291); -lean_dec(x_263); -lean_free_object(x_175); -lean_free_object(x_171); -lean_free_object(x_14); +uint8_t x_359; +lean_dec(x_294); +lean_dec(x_266); +lean_free_object(x_178); +lean_free_object(x_174); +lean_dec(x_19); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -x_363 = !lean_is_exclusive(x_356); -if (x_363 == 0) +x_359 = !lean_is_exclusive(x_350); +if (x_359 == 0) { -return x_356; +return x_350; } else { -lean_object* x_364; lean_object* x_365; -x_364 = lean_ctor_get(x_356, 0); -lean_inc(x_364); -lean_dec(x_356); -x_365 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_365, 0, x_364); -return x_365; +lean_object* x_360; lean_object* x_361; +x_360 = lean_ctor_get(x_350, 0); +lean_inc(x_360); +lean_dec(x_350); +x_361 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_361, 0, x_360); +return x_361; } } } } else { -uint8_t x_366; -lean_dec(x_264); -lean_dec(x_263); -lean_free_object(x_175); -lean_free_object(x_171); -lean_dec(x_156); -lean_dec(x_151); -lean_free_object(x_14); -lean_dec(x_16); +uint8_t x_362; +lean_dec(x_267); +lean_dec(x_266); +lean_dec(x_265); +lean_free_object(x_178); +lean_free_object(x_174); +lean_dec(x_159); +lean_dec(x_154); +lean_dec(x_19); +lean_dec_ref(x_16); +lean_dec(x_15); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -x_366 = !lean_is_exclusive(x_290); -if (x_366 == 0) +x_362 = !lean_is_exclusive(x_293); +if (x_362 == 0) { -return x_290; +return x_293; } else { -lean_object* x_367; lean_object* x_368; -x_367 = lean_ctor_get(x_290, 0); -lean_inc(x_367); -lean_dec(x_290); -x_368 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_368, 0, x_367); -return x_368; +lean_object* x_363; lean_object* x_364; +x_363 = lean_ctor_get(x_293, 0); +lean_inc(x_363); +lean_dec(x_293); +x_364 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_364, 0, x_363); +return x_364; } } } } else { -uint8_t x_369; -lean_dec(x_264); -lean_dec(x_263); -lean_dec(x_262); -lean_free_object(x_175); -lean_free_object(x_171); -lean_dec(x_156); -lean_dec_ref(x_152); -lean_dec(x_151); -lean_free_object(x_14); -lean_dec(x_16); +uint8_t x_365; +lean_dec(x_267); +lean_dec(x_266); +lean_dec(x_265); +lean_free_object(x_178); +lean_free_object(x_174); +lean_dec(x_159); +lean_dec_ref(x_155); +lean_dec(x_154); +lean_dec(x_19); +lean_dec_ref(x_16); +lean_dec(x_15); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -x_369 = !lean_is_exclusive(x_265); -if (x_369 == 0) +x_365 = !lean_is_exclusive(x_268); +if (x_365 == 0) { -return x_265; +return x_268; } else { -lean_object* x_370; lean_object* x_371; -x_370 = lean_ctor_get(x_265, 0); -lean_inc(x_370); -lean_dec(x_265); -x_371 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_371, 0, x_370); -return x_371; +lean_object* x_366; lean_object* x_367; +x_366 = lean_ctor_get(x_268, 0); +lean_inc(x_366); +lean_dec(x_268); +x_367 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_367, 0, x_366); +return x_367; } } } else { -lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; -x_372 = lean_ctor_get(x_175, 0); -lean_inc(x_372); -lean_dec(x_175); -x_373 = lean_ctor_get(x_372, 1); -lean_inc(x_373); -x_374 = lean_ctor_get(x_372, 2); -lean_inc(x_374); -x_375 = lean_ctor_get(x_372, 3); -lean_inc(x_375); -lean_dec_ref(x_372); +lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; +x_368 = lean_ctor_get(x_178, 0); +lean_inc(x_368); +lean_dec(x_178); +x_369 = lean_ctor_get(x_368, 1); +lean_inc(x_369); +x_370 = lean_ctor_get(x_368, 2); +lean_inc(x_370); +x_371 = lean_ctor_get(x_368, 3); +lean_inc(x_371); +lean_dec_ref(x_368); lean_inc(x_5); lean_inc_ref(x_4); +lean_inc(x_369); +x_372 = l_Lean_IR_hasTrivialStructure_x3f(x_369, x_4, x_5); +if (lean_obj_tag(x_372) == 0) +{ +lean_object* x_373; +x_373 = lean_ctor_get(x_372, 0); lean_inc(x_373); -x_376 = l_Lean_IR_hasTrivialStructure_x3f(x_373, x_4, x_5); -if (lean_obj_tag(x_376) == 0) +lean_dec_ref(x_372); +if (lean_obj_tag(x_373) == 1) { -lean_object* x_377; -x_377 = lean_ctor_get(x_376, 0); +lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; +lean_dec(x_371); +lean_dec(x_370); +lean_dec(x_369); +lean_free_object(x_174); +lean_dec(x_159); +lean_dec(x_154); +lean_dec(x_19); +lean_dec_ref(x_16); +x_374 = lean_ctor_get(x_373, 0); +lean_inc(x_374); +lean_dec_ref(x_373); +x_375 = lean_st_ref_take(x_3); +x_376 = lean_ctor_get(x_374, 1); +lean_inc(x_376); +x_377 = lean_ctor_get(x_374, 2); lean_inc(x_377); -lean_dec_ref(x_376); -if (lean_obj_tag(x_377) == 1) -{ -lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; -lean_dec(x_375); lean_dec(x_374); -lean_dec(x_373); -lean_free_object(x_171); -lean_dec(x_156); -lean_dec(x_151); -lean_free_object(x_14); -x_378 = lean_ctor_get(x_377, 0); -lean_inc(x_378); -lean_dec_ref(x_377); -x_379 = lean_st_ref_take(x_3); -x_380 = lean_ctor_get(x_378, 1); +x_378 = lean_ctor_get(x_375, 0); +lean_inc_ref(x_378); +x_379 = lean_ctor_get(x_375, 1); +lean_inc_ref(x_379); +x_380 = lean_ctor_get(x_375, 2); lean_inc(x_380); -x_381 = lean_ctor_get(x_378, 2); -lean_inc(x_381); -lean_dec(x_378); -x_382 = lean_ctor_get(x_379, 0); -lean_inc_ref(x_382); -x_383 = lean_ctor_get(x_379, 1); -lean_inc_ref(x_383); -x_384 = lean_ctor_get(x_379, 2); -lean_inc(x_384); -x_385 = lean_ctor_get(x_379, 3); -lean_inc_ref(x_385); -if (lean_is_exclusive(x_379)) { - lean_ctor_release(x_379, 0); - lean_ctor_release(x_379, 1); - lean_ctor_release(x_379, 2); - lean_ctor_release(x_379, 3); - x_386 = x_379; -} else { - lean_dec_ref(x_379); - x_386 = lean_box(0); -} -x_387 = lean_box(0); -x_388 = lean_nat_add(x_380, x_381); -lean_dec(x_381); -lean_dec(x_380); -x_389 = lean_array_get(x_387, x_152, x_388); -lean_dec(x_388); -lean_dec_ref(x_152); -x_390 = l_Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_ToIR_lowerParam_spec__0___redArg(x_385, x_16, x_389); -if (lean_is_scalar(x_386)) { - x_391 = lean_alloc_ctor(0, 4, 0); -} else { - x_391 = x_386; -} -lean_ctor_set(x_391, 0, x_382); -lean_ctor_set(x_391, 1, x_383); -lean_ctor_set(x_391, 2, x_384); -lean_ctor_set(x_391, 3, x_390); -x_392 = lean_st_ref_set(x_3, x_391); -x_393 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); -return x_393; -} -else -{ -lean_object* x_394; +x_381 = lean_ctor_get(x_375, 3); +lean_inc_ref(x_381); +if (lean_is_exclusive(x_375)) { + lean_ctor_release(x_375, 0); + lean_ctor_release(x_375, 1); + lean_ctor_release(x_375, 2); + lean_ctor_release(x_375, 3); + x_382 = x_375; +} else { + lean_dec_ref(x_375); + x_382 = lean_box(0); +} +x_383 = lean_box(0); +x_384 = lean_nat_add(x_376, x_377); lean_dec(x_377); -lean_dec_ref(x_152); +lean_dec(x_376); +x_385 = lean_array_get(x_383, x_155, x_384); +lean_dec(x_384); +lean_dec_ref(x_155); +x_386 = l_Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_ToIR_lowerParam_spec__0___redArg(x_381, x_15, x_385); +if (lean_is_scalar(x_382)) { + x_387 = lean_alloc_ctor(0, 4, 0); +} else { + x_387 = x_382; +} +lean_ctor_set(x_387, 0, x_378); +lean_ctor_set(x_387, 1, x_379); +lean_ctor_set(x_387, 2, x_380); +lean_ctor_set(x_387, 3, x_386); +x_388 = lean_st_ref_set(x_3, x_387); +x_389 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); +return x_389; +} +else +{ +lean_object* x_390; +lean_dec(x_373); +lean_dec_ref(x_155); lean_inc(x_5); lean_inc_ref(x_4); -x_394 = l_Lean_IR_nameToIRType(x_373, x_4, x_5); -if (lean_obj_tag(x_394) == 0) -{ -lean_object* x_395; uint8_t x_396; -x_395 = lean_ctor_get(x_394, 0); -lean_inc(x_395); -lean_dec_ref(x_394); -x_396 = l_Lean_IR_IRType_isScalar(x_395); -if (x_396 == 0) -{ -lean_object* x_397; -lean_dec(x_395); -lean_dec(x_374); -lean_free_object(x_171); +lean_inc(x_369); +x_390 = l_Lean_IR_nameToIRType(x_369, x_4, x_5); +if (lean_obj_tag(x_390) == 0) +{ +lean_object* x_391; uint8_t x_392; +x_391 = lean_ctor_get(x_390, 0); +lean_inc(x_391); +lean_dec_ref(x_390); +x_392 = l_Lean_IR_IRType_isScalar(x_391); +if (x_392 == 0) +{ +lean_object* x_393; +lean_dec(x_391); +lean_dec(x_370); +lean_free_object(x_174); lean_inc(x_5); lean_inc_ref(x_4); -x_397 = l_Lean_IR_getCtorLayout(x_151, x_4, x_5); -if (lean_obj_tag(x_397) == 0) -{ -lean_object* x_398; lean_object* x_399; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; uint8_t x_410; -x_398 = lean_ctor_get(x_397, 0); -lean_inc(x_398); -if (lean_is_exclusive(x_397)) { - lean_ctor_release(x_397, 0); - x_399 = x_397; -} else { - lean_dec_ref(x_397); - x_399 = lean_box(0); -} -x_403 = lean_ctor_get(x_398, 0); -lean_inc_ref(x_403); -x_404 = lean_ctor_get(x_398, 1); -lean_inc_ref(x_404); -if (lean_is_exclusive(x_398)) { - lean_ctor_release(x_398, 0); - lean_ctor_release(x_398, 1); - x_405 = x_398; -} else { - lean_dec_ref(x_398); - x_405 = lean_box(0); -} -x_406 = lean_array_get_size(x_156); -x_407 = l_Array_extract___redArg(x_156, x_375, x_406); -lean_dec(x_156); -x_408 = lean_array_get_size(x_407); -x_409 = lean_array_get_size(x_404); -x_410 = lean_nat_dec_eq(x_408, x_409); -if (x_410 == 0) -{ -lean_dec_ref(x_407); -lean_dec(x_405); -lean_dec_ref(x_404); +x_393 = l_Lean_IR_getCtorLayout(x_154, x_4, x_5); +if (lean_obj_tag(x_393) == 0) +{ +lean_object* x_394; lean_object* x_395; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; uint8_t x_406; +x_394 = lean_ctor_get(x_393, 0); +lean_inc(x_394); +if (lean_is_exclusive(x_393)) { + lean_ctor_release(x_393, 0); + x_395 = x_393; +} else { + lean_dec_ref(x_393); + x_395 = lean_box(0); +} +x_399 = lean_ctor_get(x_394, 0); +lean_inc_ref(x_399); +x_400 = lean_ctor_get(x_394, 1); +lean_inc_ref(x_400); +if (lean_is_exclusive(x_394)) { + lean_ctor_release(x_394, 0); + lean_ctor_release(x_394, 1); + x_401 = x_394; +} else { + lean_dec_ref(x_394); + x_401 = lean_box(0); +} +x_402 = lean_array_get_size(x_159); +x_403 = l_Array_extract___redArg(x_159, x_371, x_402); +lean_dec(x_159); +x_404 = lean_array_get_size(x_403); +x_405 = lean_array_get_size(x_400); +x_406 = lean_nat_dec_eq(x_404, x_405); +if (x_406 == 0) +{ lean_dec_ref(x_403); -lean_free_object(x_14); -lean_dec(x_16); +lean_dec(x_401); +lean_dec_ref(x_400); +lean_dec_ref(x_399); +lean_dec(x_369); +lean_dec(x_19); +lean_dec_ref(x_16); +lean_dec(x_15); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -goto block_402; -} -else -{ -if (x_396 == 0) -{ -lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; -lean_dec(x_399); -x_411 = l_Lean_IR_ToIR_getFVarValue___redArg___closed__0; -x_412 = lean_unsigned_to_nat(0u); -x_413 = l_Lean_IR_ToIR_lowerLet___closed__12; -x_414 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_ToIR_lowerLet_spec__9___redArg(x_409, x_404, x_411, x_407, x_412, x_413); -if (lean_obj_tag(x_414) == 0) -{ -lean_object* x_415; lean_object* x_416; -x_415 = lean_ctor_get(x_414, 0); -lean_inc(x_415); -lean_dec_ref(x_414); -x_416 = l_Lean_IR_ToIR_bindVar___redArg(x_16, x_3); -if (lean_obj_tag(x_416) == 0) -{ -lean_object* x_417; lean_object* x_418; -x_417 = lean_ctor_get(x_416, 0); -lean_inc(x_417); -lean_dec_ref(x_416); -lean_inc(x_417); -x_418 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_lowerNonObjectFields___redArg(x_2, x_403, x_404, x_407, x_417, x_3, x_4, x_5); -lean_dec_ref(x_407); -lean_dec_ref(x_404); -if (lean_obj_tag(x_418) == 0) -{ -lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; -x_419 = lean_ctor_get(x_418, 0); -lean_inc(x_419); -if (lean_is_exclusive(x_418)) { - lean_ctor_release(x_418, 0); - x_420 = x_418; -} else { - lean_dec_ref(x_418); - x_420 = lean_box(0); -} -x_421 = l_Lean_IR_CtorInfo_type(x_403); -if (lean_is_scalar(x_405)) { - x_422 = lean_alloc_ctor(0, 2, 0); -} else { - x_422 = x_405; -} -lean_ctor_set(x_422, 0, x_403); -lean_ctor_set(x_422, 1, x_415); -lean_ctor_set(x_14, 3, x_419); -lean_ctor_set(x_14, 2, x_422); -lean_ctor_set(x_14, 1, x_421); -lean_ctor_set(x_14, 0, x_417); -if (lean_is_scalar(x_420)) { - x_423 = lean_alloc_ctor(0, 1, 0); -} else { - x_423 = x_420; -} -lean_ctor_set(x_423, 0, x_14); -return x_423; +goto block_398; } else { -lean_dec(x_417); -lean_dec(x_415); -lean_dec(x_405); -lean_dec_ref(x_403); -lean_free_object(x_14); -return x_418; -} +if (x_392 == 0) +{ +lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; +lean_dec(x_395); +x_407 = l_Lean_IR_ToIR_getFVarValue___redArg___closed__0; +x_408 = lean_unsigned_to_nat(0u); +x_409 = l_Lean_IR_ToIR_lowerLet___closed__12; +x_410 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_ToIR_lowerLet_spec__9___redArg(x_405, x_400, x_407, x_403, x_408, x_409); +if (lean_obj_tag(x_410) == 0) +{ +lean_object* x_411; lean_object* x_412; +x_411 = lean_ctor_get(x_410, 0); +lean_inc(x_411); +lean_dec_ref(x_410); +x_412 = l_Lean_IR_ToIR_bindVar___redArg(x_15, x_3); +if (lean_obj_tag(x_412) == 0) +{ +lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_426; lean_object* x_427; uint8_t x_428; +x_413 = lean_ctor_get(x_412, 0); +lean_inc(x_413); +lean_dec_ref(x_412); +x_426 = lean_st_ref_get(x_5); +x_427 = lean_ctor_get(x_426, 0); +lean_inc_ref(x_427); +lean_dec(x_426); +x_428 = l_Lean_IR_UnboxResult_hasUnboxAttr(x_427, x_369); +if (x_428 == 0) +{ +lean_object* x_429; +lean_dec_ref(x_16); +x_429 = l_Lean_IR_CtorInfo_type(x_399); +x_414 = x_429; +x_415 = x_3; +x_416 = x_4; +x_417 = x_5; +x_418 = lean_box(0); +goto block_425; } else { -lean_object* x_424; lean_object* x_425; lean_object* x_426; -lean_dec(x_415); -lean_dec_ref(x_407); -lean_dec(x_405); -lean_dec_ref(x_404); -lean_dec_ref(x_403); -lean_free_object(x_14); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -x_424 = lean_ctor_get(x_416, 0); -lean_inc(x_424); -if (lean_is_exclusive(x_416)) { - lean_ctor_release(x_416, 0); - x_425 = x_416; -} else { - lean_dec_ref(x_416); - x_425 = lean_box(0); -} -if (lean_is_scalar(x_425)) { - x_426 = lean_alloc_ctor(1, 1, 0); -} else { - x_426 = x_425; -} -lean_ctor_set(x_426, 0, x_424); -return x_426; -} +lean_object* x_430; +lean_inc(x_5); +lean_inc_ref(x_4); +x_430 = l_Lean_IR_toIRType(x_16, x_4, x_5); +if (lean_obj_tag(x_430) == 0) +{ +lean_object* x_431; +x_431 = lean_ctor_get(x_430, 0); +lean_inc(x_431); +lean_dec_ref(x_430); +x_414 = x_431; +x_415 = x_3; +x_416 = x_4; +x_417 = x_5; +x_418 = lean_box(0); +goto block_425; } else { -lean_object* x_427; lean_object* x_428; lean_object* x_429; -lean_dec_ref(x_407); -lean_dec(x_405); -lean_dec_ref(x_404); +lean_object* x_432; lean_object* x_433; lean_object* x_434; +lean_dec(x_413); +lean_dec(x_411); lean_dec_ref(x_403); -lean_free_object(x_14); -lean_dec(x_16); +lean_dec(x_401); +lean_dec_ref(x_400); +lean_dec_ref(x_399); +lean_dec(x_19); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -x_427 = lean_ctor_get(x_414, 0); -lean_inc(x_427); -if (lean_is_exclusive(x_414)) { - lean_ctor_release(x_414, 0); - x_428 = x_414; +x_432 = lean_ctor_get(x_430, 0); +lean_inc(x_432); +if (lean_is_exclusive(x_430)) { + lean_ctor_release(x_430, 0); + x_433 = x_430; } else { - lean_dec_ref(x_414); - x_428 = lean_box(0); + lean_dec_ref(x_430); + x_433 = lean_box(0); } -if (lean_is_scalar(x_428)) { - x_429 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_433)) { + x_434 = lean_alloc_ctor(1, 1, 0); } else { - x_429 = x_428; + x_434 = x_433; } -lean_ctor_set(x_429, 0, x_427); -return x_429; +lean_ctor_set(x_434, 0, x_432); +return x_434; } } -else +block_425: { -lean_dec_ref(x_407); -lean_dec(x_405); -lean_dec_ref(x_404); +lean_object* x_419; +lean_inc(x_413); +lean_inc_ref(x_399); +x_419 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_lowerNonObjectFields___redArg(x_2, x_399, x_400, x_403, x_413, x_415, x_416, x_417); lean_dec_ref(x_403); -lean_free_object(x_14); -lean_dec(x_16); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -goto block_402; -} -} -block_402: -{ -lean_object* x_400; lean_object* x_401; -x_400 = lean_box(12); -if (lean_is_scalar(x_399)) { - x_401 = lean_alloc_ctor(0, 1, 0); -} else { - x_401 = x_399; -} -lean_ctor_set(x_401, 0, x_400); -return x_401; -} -} -else +lean_dec_ref(x_400); +if (lean_obj_tag(x_419) == 0) { -lean_object* x_430; lean_object* x_431; lean_object* x_432; -lean_dec(x_375); -lean_dec(x_156); -lean_free_object(x_14); -lean_dec(x_16); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -x_430 = lean_ctor_get(x_397, 0); -lean_inc(x_430); -if (lean_is_exclusive(x_397)) { - lean_ctor_release(x_397, 0); - x_431 = x_397; +lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; +x_420 = lean_ctor_get(x_419, 0); +lean_inc(x_420); +if (lean_is_exclusive(x_419)) { + lean_ctor_release(x_419, 0); + x_421 = x_419; } else { - lean_dec_ref(x_397); - x_431 = lean_box(0); + lean_dec_ref(x_419); + x_421 = lean_box(0); } -if (lean_is_scalar(x_431)) { - x_432 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_401)) { + x_422 = lean_alloc_ctor(0, 2, 0); } else { - x_432 = x_431; + x_422 = x_401; } -lean_ctor_set(x_432, 0, x_430); -return x_432; -} -} -else -{ -lean_object* x_433; -lean_dec(x_375); -lean_dec(x_156); -lean_dec(x_151); -x_433 = l_Lean_IR_ToIR_bindVar___redArg(x_16, x_3); -if (lean_obj_tag(x_433) == 0) -{ -lean_object* x_434; lean_object* x_435; -x_434 = lean_ctor_get(x_433, 0); -lean_inc(x_434); -lean_dec_ref(x_433); -x_435 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); -if (lean_obj_tag(x_435) == 0) -{ -lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; -x_436 = lean_ctor_get(x_435, 0); -lean_inc(x_436); -if (lean_is_exclusive(x_435)) { - lean_ctor_release(x_435, 0); - x_437 = x_435; +lean_ctor_set(x_422, 0, x_399); +lean_ctor_set(x_422, 1, x_411); +if (lean_is_scalar(x_19)) { + x_423 = lean_alloc_ctor(0, 4, 0); } else { - lean_dec_ref(x_435); - x_437 = lean_box(0); + x_423 = x_19; } -x_438 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_438, 0, x_374); -lean_ctor_set_tag(x_171, 11); -lean_ctor_set(x_171, 0, x_438); -lean_ctor_set(x_14, 3, x_436); -lean_ctor_set(x_14, 2, x_171); -lean_ctor_set(x_14, 1, x_395); -lean_ctor_set(x_14, 0, x_434); -if (lean_is_scalar(x_437)) { - x_439 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_423, 0, x_413); +lean_ctor_set(x_423, 1, x_414); +lean_ctor_set(x_423, 2, x_422); +lean_ctor_set(x_423, 3, x_420); +if (lean_is_scalar(x_421)) { + x_424 = lean_alloc_ctor(0, 1, 0); } else { - x_439 = x_437; + x_424 = x_421; } -lean_ctor_set(x_439, 0, x_14); -return x_439; +lean_ctor_set(x_424, 0, x_423); +return x_424; } else { -lean_dec(x_434); -lean_dec(x_395); -lean_dec(x_374); -lean_free_object(x_171); -lean_free_object(x_14); -return x_435; +lean_dec(x_414); +lean_dec(x_413); +lean_dec(x_411); +lean_dec(x_401); +lean_dec_ref(x_399); +lean_dec(x_19); +return x_419; +} } } else { -lean_object* x_440; lean_object* x_441; lean_object* x_442; -lean_dec(x_395); -lean_dec(x_374); -lean_free_object(x_171); -lean_free_object(x_14); +lean_object* x_435; lean_object* x_436; lean_object* x_437; +lean_dec(x_411); +lean_dec_ref(x_403); +lean_dec(x_401); +lean_dec_ref(x_400); +lean_dec_ref(x_399); +lean_dec(x_369); +lean_dec(x_19); +lean_dec_ref(x_16); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -x_440 = lean_ctor_get(x_433, 0); -lean_inc(x_440); -if (lean_is_exclusive(x_433)) { - lean_ctor_release(x_433, 0); - x_441 = x_433; +x_435 = lean_ctor_get(x_412, 0); +lean_inc(x_435); +if (lean_is_exclusive(x_412)) { + lean_ctor_release(x_412, 0); + x_436 = x_412; } else { - lean_dec_ref(x_433); - x_441 = lean_box(0); + lean_dec_ref(x_412); + x_436 = lean_box(0); } -if (lean_is_scalar(x_441)) { - x_442 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_436)) { + x_437 = lean_alloc_ctor(1, 1, 0); } else { - x_442 = x_441; -} -lean_ctor_set(x_442, 0, x_440); -return x_442; + x_437 = x_436; } +lean_ctor_set(x_437, 0, x_435); +return x_437; } } else { -lean_object* x_443; lean_object* x_444; lean_object* x_445; -lean_dec(x_375); -lean_dec(x_374); -lean_free_object(x_171); -lean_dec(x_156); -lean_dec(x_151); -lean_free_object(x_14); -lean_dec(x_16); +lean_object* x_438; lean_object* x_439; lean_object* x_440; +lean_dec_ref(x_403); +lean_dec(x_401); +lean_dec_ref(x_400); +lean_dec_ref(x_399); +lean_dec(x_369); +lean_dec(x_19); +lean_dec_ref(x_16); +lean_dec(x_15); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -x_443 = lean_ctor_get(x_394, 0); -lean_inc(x_443); -if (lean_is_exclusive(x_394)) { - lean_ctor_release(x_394, 0); - x_444 = x_394; +x_438 = lean_ctor_get(x_410, 0); +lean_inc(x_438); +if (lean_is_exclusive(x_410)) { + lean_ctor_release(x_410, 0); + x_439 = x_410; } else { - lean_dec_ref(x_394); - x_444 = lean_box(0); + lean_dec_ref(x_410); + x_439 = lean_box(0); } -if (lean_is_scalar(x_444)) { - x_445 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_439)) { + x_440 = lean_alloc_ctor(1, 1, 0); } else { - x_445 = x_444; -} -lean_ctor_set(x_445, 0, x_443); -return x_445; + x_440 = x_439; } +lean_ctor_set(x_440, 0, x_438); +return x_440; } } else { -lean_object* x_446; lean_object* x_447; lean_object* x_448; -lean_dec(x_375); -lean_dec(x_374); -lean_dec(x_373); -lean_free_object(x_171); -lean_dec(x_156); -lean_dec_ref(x_152); -lean_dec(x_151); -lean_free_object(x_14); -lean_dec(x_16); +lean_dec_ref(x_403); +lean_dec(x_401); +lean_dec_ref(x_400); +lean_dec_ref(x_399); +lean_dec(x_369); +lean_dec(x_19); +lean_dec_ref(x_16); +lean_dec(x_15); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -x_446 = lean_ctor_get(x_376, 0); -lean_inc(x_446); -if (lean_is_exclusive(x_376)) { - lean_ctor_release(x_376, 0); - x_447 = x_376; -} else { - lean_dec_ref(x_376); - x_447 = lean_box(0); +goto block_398; } -if (lean_is_scalar(x_447)) { - x_448 = lean_alloc_ctor(1, 1, 0); -} else { - x_448 = x_447; } -lean_ctor_set(x_448, 0, x_446); -return x_448; +block_398: +{ +lean_object* x_396; lean_object* x_397; +x_396 = lean_box(12); +if (lean_is_scalar(x_395)) { + x_397 = lean_alloc_ctor(0, 1, 0); +} else { + x_397 = x_395; } +lean_ctor_set(x_397, 0, x_396); +return x_397; } } -case 7: +else { -uint8_t x_449; -lean_free_object(x_171); -lean_dec(x_156); -lean_dec_ref(x_152); -lean_free_object(x_14); -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_449 = !lean_is_exclusive(x_175); -if (x_449 == 0) -{ -lean_object* x_450; lean_object* x_451; uint8_t x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; -x_450 = lean_ctor_get(x_175, 0); -lean_dec(x_450); -x_451 = l_Lean_IR_ToIR_lowerLet___closed__14; -x_452 = 1; -x_453 = l_Lean_Name_toString(x_151, x_452); -lean_ctor_set_tag(x_175, 3); -lean_ctor_set(x_175, 0, x_453); -x_454 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_454, 0, x_451); -lean_ctor_set(x_454, 1, x_175); -x_455 = l_Lean_IR_ToIR_lowerLet___closed__16; -x_456 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_456, 0, x_454); -lean_ctor_set(x_456, 1, x_455); -x_457 = l_Lean_MessageData_ofFormat(x_456); -x_458 = l_Lean_throwError___at___00Lean_IR_ToIR_lowerLet_spec__10___redArg(x_457, x_4, x_5); -lean_dec(x_5); -lean_dec_ref(x_4); -return x_458; -} -else -{ -lean_object* x_459; uint8_t x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; -lean_dec(x_175); -x_459 = l_Lean_IR_ToIR_lowerLet___closed__14; -x_460 = 1; -x_461 = l_Lean_Name_toString(x_151, x_460); -x_462 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_462, 0, x_461); -x_463 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_463, 0, x_459); -lean_ctor_set(x_463, 1, x_462); -x_464 = l_Lean_IR_ToIR_lowerLet___closed__16; -x_465 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_465, 0, x_463); -lean_ctor_set(x_465, 1, x_464); -x_466 = l_Lean_MessageData_ofFormat(x_465); -x_467 = l_Lean_throwError___at___00Lean_IR_ToIR_lowerLet_spec__10___redArg(x_466, x_4, x_5); +lean_object* x_441; lean_object* x_442; lean_object* x_443; +lean_dec(x_371); +lean_dec(x_369); +lean_dec(x_159); +lean_dec(x_19); +lean_dec_ref(x_16); +lean_dec(x_15); lean_dec(x_5); lean_dec_ref(x_4); -return x_467; -} +lean_dec(x_3); +lean_dec_ref(x_2); +x_441 = lean_ctor_get(x_393, 0); +lean_inc(x_441); +if (lean_is_exclusive(x_393)) { + lean_ctor_release(x_393, 0); + x_442 = x_393; +} else { + lean_dec_ref(x_393); + x_442 = lean_box(0); } -default: -{ -lean_object* x_468; -lean_free_object(x_171); -lean_dec(x_175); -lean_dec_ref(x_152); -lean_free_object(x_14); -x_468 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_mkFap(x_1, x_2, x_151, x_156, x_3, x_4, x_5); -return x_468; +if (lean_is_scalar(x_442)) { + x_443 = lean_alloc_ctor(1, 1, 0); +} else { + x_443 = x_442; } +lean_ctor_set(x_443, 0, x_441); +return x_443; } } else { -lean_object* x_469; -x_469 = lean_ctor_get(x_171, 0); -lean_inc(x_469); -lean_dec(x_171); -switch (lean_obj_tag(x_469)) { -case 0: +lean_object* x_444; +lean_dec(x_371); +lean_dec(x_369); +lean_dec(x_159); +lean_dec(x_154); +lean_dec_ref(x_16); +x_444 = l_Lean_IR_ToIR_bindVar___redArg(x_15, x_3); +if (lean_obj_tag(x_444) == 0) { -lean_object* x_470; lean_object* x_471; lean_object* x_472; uint8_t x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; -lean_dec(x_156); -lean_dec_ref(x_152); -lean_free_object(x_14); -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -if (lean_is_exclusive(x_469)) { - lean_ctor_release(x_469, 0); - x_470 = x_469; -} else { - lean_dec_ref(x_469); - x_470 = lean_box(0); -} -x_471 = l_Lean_IR_ToIR_lowerLet___closed__7; -x_472 = l_Lean_IR_ToIR_lowerLet___closed__9; -x_473 = 1; -x_474 = l_Lean_Name_toString(x_151, x_473); -if (lean_is_scalar(x_470)) { - x_475 = lean_alloc_ctor(3, 1, 0); -} else { - x_475 = x_470; - lean_ctor_set_tag(x_475, 3); -} -lean_ctor_set(x_475, 0, x_474); -x_476 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_476, 0, x_472); -lean_ctor_set(x_476, 1, x_475); -x_477 = l_Lean_IR_ToIR_lowerLet___closed__11; -x_478 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_478, 0, x_476); -lean_ctor_set(x_478, 1, x_477); -x_479 = l_Lean_MessageData_ofFormat(x_478); -x_480 = l_Lean_throwNamedError___at___00Lean_IR_ToIR_lowerLet_spec__8___redArg(x_471, x_479, x_4, x_5); -lean_dec(x_5); -lean_dec_ref(x_4); -return x_480; -} -case 2: -{ -lean_object* x_481; lean_object* x_482; lean_object* x_483; uint8_t x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; -lean_dec(x_156); -lean_dec_ref(x_152); -lean_free_object(x_14); -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -if (lean_is_exclusive(x_469)) { - lean_ctor_release(x_469, 0); - x_481 = x_469; -} else { - lean_dec_ref(x_469); - x_481 = lean_box(0); -} -x_482 = l_Lean_IR_ToIR_lowerLet___closed__7; -x_483 = l_Lean_IR_ToIR_lowerLet___closed__9; -x_484 = 1; -x_485 = l_Lean_Name_toString(x_151, x_484); -if (lean_is_scalar(x_481)) { - x_486 = lean_alloc_ctor(3, 1, 0); -} else { - x_486 = x_481; - lean_ctor_set_tag(x_486, 3); -} -lean_ctor_set(x_486, 0, x_485); -x_487 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_487, 0, x_483); -lean_ctor_set(x_487, 1, x_486); -x_488 = l_Lean_IR_ToIR_lowerLet___closed__11; -x_489 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_489, 0, x_487); -lean_ctor_set(x_489, 1, x_488); -x_490 = l_Lean_MessageData_ofFormat(x_489); -x_491 = l_Lean_throwNamedError___at___00Lean_IR_ToIR_lowerLet_spec__8___redArg(x_482, x_490, x_4, x_5); -lean_dec(x_5); -lean_dec_ref(x_4); -return x_491; -} -case 4: -{ -lean_object* x_492; lean_object* x_493; lean_object* x_494; uint8_t x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; -lean_dec(x_156); -lean_dec_ref(x_152); -lean_free_object(x_14); -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -if (lean_is_exclusive(x_469)) { - lean_ctor_release(x_469, 0); - x_492 = x_469; -} else { - lean_dec_ref(x_469); - x_492 = lean_box(0); -} -x_493 = l_Lean_IR_ToIR_lowerLet___closed__7; -x_494 = l_Lean_IR_ToIR_lowerLet___closed__9; -x_495 = 1; -x_496 = l_Lean_Name_toString(x_151, x_495); -if (lean_is_scalar(x_492)) { - x_497 = lean_alloc_ctor(3, 1, 0); -} else { - x_497 = x_492; - lean_ctor_set_tag(x_497, 3); -} -lean_ctor_set(x_497, 0, x_496); -x_498 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_498, 0, x_494); -lean_ctor_set(x_498, 1, x_497); -x_499 = l_Lean_IR_ToIR_lowerLet___closed__11; -x_500 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_500, 0, x_498); -lean_ctor_set(x_500, 1, x_499); -x_501 = l_Lean_MessageData_ofFormat(x_500); -x_502 = l_Lean_throwNamedError___at___00Lean_IR_ToIR_lowerLet_spec__8___redArg(x_493, x_501, x_4, x_5); -lean_dec(x_5); -lean_dec_ref(x_4); -return x_502; -} -case 5: -{ -lean_object* x_503; lean_object* x_504; lean_object* x_505; uint8_t x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; -lean_dec(x_156); -lean_dec_ref(x_152); -lean_free_object(x_14); -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -if (lean_is_exclusive(x_469)) { - lean_ctor_release(x_469, 0); - x_503 = x_469; -} else { - lean_dec_ref(x_469); - x_503 = lean_box(0); -} -x_504 = l_Lean_IR_ToIR_lowerLet___closed__7; -x_505 = l_Lean_IR_ToIR_lowerLet___closed__9; -x_506 = 1; -x_507 = l_Lean_Name_toString(x_151, x_506); -if (lean_is_scalar(x_503)) { - x_508 = lean_alloc_ctor(3, 1, 0); -} else { - x_508 = x_503; - lean_ctor_set_tag(x_508, 3); -} -lean_ctor_set(x_508, 0, x_507); -x_509 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_509, 0, x_505); -lean_ctor_set(x_509, 1, x_508); -x_510 = l_Lean_IR_ToIR_lowerLet___closed__11; -x_511 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_511, 0, x_509); -lean_ctor_set(x_511, 1, x_510); -x_512 = l_Lean_MessageData_ofFormat(x_511); -x_513 = l_Lean_throwNamedError___at___00Lean_IR_ToIR_lowerLet_spec__8___redArg(x_504, x_512, x_4, x_5); -lean_dec(x_5); -lean_dec_ref(x_4); -return x_513; -} -case 6: -{ -lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; -lean_inc(x_16); -lean_dec_ref(x_1); -x_514 = lean_ctor_get(x_469, 0); -lean_inc_ref(x_514); -if (lean_is_exclusive(x_469)) { - lean_ctor_release(x_469, 0); - x_515 = x_469; -} else { - lean_dec_ref(x_469); - x_515 = lean_box(0); -} -x_516 = lean_ctor_get(x_514, 1); -lean_inc(x_516); -x_517 = lean_ctor_get(x_514, 2); -lean_inc(x_517); -x_518 = lean_ctor_get(x_514, 3); -lean_inc(x_518); -lean_dec_ref(x_514); -lean_inc(x_5); -lean_inc_ref(x_4); -lean_inc(x_516); -x_519 = l_Lean_IR_hasTrivialStructure_x3f(x_516, x_4, x_5); -if (lean_obj_tag(x_519) == 0) -{ -lean_object* x_520; -x_520 = lean_ctor_get(x_519, 0); -lean_inc(x_520); -lean_dec_ref(x_519); -if (lean_obj_tag(x_520) == 1) -{ -lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; -lean_dec(x_518); -lean_dec(x_517); -lean_dec(x_516); -lean_dec(x_515); -lean_dec(x_156); -lean_dec(x_151); -lean_free_object(x_14); -x_521 = lean_ctor_get(x_520, 0); -lean_inc(x_521); -lean_dec_ref(x_520); -x_522 = lean_st_ref_take(x_3); -x_523 = lean_ctor_get(x_521, 1); -lean_inc(x_523); -x_524 = lean_ctor_get(x_521, 2); -lean_inc(x_524); -lean_dec(x_521); -x_525 = lean_ctor_get(x_522, 0); -lean_inc_ref(x_525); -x_526 = lean_ctor_get(x_522, 1); -lean_inc_ref(x_526); -x_527 = lean_ctor_get(x_522, 2); -lean_inc(x_527); -x_528 = lean_ctor_get(x_522, 3); -lean_inc_ref(x_528); -if (lean_is_exclusive(x_522)) { - lean_ctor_release(x_522, 0); - lean_ctor_release(x_522, 1); - lean_ctor_release(x_522, 2); - lean_ctor_release(x_522, 3); - x_529 = x_522; -} else { - lean_dec_ref(x_522); - x_529 = lean_box(0); -} -x_530 = lean_box(0); -x_531 = lean_nat_add(x_523, x_524); -lean_dec(x_524); -lean_dec(x_523); -x_532 = lean_array_get(x_530, x_152, x_531); -lean_dec(x_531); -lean_dec_ref(x_152); -x_533 = l_Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_ToIR_lowerParam_spec__0___redArg(x_528, x_16, x_532); -if (lean_is_scalar(x_529)) { - x_534 = lean_alloc_ctor(0, 4, 0); -} else { - x_534 = x_529; -} -lean_ctor_set(x_534, 0, x_525); -lean_ctor_set(x_534, 1, x_526); -lean_ctor_set(x_534, 2, x_527); -lean_ctor_set(x_534, 3, x_533); -x_535 = lean_st_ref_set(x_3, x_534); -x_536 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); -return x_536; -} -else -{ -lean_object* x_537; -lean_dec(x_520); -lean_dec_ref(x_152); -lean_inc(x_5); -lean_inc_ref(x_4); -x_537 = l_Lean_IR_nameToIRType(x_516, x_4, x_5); -if (lean_obj_tag(x_537) == 0) -{ -lean_object* x_538; uint8_t x_539; -x_538 = lean_ctor_get(x_537, 0); -lean_inc(x_538); -lean_dec_ref(x_537); -x_539 = l_Lean_IR_IRType_isScalar(x_538); -if (x_539 == 0) -{ -lean_object* x_540; -lean_dec(x_538); -lean_dec(x_517); -lean_dec(x_515); -lean_inc(x_5); -lean_inc_ref(x_4); -x_540 = l_Lean_IR_getCtorLayout(x_151, x_4, x_5); -if (lean_obj_tag(x_540) == 0) -{ -lean_object* x_541; lean_object* x_542; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; uint8_t x_553; -x_541 = lean_ctor_get(x_540, 0); -lean_inc(x_541); -if (lean_is_exclusive(x_540)) { - lean_ctor_release(x_540, 0); - x_542 = x_540; -} else { - lean_dec_ref(x_540); - x_542 = lean_box(0); -} -x_546 = lean_ctor_get(x_541, 0); -lean_inc_ref(x_546); -x_547 = lean_ctor_get(x_541, 1); -lean_inc_ref(x_547); -if (lean_is_exclusive(x_541)) { - lean_ctor_release(x_541, 0); - lean_ctor_release(x_541, 1); - x_548 = x_541; -} else { - lean_dec_ref(x_541); - x_548 = lean_box(0); -} -x_549 = lean_array_get_size(x_156); -x_550 = l_Array_extract___redArg(x_156, x_518, x_549); -lean_dec(x_156); -x_551 = lean_array_get_size(x_550); -x_552 = lean_array_get_size(x_547); -x_553 = lean_nat_dec_eq(x_551, x_552); -if (x_553 == 0) -{ -lean_dec_ref(x_550); -lean_dec(x_548); -lean_dec_ref(x_547); -lean_dec_ref(x_546); -lean_free_object(x_14); -lean_dec(x_16); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -goto block_545; -} -else -{ -if (x_539 == 0) -{ -lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; -lean_dec(x_542); -x_554 = l_Lean_IR_ToIR_getFVarValue___redArg___closed__0; -x_555 = lean_unsigned_to_nat(0u); -x_556 = l_Lean_IR_ToIR_lowerLet___closed__12; -x_557 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_ToIR_lowerLet_spec__9___redArg(x_552, x_547, x_554, x_550, x_555, x_556); -if (lean_obj_tag(x_557) == 0) -{ -lean_object* x_558; lean_object* x_559; -x_558 = lean_ctor_get(x_557, 0); -lean_inc(x_558); -lean_dec_ref(x_557); -x_559 = l_Lean_IR_ToIR_bindVar___redArg(x_16, x_3); -if (lean_obj_tag(x_559) == 0) -{ -lean_object* x_560; lean_object* x_561; -x_560 = lean_ctor_get(x_559, 0); -lean_inc(x_560); -lean_dec_ref(x_559); -lean_inc(x_560); -x_561 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_lowerNonObjectFields___redArg(x_2, x_546, x_547, x_550, x_560, x_3, x_4, x_5); -lean_dec_ref(x_550); -lean_dec_ref(x_547); -if (lean_obj_tag(x_561) == 0) -{ -lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; lean_object* x_566; -x_562 = lean_ctor_get(x_561, 0); -lean_inc(x_562); -if (lean_is_exclusive(x_561)) { - lean_ctor_release(x_561, 0); - x_563 = x_561; -} else { - lean_dec_ref(x_561); - x_563 = lean_box(0); -} -x_564 = l_Lean_IR_CtorInfo_type(x_546); -if (lean_is_scalar(x_548)) { - x_565 = lean_alloc_ctor(0, 2, 0); -} else { - x_565 = x_548; -} -lean_ctor_set(x_565, 0, x_546); -lean_ctor_set(x_565, 1, x_558); -lean_ctor_set(x_14, 3, x_562); -lean_ctor_set(x_14, 2, x_565); -lean_ctor_set(x_14, 1, x_564); -lean_ctor_set(x_14, 0, x_560); -if (lean_is_scalar(x_563)) { - x_566 = lean_alloc_ctor(0, 1, 0); -} else { - x_566 = x_563; -} -lean_ctor_set(x_566, 0, x_14); -return x_566; -} -else -{ -lean_dec(x_560); -lean_dec(x_558); -lean_dec(x_548); -lean_dec_ref(x_546); -lean_free_object(x_14); -return x_561; -} -} -else -{ -lean_object* x_567; lean_object* x_568; lean_object* x_569; -lean_dec(x_558); -lean_dec_ref(x_550); -lean_dec(x_548); -lean_dec_ref(x_547); -lean_dec_ref(x_546); -lean_free_object(x_14); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -x_567 = lean_ctor_get(x_559, 0); -lean_inc(x_567); -if (lean_is_exclusive(x_559)) { - lean_ctor_release(x_559, 0); - x_568 = x_559; -} else { - lean_dec_ref(x_559); - x_568 = lean_box(0); -} -if (lean_is_scalar(x_568)) { - x_569 = lean_alloc_ctor(1, 1, 0); -} else { - x_569 = x_568; -} -lean_ctor_set(x_569, 0, x_567); -return x_569; -} -} -else -{ -lean_object* x_570; lean_object* x_571; lean_object* x_572; -lean_dec_ref(x_550); -lean_dec(x_548); -lean_dec_ref(x_547); -lean_dec_ref(x_546); -lean_free_object(x_14); -lean_dec(x_16); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -x_570 = lean_ctor_get(x_557, 0); -lean_inc(x_570); -if (lean_is_exclusive(x_557)) { - lean_ctor_release(x_557, 0); - x_571 = x_557; -} else { - lean_dec_ref(x_557); - x_571 = lean_box(0); -} -if (lean_is_scalar(x_571)) { - x_572 = lean_alloc_ctor(1, 1, 0); -} else { - x_572 = x_571; -} -lean_ctor_set(x_572, 0, x_570); -return x_572; -} -} -else -{ -lean_dec_ref(x_550); -lean_dec(x_548); -lean_dec_ref(x_547); -lean_dec_ref(x_546); -lean_free_object(x_14); -lean_dec(x_16); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -goto block_545; -} -} -block_545: -{ -lean_object* x_543; lean_object* x_544; -x_543 = lean_box(12); -if (lean_is_scalar(x_542)) { - x_544 = lean_alloc_ctor(0, 1, 0); -} else { - x_544 = x_542; -} -lean_ctor_set(x_544, 0, x_543); -return x_544; -} -} -else -{ -lean_object* x_573; lean_object* x_574; lean_object* x_575; -lean_dec(x_518); -lean_dec(x_156); -lean_free_object(x_14); -lean_dec(x_16); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -x_573 = lean_ctor_get(x_540, 0); -lean_inc(x_573); -if (lean_is_exclusive(x_540)) { - lean_ctor_release(x_540, 0); - x_574 = x_540; -} else { - lean_dec_ref(x_540); - x_574 = lean_box(0); -} -if (lean_is_scalar(x_574)) { - x_575 = lean_alloc_ctor(1, 1, 0); -} else { - x_575 = x_574; -} -lean_ctor_set(x_575, 0, x_573); -return x_575; -} -} -else -{ -lean_object* x_576; -lean_dec(x_518); -lean_dec(x_156); -lean_dec(x_151); -x_576 = l_Lean_IR_ToIR_bindVar___redArg(x_16, x_3); -if (lean_obj_tag(x_576) == 0) -{ -lean_object* x_577; lean_object* x_578; -x_577 = lean_ctor_get(x_576, 0); -lean_inc(x_577); -lean_dec_ref(x_576); -x_578 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); -if (lean_obj_tag(x_578) == 0) -{ -lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; -x_579 = lean_ctor_get(x_578, 0); -lean_inc(x_579); -if (lean_is_exclusive(x_578)) { - lean_ctor_release(x_578, 0); - x_580 = x_578; -} else { - lean_dec_ref(x_578); - x_580 = lean_box(0); -} -if (lean_is_scalar(x_515)) { - x_581 = lean_alloc_ctor(0, 1, 0); -} else { - x_581 = x_515; - lean_ctor_set_tag(x_581, 0); -} -lean_ctor_set(x_581, 0, x_517); -x_582 = lean_alloc_ctor(11, 1, 0); -lean_ctor_set(x_582, 0, x_581); -lean_ctor_set(x_14, 3, x_579); -lean_ctor_set(x_14, 2, x_582); -lean_ctor_set(x_14, 1, x_538); -lean_ctor_set(x_14, 0, x_577); -if (lean_is_scalar(x_580)) { - x_583 = lean_alloc_ctor(0, 1, 0); -} else { - x_583 = x_580; -} -lean_ctor_set(x_583, 0, x_14); -return x_583; -} -else -{ -lean_dec(x_577); -lean_dec(x_538); -lean_dec(x_517); -lean_dec(x_515); -lean_free_object(x_14); -return x_578; -} -} -else -{ -lean_object* x_584; lean_object* x_585; lean_object* x_586; -lean_dec(x_538); -lean_dec(x_517); -lean_dec(x_515); -lean_free_object(x_14); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -x_584 = lean_ctor_get(x_576, 0); -lean_inc(x_584); -if (lean_is_exclusive(x_576)) { - lean_ctor_release(x_576, 0); - x_585 = x_576; -} else { - lean_dec_ref(x_576); - x_585 = lean_box(0); -} -if (lean_is_scalar(x_585)) { - x_586 = lean_alloc_ctor(1, 1, 0); -} else { - x_586 = x_585; -} -lean_ctor_set(x_586, 0, x_584); -return x_586; -} -} -} -else -{ -lean_object* x_587; lean_object* x_588; lean_object* x_589; -lean_dec(x_518); -lean_dec(x_517); -lean_dec(x_515); -lean_dec(x_156); -lean_dec(x_151); -lean_free_object(x_14); -lean_dec(x_16); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -x_587 = lean_ctor_get(x_537, 0); -lean_inc(x_587); -if (lean_is_exclusive(x_537)) { - lean_ctor_release(x_537, 0); - x_588 = x_537; -} else { - lean_dec_ref(x_537); - x_588 = lean_box(0); -} -if (lean_is_scalar(x_588)) { - x_589 = lean_alloc_ctor(1, 1, 0); -} else { - x_589 = x_588; -} -lean_ctor_set(x_589, 0, x_587); -return x_589; -} -} -} -else -{ -lean_object* x_590; lean_object* x_591; lean_object* x_592; -lean_dec(x_518); -lean_dec(x_517); -lean_dec(x_516); -lean_dec(x_515); -lean_dec(x_156); -lean_dec_ref(x_152); -lean_dec(x_151); -lean_free_object(x_14); -lean_dec(x_16); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -x_590 = lean_ctor_get(x_519, 0); -lean_inc(x_590); -if (lean_is_exclusive(x_519)) { - lean_ctor_release(x_519, 0); - x_591 = x_519; -} else { - lean_dec_ref(x_519); - x_591 = lean_box(0); -} -if (lean_is_scalar(x_591)) { - x_592 = lean_alloc_ctor(1, 1, 0); -} else { - x_592 = x_591; -} -lean_ctor_set(x_592, 0, x_590); -return x_592; -} -} -case 7: -{ -lean_object* x_593; lean_object* x_594; uint8_t x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; -lean_dec(x_156); -lean_dec_ref(x_152); -lean_free_object(x_14); -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -if (lean_is_exclusive(x_469)) { - lean_ctor_release(x_469, 0); - x_593 = x_469; -} else { - lean_dec_ref(x_469); - x_593 = lean_box(0); -} -x_594 = l_Lean_IR_ToIR_lowerLet___closed__14; -x_595 = 1; -x_596 = l_Lean_Name_toString(x_151, x_595); -if (lean_is_scalar(x_593)) { - x_597 = lean_alloc_ctor(3, 1, 0); -} else { - x_597 = x_593; - lean_ctor_set_tag(x_597, 3); -} -lean_ctor_set(x_597, 0, x_596); -x_598 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_598, 0, x_594); -lean_ctor_set(x_598, 1, x_597); -x_599 = l_Lean_IR_ToIR_lowerLet___closed__16; -x_600 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_600, 0, x_598); -lean_ctor_set(x_600, 1, x_599); -x_601 = l_Lean_MessageData_ofFormat(x_600); -x_602 = l_Lean_throwError___at___00Lean_IR_ToIR_lowerLet_spec__10___redArg(x_601, x_4, x_5); -lean_dec(x_5); -lean_dec_ref(x_4); -return x_602; -} -default: -{ -lean_object* x_603; -lean_dec(x_469); -lean_dec_ref(x_152); -lean_free_object(x_14); -x_603 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_mkFap(x_1, x_2, x_151, x_156, x_3, x_4, x_5); -return x_603; -} -} -} -} -} -} -else -{ -uint8_t x_604; -lean_dec(x_156); -lean_dec_ref(x_152); -lean_dec(x_151); -lean_free_object(x_14); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_604 = !lean_is_exclusive(x_163); -if (x_604 == 0) -{ -return x_163; -} -else -{ -lean_object* x_605; lean_object* x_606; -x_605 = lean_ctor_get(x_163, 0); -lean_inc(x_605); -lean_dec(x_163); -x_606 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_606, 0, x_605); -return x_606; -} -} -} -} -else -{ -uint8_t x_607; -lean_dec(x_156); -lean_dec_ref(x_152); -lean_dec(x_151); -lean_free_object(x_14); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_607 = !lean_is_exclusive(x_157); -if (x_607 == 0) -{ -return x_157; -} -else -{ -lean_object* x_608; lean_object* x_609; -x_608 = lean_ctor_get(x_157, 0); -lean_inc(x_608); -lean_dec(x_157); -x_609 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_609, 0, x_608); -return x_609; -} -} -} -else -{ -uint8_t x_610; -lean_dec_ref(x_152); -lean_dec(x_151); -lean_free_object(x_14); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_610 = !lean_is_exclusive(x_155); -if (x_610 == 0) -{ -return x_155; -} -else -{ -lean_object* x_611; lean_object* x_612; -x_611 = lean_ctor_get(x_155, 0); -lean_inc(x_611); -lean_dec(x_155); -x_612 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_612, 0, x_611); -return x_612; -} -} -} -default: -{ -lean_object* x_613; lean_object* x_614; lean_object* x_615; -lean_free_object(x_14); -x_613 = lean_ctor_get(x_23, 0); -lean_inc(x_613); -x_614 = lean_ctor_get(x_23, 1); -lean_inc_ref(x_614); -lean_dec_ref(x_23); -x_615 = l_Lean_IR_ToIR_getFVarValue___redArg(x_613, x_3); -if (lean_obj_tag(x_615) == 0) -{ -lean_object* x_616; -x_616 = lean_ctor_get(x_615, 0); -lean_inc(x_616); -lean_dec_ref(x_615); -if (lean_obj_tag(x_616) == 0) -{ -lean_object* x_617; size_t x_618; size_t x_619; lean_object* x_620; -x_617 = lean_ctor_get(x_616, 0); -lean_inc(x_617); -lean_dec_ref(x_616); -x_618 = lean_array_size(x_614); -x_619 = 0; -x_620 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ToIR_lowerCode_spec__3___redArg(x_618, x_619, x_614, x_3); -if (lean_obj_tag(x_620) == 0) -{ -lean_object* x_621; lean_object* x_622; -x_621 = lean_ctor_get(x_620, 0); -lean_inc(x_621); -lean_dec_ref(x_620); -x_622 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_mkAp(x_1, x_2, x_617, x_621, x_3, x_4, x_5); -return x_622; -} -else -{ -uint8_t x_623; -lean_dec(x_617); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_623 = !lean_is_exclusive(x_620); -if (x_623 == 0) +lean_object* x_445; lean_object* x_446; +x_445 = lean_ctor_get(x_444, 0); +lean_inc(x_445); +lean_dec_ref(x_444); +x_446 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); +if (lean_obj_tag(x_446) == 0) { -return x_620; -} -else -{ -lean_object* x_624; lean_object* x_625; -x_624 = lean_ctor_get(x_620, 0); -lean_inc(x_624); -lean_dec(x_620); -x_625 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_625, 0, x_624); -return x_625; -} -} -} -else -{ -lean_object* x_626; -lean_dec_ref(x_614); -x_626 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_mkErased___redArg(x_1, x_2, x_3, x_4, x_5); -return x_626; -} -} -else -{ -uint8_t x_627; -lean_dec_ref(x_614); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_627 = !lean_is_exclusive(x_615); -if (x_627 == 0) -{ -return x_615; -} -else -{ -lean_object* x_628; lean_object* x_629; -x_628 = lean_ctor_get(x_615, 0); -lean_inc(x_628); -lean_dec(x_615); -x_629 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_629, 0, x_628); -return x_629; -} -} -} -} -} -else -{ -lean_object* x_630; lean_object* x_631; lean_object* x_632; uint8_t x_633; lean_object* x_634; -x_630 = lean_ctor_get(x_1, 0); -x_631 = lean_ctor_get(x_1, 3); -x_632 = lean_ctor_get(x_14, 3); -lean_inc(x_632); -lean_dec(x_14); -x_633 = 0; -lean_inc(x_631); -x_634 = l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_normLetValueImp(x_632, x_631, x_633); -lean_dec_ref(x_632); -switch (lean_obj_tag(x_634)) { -case 0: -{ -lean_object* x_635; lean_object* x_636; lean_object* x_637; -lean_inc(x_630); -lean_dec_ref(x_1); -x_635 = lean_ctor_get(x_634, 0); -lean_inc_ref(x_635); -if (lean_is_exclusive(x_634)) { - lean_ctor_release(x_634, 0); - x_636 = x_634; -} else { - lean_dec_ref(x_634); - x_636 = lean_box(0); -} -x_637 = l_Lean_IR_ToIR_bindVar___redArg(x_630, x_3); -if (lean_obj_tag(x_637) == 0) -{ -lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; -x_638 = lean_ctor_get(x_637, 0); -lean_inc(x_638); -lean_dec_ref(x_637); -x_639 = l_Lean_IR_ToIR_lowerLitValue(x_635); -x_640 = lean_ctor_get(x_639, 0); -lean_inc(x_640); -x_641 = lean_ctor_get(x_639, 1); -lean_inc(x_641); -lean_dec_ref(x_639); -x_642 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); -if (lean_obj_tag(x_642) == 0) -{ -lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; -x_643 = lean_ctor_get(x_642, 0); -lean_inc(x_643); -if (lean_is_exclusive(x_642)) { - lean_ctor_release(x_642, 0); - x_644 = x_642; +lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; +x_447 = lean_ctor_get(x_446, 0); +lean_inc(x_447); +if (lean_is_exclusive(x_446)) { + lean_ctor_release(x_446, 0); + x_448 = x_446; } else { - lean_dec_ref(x_642); - x_644 = lean_box(0); + lean_dec_ref(x_446); + x_448 = lean_box(0); } -if (lean_is_scalar(x_636)) { - x_645 = lean_alloc_ctor(11, 1, 0); -} else { - x_645 = x_636; - lean_ctor_set_tag(x_645, 11); -} -lean_ctor_set(x_645, 0, x_640); -x_646 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_646, 0, x_638); -lean_ctor_set(x_646, 1, x_641); -lean_ctor_set(x_646, 2, x_645); -lean_ctor_set(x_646, 3, x_643); -if (lean_is_scalar(x_644)) { - x_647 = lean_alloc_ctor(0, 1, 0); -} else { - x_647 = x_644; -} -lean_ctor_set(x_647, 0, x_646); -return x_647; -} -else -{ -lean_dec(x_641); -lean_dec(x_640); -lean_dec(x_638); -lean_dec(x_636); -return x_642; -} -} -else -{ -lean_object* x_648; lean_object* x_649; lean_object* x_650; -lean_dec(x_636); -lean_dec_ref(x_635); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -x_648 = lean_ctor_get(x_637, 0); -lean_inc(x_648); -if (lean_is_exclusive(x_637)) { - lean_ctor_release(x_637, 0); - x_649 = x_637; -} else { - lean_dec_ref(x_637); - x_649 = lean_box(0); -} -if (lean_is_scalar(x_649)) { - x_650 = lean_alloc_ctor(1, 1, 0); -} else { - x_650 = x_649; -} -lean_ctor_set(x_650, 0, x_648); -return x_650; -} -} -case 1: -{ -lean_object* x_651; -x_651 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_mkErased___redArg(x_1, x_2, x_3, x_4, x_5); -return x_651; -} -case 2: -{ -lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; -lean_inc(x_630); -lean_dec_ref(x_1); -x_652 = lean_ctor_get(x_634, 0); -lean_inc(x_652); -x_653 = lean_ctor_get(x_634, 1); -lean_inc(x_653); -x_654 = lean_ctor_get(x_634, 2); -lean_inc(x_654); -lean_dec_ref(x_634); -lean_inc(x_5); -lean_inc_ref(x_4); -lean_inc(x_652); -x_655 = l_Lean_IR_hasTrivialStructure_x3f(x_652, x_4, x_5); -if (lean_obj_tag(x_655) == 0) -{ -lean_object* x_656; -x_656 = lean_ctor_get(x_655, 0); -lean_inc(x_656); -lean_dec_ref(x_655); -if (lean_obj_tag(x_656) == 1) -{ -lean_object* x_657; lean_object* x_658; lean_object* x_659; uint8_t x_660; -lean_dec(x_652); -x_657 = lean_ctor_get(x_656, 0); -lean_inc(x_657); -if (lean_is_exclusive(x_656)) { - lean_ctor_release(x_656, 0); - x_658 = x_656; -} else { - lean_dec_ref(x_656); - x_658 = lean_box(0); -} -x_659 = lean_ctor_get(x_657, 2); -lean_inc(x_659); -lean_dec(x_657); -x_660 = lean_nat_dec_eq(x_659, x_653); -lean_dec(x_653); -lean_dec(x_659); -if (x_660 == 0) -{ -lean_object* x_661; -lean_dec(x_658); -lean_dec(x_654); -x_661 = l_Lean_IR_ToIR_bindErased___redArg(x_630, x_3); -if (lean_obj_tag(x_661) == 0) -{ -lean_object* x_662; -lean_dec_ref(x_661); -x_662 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); -return x_662; -} -else -{ -lean_object* x_663; lean_object* x_664; lean_object* x_665; -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -x_663 = lean_ctor_get(x_661, 0); -lean_inc(x_663); -if (lean_is_exclusive(x_661)) { - lean_ctor_release(x_661, 0); - x_664 = x_661; -} else { - lean_dec_ref(x_661); - x_664 = lean_box(0); -} -if (lean_is_scalar(x_664)) { - x_665 = lean_alloc_ctor(1, 1, 0); -} else { - x_665 = x_664; -} -lean_ctor_set(x_665, 0, x_663); -return x_665; -} -} -else -{ -lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; -x_666 = lean_st_ref_take(x_3); -x_667 = lean_ctor_get(x_666, 0); -lean_inc_ref(x_667); -x_668 = lean_ctor_get(x_666, 1); -lean_inc_ref(x_668); -x_669 = lean_ctor_get(x_666, 2); -lean_inc(x_669); -x_670 = lean_ctor_get(x_666, 3); -lean_inc_ref(x_670); -if (lean_is_exclusive(x_666)) { - lean_ctor_release(x_666, 0); - lean_ctor_release(x_666, 1); - lean_ctor_release(x_666, 2); - lean_ctor_release(x_666, 3); - x_671 = x_666; -} else { - lean_dec_ref(x_666); - x_671 = lean_box(0); -} -if (lean_is_scalar(x_658)) { - x_672 = lean_alloc_ctor(1, 1, 0); -} else { - x_672 = x_658; -} -lean_ctor_set(x_672, 0, x_654); -x_673 = l_Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_ToIR_lowerParam_spec__0___redArg(x_670, x_630, x_672); -if (lean_is_scalar(x_671)) { - x_674 = lean_alloc_ctor(0, 4, 0); -} else { - x_674 = x_671; -} -lean_ctor_set(x_674, 0, x_667); -lean_ctor_set(x_674, 1, x_668); -lean_ctor_set(x_674, 2, x_669); -lean_ctor_set(x_674, 3, x_673); -x_675 = lean_st_ref_set(x_3, x_674); -x_676 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); -return x_676; -} -} -else -{ -lean_object* x_677; -lean_dec(x_656); -x_677 = l_Lean_IR_ToIR_getFVarValue___redArg(x_654, x_3); -if (lean_obj_tag(x_677) == 0) -{ -lean_object* x_678; -x_678 = lean_ctor_get(x_677, 0); -lean_inc(x_678); -lean_dec_ref(x_677); -if (lean_obj_tag(x_678) == 0) -{ -lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_object* x_682; -x_679 = lean_ctor_get(x_678, 0); -lean_inc(x_679); -lean_dec_ref(x_678); -x_680 = lean_st_ref_get(x_5); -x_681 = lean_ctor_get(x_680, 0); -lean_inc_ref(x_681); -lean_dec(x_680); -x_682 = l_Lean_Environment_find_x3f(x_681, x_652, x_633); -if (lean_obj_tag(x_682) == 1) -{ -lean_object* x_683; -x_683 = lean_ctor_get(x_682, 0); -lean_inc(x_683); -lean_dec_ref(x_682); -if (lean_obj_tag(x_683) == 5) -{ -lean_object* x_684; lean_object* x_685; -x_684 = lean_ctor_get(x_683, 0); -lean_inc_ref(x_684); -lean_dec_ref(x_683); -x_685 = lean_ctor_get(x_684, 4); -lean_inc(x_685); -lean_dec_ref(x_684); -if (lean_obj_tag(x_685) == 1) -{ -lean_object* x_686; -x_686 = lean_ctor_get(x_685, 1); -if (lean_obj_tag(x_686) == 0) -{ -lean_object* x_687; lean_object* x_688; -x_687 = lean_ctor_get(x_685, 0); -lean_inc(x_687); -lean_dec_ref(x_685); -lean_inc(x_5); -lean_inc_ref(x_4); -x_688 = l_Lean_IR_getCtorLayout(x_687, x_4, x_5); -if (lean_obj_tag(x_688) == 0) -{ -lean_object* x_689; lean_object* x_690; lean_object* x_691; lean_object* x_692; lean_object* x_693; lean_object* x_694; lean_object* x_695; -x_689 = lean_ctor_get(x_688, 0); -lean_inc(x_689); -lean_dec_ref(x_688); -x_690 = lean_ctor_get(x_689, 0); -lean_inc_ref(x_690); -x_691 = lean_ctor_get(x_689, 1); -lean_inc_ref(x_691); -lean_dec(x_689); -x_692 = lean_box(0); -x_693 = lean_array_get(x_692, x_691, x_653); -lean_dec(x_653); -lean_dec_ref(x_691); -x_694 = l_Lean_IR_ToIR_lowerProj(x_679, x_690, x_693); -lean_dec_ref(x_690); -x_695 = lean_ctor_get(x_694, 0); -lean_inc(x_695); -if (lean_obj_tag(x_695) == 0) -{ -lean_object* x_696; lean_object* x_697; lean_object* x_698; -x_696 = lean_ctor_get(x_694, 1); -lean_inc(x_696); -lean_dec_ref(x_694); -x_697 = lean_ctor_get(x_695, 0); -lean_inc_ref(x_697); -lean_dec_ref(x_695); -x_698 = l_Lean_IR_ToIR_bindVar___redArg(x_630, x_3); -if (lean_obj_tag(x_698) == 0) -{ -lean_object* x_699; lean_object* x_700; -x_699 = lean_ctor_get(x_698, 0); -lean_inc(x_699); -lean_dec_ref(x_698); -x_700 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); -if (lean_obj_tag(x_700) == 0) -{ -lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; -x_701 = lean_ctor_get(x_700, 0); -lean_inc(x_701); -if (lean_is_exclusive(x_700)) { - lean_ctor_release(x_700, 0); - x_702 = x_700; -} else { - lean_dec_ref(x_700); - x_702 = lean_box(0); -} -x_703 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_703, 0, x_699); -lean_ctor_set(x_703, 1, x_696); -lean_ctor_set(x_703, 2, x_697); -lean_ctor_set(x_703, 3, x_701); -if (lean_is_scalar(x_702)) { - x_704 = lean_alloc_ctor(0, 1, 0); -} else { - x_704 = x_702; -} -lean_ctor_set(x_704, 0, x_703); -return x_704; -} -else -{ -lean_dec(x_699); -lean_dec_ref(x_697); -lean_dec(x_696); -return x_700; -} -} -else -{ -lean_object* x_705; lean_object* x_706; lean_object* x_707; -lean_dec_ref(x_697); -lean_dec(x_696); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -x_705 = lean_ctor_get(x_698, 0); -lean_inc(x_705); -if (lean_is_exclusive(x_698)) { - lean_ctor_release(x_698, 0); - x_706 = x_698; -} else { - lean_dec_ref(x_698); - x_706 = lean_box(0); -} -if (lean_is_scalar(x_706)) { - x_707 = lean_alloc_ctor(1, 1, 0); -} else { - x_707 = x_706; -} -lean_ctor_set(x_707, 0, x_705); -return x_707; -} -} -else -{ -lean_object* x_708; -lean_dec_ref(x_694); -x_708 = l_Lean_IR_ToIR_bindErased___redArg(x_630, x_3); -if (lean_obj_tag(x_708) == 0) -{ -lean_object* x_709; -lean_dec_ref(x_708); -x_709 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); -return x_709; -} -else -{ -lean_object* x_710; lean_object* x_711; lean_object* x_712; -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -x_710 = lean_ctor_get(x_708, 0); -lean_inc(x_710); -if (lean_is_exclusive(x_708)) { - lean_ctor_release(x_708, 0); - x_711 = x_708; -} else { - lean_dec_ref(x_708); - x_711 = lean_box(0); -} -if (lean_is_scalar(x_711)) { - x_712 = lean_alloc_ctor(1, 1, 0); -} else { - x_712 = x_711; -} -lean_ctor_set(x_712, 0, x_710); -return x_712; -} -} -} -else -{ -lean_object* x_713; lean_object* x_714; lean_object* x_715; -lean_dec(x_679); -lean_dec(x_653); -lean_dec(x_630); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -x_713 = lean_ctor_get(x_688, 0); -lean_inc(x_713); -if (lean_is_exclusive(x_688)) { - lean_ctor_release(x_688, 0); - x_714 = x_688; -} else { - lean_dec_ref(x_688); - x_714 = lean_box(0); -} -if (lean_is_scalar(x_714)) { - x_715 = lean_alloc_ctor(1, 1, 0); -} else { - x_715 = x_714; -} -lean_ctor_set(x_715, 0, x_713); -return x_715; -} -} -else -{ -lean_dec_ref(x_685); -lean_dec(x_679); -lean_dec(x_653); -lean_dec(x_630); -lean_dec_ref(x_2); -x_7 = x_3; -x_8 = x_4; -x_9 = x_5; -x_10 = lean_box(0); -goto block_13; -} -} -else -{ -lean_dec(x_685); -lean_dec(x_679); -lean_dec(x_653); -lean_dec(x_630); -lean_dec_ref(x_2); -x_7 = x_3; -x_8 = x_4; -x_9 = x_5; -x_10 = lean_box(0); -goto block_13; -} -} -else -{ -lean_dec(x_683); -lean_dec(x_679); -lean_dec(x_653); -lean_dec(x_630); -lean_dec_ref(x_2); -x_7 = x_3; -x_8 = x_4; -x_9 = x_5; -x_10 = lean_box(0); -goto block_13; -} -} -else -{ -lean_dec(x_682); -lean_dec(x_679); -lean_dec(x_653); -lean_dec(x_630); -lean_dec_ref(x_2); -x_7 = x_3; -x_8 = x_4; -x_9 = x_5; -x_10 = lean_box(0); -goto block_13; +x_449 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_449, 0, x_370); +lean_ctor_set_tag(x_174, 11); +lean_ctor_set(x_174, 0, x_449); +if (lean_is_scalar(x_19)) { + x_450 = lean_alloc_ctor(0, 4, 0); +} else { + x_450 = x_19; +} +lean_ctor_set(x_450, 0, x_445); +lean_ctor_set(x_450, 1, x_391); +lean_ctor_set(x_450, 2, x_174); +lean_ctor_set(x_450, 3, x_447); +if (lean_is_scalar(x_448)) { + x_451 = lean_alloc_ctor(0, 1, 0); +} else { + x_451 = x_448; } +lean_ctor_set(x_451, 0, x_450); +return x_451; } else { -lean_object* x_716; -lean_dec(x_653); -lean_dec(x_652); -x_716 = l_Lean_IR_ToIR_bindErased___redArg(x_630, x_3); -if (lean_obj_tag(x_716) == 0) -{ -lean_object* x_717; -lean_dec_ref(x_716); -x_717 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); -return x_717; +lean_dec(x_445); +lean_dec(x_391); +lean_dec(x_370); +lean_free_object(x_174); +lean_dec(x_19); +return x_446; +} } else { -lean_object* x_718; lean_object* x_719; lean_object* x_720; +lean_object* x_452; lean_object* x_453; lean_object* x_454; +lean_dec(x_391); +lean_dec(x_370); +lean_free_object(x_174); +lean_dec(x_19); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -x_718 = lean_ctor_get(x_716, 0); -lean_inc(x_718); -if (lean_is_exclusive(x_716)) { - lean_ctor_release(x_716, 0); - x_719 = x_716; +x_452 = lean_ctor_get(x_444, 0); +lean_inc(x_452); +if (lean_is_exclusive(x_444)) { + lean_ctor_release(x_444, 0); + x_453 = x_444; } else { - lean_dec_ref(x_716); - x_719 = lean_box(0); + lean_dec_ref(x_444); + x_453 = lean_box(0); } -if (lean_is_scalar(x_719)) { - x_720 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_453)) { + x_454 = lean_alloc_ctor(1, 1, 0); } else { - x_720 = x_719; + x_454 = x_453; } -lean_ctor_set(x_720, 0, x_718); -return x_720; +lean_ctor_set(x_454, 0, x_452); +return x_454; } } } else { -lean_object* x_721; lean_object* x_722; lean_object* x_723; -lean_dec(x_653); -lean_dec(x_652); -lean_dec(x_630); +lean_object* x_455; lean_object* x_456; lean_object* x_457; +lean_dec(x_371); +lean_dec(x_370); +lean_dec(x_369); +lean_free_object(x_174); +lean_dec(x_159); +lean_dec(x_154); +lean_dec(x_19); +lean_dec_ref(x_16); +lean_dec(x_15); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -x_721 = lean_ctor_get(x_677, 0); -lean_inc(x_721); -if (lean_is_exclusive(x_677)) { - lean_ctor_release(x_677, 0); - x_722 = x_677; +x_455 = lean_ctor_get(x_390, 0); +lean_inc(x_455); +if (lean_is_exclusive(x_390)) { + lean_ctor_release(x_390, 0); + x_456 = x_390; } else { - lean_dec_ref(x_677); - x_722 = lean_box(0); + lean_dec_ref(x_390); + x_456 = lean_box(0); } -if (lean_is_scalar(x_722)) { - x_723 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_456)) { + x_457 = lean_alloc_ctor(1, 1, 0); } else { - x_723 = x_722; + x_457 = x_456; } -lean_ctor_set(x_723, 0, x_721); -return x_723; +lean_ctor_set(x_457, 0, x_455); +return x_457; } } } else { -lean_object* x_724; lean_object* x_725; lean_object* x_726; -lean_dec(x_654); -lean_dec(x_653); -lean_dec(x_652); -lean_dec(x_630); +lean_object* x_458; lean_object* x_459; lean_object* x_460; +lean_dec(x_371); +lean_dec(x_370); +lean_dec(x_369); +lean_free_object(x_174); +lean_dec(x_159); +lean_dec_ref(x_155); +lean_dec(x_154); +lean_dec(x_19); +lean_dec_ref(x_16); +lean_dec(x_15); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -x_724 = lean_ctor_get(x_655, 0); -lean_inc(x_724); -if (lean_is_exclusive(x_655)) { - lean_ctor_release(x_655, 0); - x_725 = x_655; +x_458 = lean_ctor_get(x_372, 0); +lean_inc(x_458); +if (lean_is_exclusive(x_372)) { + lean_ctor_release(x_372, 0); + x_459 = x_372; } else { - lean_dec_ref(x_655); - x_725 = lean_box(0); + lean_dec_ref(x_372); + x_459 = lean_box(0); } -if (lean_is_scalar(x_725)) { - x_726 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_459)) { + x_460 = lean_alloc_ctor(1, 1, 0); } else { - x_726 = x_725; + x_460 = x_459; } -lean_ctor_set(x_726, 0, x_724); -return x_726; +lean_ctor_set(x_460, 0, x_458); +return x_460; } } -case 3: +} +case 7: { -lean_object* x_727; lean_object* x_728; size_t x_729; size_t x_730; lean_object* x_731; -x_727 = lean_ctor_get(x_634, 0); -lean_inc(x_727); -x_728 = lean_ctor_get(x_634, 2); -lean_inc_ref(x_728); -lean_dec_ref(x_634); -x_729 = lean_array_size(x_728); -x_730 = 0; -lean_inc_ref(x_728); -x_731 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ToIR_lowerCode_spec__3___redArg(x_729, x_730, x_728, x_3); -if (lean_obj_tag(x_731) == 0) -{ -lean_object* x_732; lean_object* x_733; -x_732 = lean_ctor_get(x_731, 0); -lean_inc(x_732); -lean_dec_ref(x_731); -lean_inc(x_727); -x_733 = l_Lean_IR_ToIR_findDecl___redArg(x_727, x_5); -if (lean_obj_tag(x_733) == 0) -{ -lean_object* x_734; -x_734 = lean_ctor_get(x_733, 0); -lean_inc(x_734); -lean_dec_ref(x_733); -if (lean_obj_tag(x_734) == 1) -{ -lean_object* x_735; lean_object* x_736; lean_object* x_737; lean_object* x_738; -lean_dec_ref(x_728); -x_735 = lean_ctor_get(x_734, 0); -lean_inc(x_735); -lean_dec_ref(x_734); -x_736 = l_Lean_IR_Decl_params(x_735); -lean_dec(x_735); -x_737 = lean_array_get_size(x_736); -lean_dec_ref(x_736); -x_738 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_mkApplication(x_1, x_2, x_727, x_737, x_732, x_3, x_4, x_5); -return x_738; -} -else -{ -lean_object* x_739; -lean_dec(x_734); -lean_inc(x_727); -x_739 = l_Lean_Compiler_LCNF_getMonoDecl_x3f___redArg(x_727, x_5); -if (lean_obj_tag(x_739) == 0) -{ -lean_object* x_740; -x_740 = lean_ctor_get(x_739, 0); -lean_inc(x_740); -lean_dec_ref(x_739); -if (lean_obj_tag(x_740) == 1) -{ -lean_object* x_741; lean_object* x_742; lean_object* x_743; lean_object* x_744; -lean_dec_ref(x_728); -x_741 = lean_ctor_get(x_740, 0); -lean_inc(x_741); -lean_dec_ref(x_740); -x_742 = lean_ctor_get(x_741, 3); -lean_inc_ref(x_742); -lean_dec(x_741); -x_743 = lean_array_get_size(x_742); -lean_dec_ref(x_742); -x_744 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_mkApplication(x_1, x_2, x_727, x_743, x_732, x_3, x_4, x_5); -return x_744; -} -else -{ -lean_object* x_745; lean_object* x_746; lean_object* x_747; -lean_dec(x_740); -x_745 = lean_st_ref_get(x_5); -x_746 = lean_ctor_get(x_745, 0); -lean_inc_ref(x_746); -lean_dec(x_745); -lean_inc(x_727); -x_747 = l_Lean_Environment_find_x3f(x_746, x_727, x_633); -if (lean_obj_tag(x_747) == 0) -{ -lean_object* x_748; lean_object* x_749; -lean_dec(x_732); -lean_dec_ref(x_728); -lean_dec(x_727); +uint8_t x_461; +lean_free_object(x_174); +lean_dec(x_159); +lean_dec_ref(x_155); +lean_dec(x_19); +lean_dec(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_748 = l_Lean_IR_ToIR_lowerLet___closed__4; -x_749 = l_panic___at___00__private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerAlt_loop_spec__0(x_748, x_3, x_4, x_5); -return x_749; +x_461 = !lean_is_exclusive(x_178); +if (x_461 == 0) +{ +lean_object* x_462; lean_object* x_463; uint8_t x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; +x_462 = lean_ctor_get(x_178, 0); +lean_dec(x_462); +x_463 = l_Lean_IR_ToIR_lowerLet___closed__14; +x_464 = 1; +x_465 = l_Lean_Name_toString(x_154, x_464); +lean_ctor_set_tag(x_178, 3); +lean_ctor_set(x_178, 0, x_465); +x_466 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_466, 0, x_463); +lean_ctor_set(x_466, 1, x_178); +x_467 = l_Lean_IR_ToIR_lowerLet___closed__16; +x_468 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_468, 0, x_466); +lean_ctor_set(x_468, 1, x_467); +x_469 = l_Lean_MessageData_ofFormat(x_468); +x_470 = l_Lean_throwError___at___00Lean_IR_ToIR_lowerLet_spec__10___redArg(x_469, x_4, x_5); +lean_dec(x_5); +lean_dec_ref(x_4); +return x_470; +} +else +{ +lean_object* x_471; uint8_t x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; +lean_dec(x_178); +x_471 = l_Lean_IR_ToIR_lowerLet___closed__14; +x_472 = 1; +x_473 = l_Lean_Name_toString(x_154, x_472); +x_474 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_474, 0, x_473); +x_475 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_475, 0, x_471); +lean_ctor_set(x_475, 1, x_474); +x_476 = l_Lean_IR_ToIR_lowerLet___closed__16; +x_477 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_477, 0, x_475); +lean_ctor_set(x_477, 1, x_476); +x_478 = l_Lean_MessageData_ofFormat(x_477); +x_479 = l_Lean_throwError___at___00Lean_IR_ToIR_lowerLet_spec__10___redArg(x_478, x_4, x_5); +lean_dec(x_5); +lean_dec_ref(x_4); +return x_479; } -else +} +default: { -lean_object* x_750; lean_object* x_751; -x_750 = lean_ctor_get(x_747, 0); -lean_inc(x_750); -if (lean_is_exclusive(x_747)) { - lean_ctor_release(x_747, 0); - x_751 = x_747; -} else { - lean_dec_ref(x_747); - x_751 = lean_box(0); +lean_object* x_480; +lean_free_object(x_174); +lean_dec(x_178); +lean_dec_ref(x_155); +lean_dec(x_19); +x_480 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_mkFap(x_1, x_2, x_154, x_159, x_3, x_4, x_5); +return x_480; +} +} } -switch (lean_obj_tag(x_750)) { +else +{ +lean_object* x_481; +x_481 = lean_ctor_get(x_174, 0); +lean_inc(x_481); +lean_dec(x_174); +switch (lean_obj_tag(x_481)) { case 0: { -lean_object* x_752; lean_object* x_753; lean_object* x_754; uint8_t x_755; lean_object* x_756; lean_object* x_757; lean_object* x_758; lean_object* x_759; lean_object* x_760; lean_object* x_761; lean_object* x_762; -lean_dec(x_751); -lean_dec(x_732); -lean_dec_ref(x_728); +lean_object* x_482; lean_object* x_483; lean_object* x_484; uint8_t x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; +lean_dec(x_159); +lean_dec_ref(x_155); +lean_dec(x_19); lean_dec(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -if (lean_is_exclusive(x_750)) { - lean_ctor_release(x_750, 0); - x_752 = x_750; -} else { - lean_dec_ref(x_750); - x_752 = lean_box(0); -} -x_753 = l_Lean_IR_ToIR_lowerLet___closed__7; -x_754 = l_Lean_IR_ToIR_lowerLet___closed__9; -x_755 = 1; -x_756 = l_Lean_Name_toString(x_727, x_755); -if (lean_is_scalar(x_752)) { - x_757 = lean_alloc_ctor(3, 1, 0); -} else { - x_757 = x_752; - lean_ctor_set_tag(x_757, 3); -} -lean_ctor_set(x_757, 0, x_756); -x_758 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_758, 0, x_754); -lean_ctor_set(x_758, 1, x_757); -x_759 = l_Lean_IR_ToIR_lowerLet___closed__11; -x_760 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_760, 0, x_758); -lean_ctor_set(x_760, 1, x_759); -x_761 = l_Lean_MessageData_ofFormat(x_760); -x_762 = l_Lean_throwNamedError___at___00Lean_IR_ToIR_lowerLet_spec__8___redArg(x_753, x_761, x_4, x_5); +if (lean_is_exclusive(x_481)) { + lean_ctor_release(x_481, 0); + x_482 = x_481; +} else { + lean_dec_ref(x_481); + x_482 = lean_box(0); +} +x_483 = l_Lean_IR_ToIR_lowerLet___closed__7; +x_484 = l_Lean_IR_ToIR_lowerLet___closed__9; +x_485 = 1; +x_486 = l_Lean_Name_toString(x_154, x_485); +if (lean_is_scalar(x_482)) { + x_487 = lean_alloc_ctor(3, 1, 0); +} else { + x_487 = x_482; + lean_ctor_set_tag(x_487, 3); +} +lean_ctor_set(x_487, 0, x_486); +x_488 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_488, 0, x_484); +lean_ctor_set(x_488, 1, x_487); +x_489 = l_Lean_IR_ToIR_lowerLet___closed__11; +x_490 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_490, 0, x_488); +lean_ctor_set(x_490, 1, x_489); +x_491 = l_Lean_MessageData_ofFormat(x_490); +x_492 = l_Lean_throwNamedError___at___00Lean_IR_ToIR_lowerLet_spec__8___redArg(x_483, x_491, x_4, x_5); lean_dec(x_5); lean_dec_ref(x_4); -return x_762; +return x_492; } case 2: { -lean_object* x_763; lean_object* x_764; lean_object* x_765; uint8_t x_766; lean_object* x_767; 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_dec(x_751); -lean_dec(x_732); -lean_dec_ref(x_728); +lean_object* x_493; lean_object* x_494; lean_object* x_495; uint8_t x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; +lean_dec(x_159); +lean_dec_ref(x_155); +lean_dec(x_19); lean_dec(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -if (lean_is_exclusive(x_750)) { - lean_ctor_release(x_750, 0); - x_763 = x_750; -} else { - lean_dec_ref(x_750); - x_763 = lean_box(0); -} -x_764 = l_Lean_IR_ToIR_lowerLet___closed__7; -x_765 = l_Lean_IR_ToIR_lowerLet___closed__9; -x_766 = 1; -x_767 = l_Lean_Name_toString(x_727, x_766); -if (lean_is_scalar(x_763)) { - x_768 = lean_alloc_ctor(3, 1, 0); -} else { - x_768 = x_763; - lean_ctor_set_tag(x_768, 3); -} -lean_ctor_set(x_768, 0, x_767); -x_769 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_769, 0, x_765); -lean_ctor_set(x_769, 1, x_768); -x_770 = l_Lean_IR_ToIR_lowerLet___closed__11; -x_771 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_771, 0, x_769); -lean_ctor_set(x_771, 1, x_770); -x_772 = l_Lean_MessageData_ofFormat(x_771); -x_773 = l_Lean_throwNamedError___at___00Lean_IR_ToIR_lowerLet_spec__8___redArg(x_764, x_772, x_4, x_5); +if (lean_is_exclusive(x_481)) { + lean_ctor_release(x_481, 0); + x_493 = x_481; +} else { + lean_dec_ref(x_481); + x_493 = lean_box(0); +} +x_494 = l_Lean_IR_ToIR_lowerLet___closed__7; +x_495 = l_Lean_IR_ToIR_lowerLet___closed__9; +x_496 = 1; +x_497 = l_Lean_Name_toString(x_154, x_496); +if (lean_is_scalar(x_493)) { + x_498 = lean_alloc_ctor(3, 1, 0); +} else { + x_498 = x_493; + lean_ctor_set_tag(x_498, 3); +} +lean_ctor_set(x_498, 0, x_497); +x_499 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_499, 0, x_495); +lean_ctor_set(x_499, 1, x_498); +x_500 = l_Lean_IR_ToIR_lowerLet___closed__11; +x_501 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_501, 0, x_499); +lean_ctor_set(x_501, 1, x_500); +x_502 = l_Lean_MessageData_ofFormat(x_501); +x_503 = l_Lean_throwNamedError___at___00Lean_IR_ToIR_lowerLet_spec__8___redArg(x_494, x_502, x_4, x_5); lean_dec(x_5); lean_dec_ref(x_4); -return x_773; +return x_503; } case 4: { -lean_object* x_774; lean_object* x_775; lean_object* x_776; uint8_t x_777; lean_object* x_778; lean_object* x_779; lean_object* x_780; lean_object* x_781; lean_object* x_782; lean_object* x_783; lean_object* x_784; -lean_dec(x_751); -lean_dec(x_732); -lean_dec_ref(x_728); +lean_object* x_504; lean_object* x_505; lean_object* x_506; uint8_t x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; +lean_dec(x_159); +lean_dec_ref(x_155); +lean_dec(x_19); lean_dec(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -if (lean_is_exclusive(x_750)) { - lean_ctor_release(x_750, 0); - x_774 = x_750; -} else { - lean_dec_ref(x_750); - x_774 = lean_box(0); -} -x_775 = l_Lean_IR_ToIR_lowerLet___closed__7; -x_776 = l_Lean_IR_ToIR_lowerLet___closed__9; -x_777 = 1; -x_778 = l_Lean_Name_toString(x_727, x_777); -if (lean_is_scalar(x_774)) { - x_779 = lean_alloc_ctor(3, 1, 0); -} else { - x_779 = x_774; - lean_ctor_set_tag(x_779, 3); -} -lean_ctor_set(x_779, 0, x_778); -x_780 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_780, 0, x_776); -lean_ctor_set(x_780, 1, x_779); -x_781 = l_Lean_IR_ToIR_lowerLet___closed__11; -x_782 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_782, 0, x_780); -lean_ctor_set(x_782, 1, x_781); -x_783 = l_Lean_MessageData_ofFormat(x_782); -x_784 = l_Lean_throwNamedError___at___00Lean_IR_ToIR_lowerLet_spec__8___redArg(x_775, x_783, x_4, x_5); +if (lean_is_exclusive(x_481)) { + lean_ctor_release(x_481, 0); + x_504 = x_481; +} else { + lean_dec_ref(x_481); + x_504 = lean_box(0); +} +x_505 = l_Lean_IR_ToIR_lowerLet___closed__7; +x_506 = l_Lean_IR_ToIR_lowerLet___closed__9; +x_507 = 1; +x_508 = l_Lean_Name_toString(x_154, x_507); +if (lean_is_scalar(x_504)) { + x_509 = lean_alloc_ctor(3, 1, 0); +} else { + x_509 = x_504; + lean_ctor_set_tag(x_509, 3); +} +lean_ctor_set(x_509, 0, x_508); +x_510 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_510, 0, x_506); +lean_ctor_set(x_510, 1, x_509); +x_511 = l_Lean_IR_ToIR_lowerLet___closed__11; +x_512 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_512, 0, x_510); +lean_ctor_set(x_512, 1, x_511); +x_513 = l_Lean_MessageData_ofFormat(x_512); +x_514 = l_Lean_throwNamedError___at___00Lean_IR_ToIR_lowerLet_spec__8___redArg(x_505, x_513, x_4, x_5); lean_dec(x_5); lean_dec_ref(x_4); -return x_784; +return x_514; } case 5: { -lean_object* x_785; lean_object* x_786; lean_object* x_787; uint8_t x_788; lean_object* x_789; lean_object* x_790; lean_object* x_791; lean_object* x_792; lean_object* x_793; lean_object* x_794; lean_object* x_795; -lean_dec(x_751); -lean_dec(x_732); -lean_dec_ref(x_728); +lean_object* x_515; lean_object* x_516; lean_object* x_517; uint8_t x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; +lean_dec(x_159); +lean_dec_ref(x_155); +lean_dec(x_19); lean_dec(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -if (lean_is_exclusive(x_750)) { - lean_ctor_release(x_750, 0); - x_785 = x_750; -} else { - lean_dec_ref(x_750); - x_785 = lean_box(0); -} -x_786 = l_Lean_IR_ToIR_lowerLet___closed__7; -x_787 = l_Lean_IR_ToIR_lowerLet___closed__9; -x_788 = 1; -x_789 = l_Lean_Name_toString(x_727, x_788); -if (lean_is_scalar(x_785)) { - x_790 = lean_alloc_ctor(3, 1, 0); -} else { - x_790 = x_785; - lean_ctor_set_tag(x_790, 3); -} -lean_ctor_set(x_790, 0, x_789); -x_791 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_791, 0, x_787); -lean_ctor_set(x_791, 1, x_790); -x_792 = l_Lean_IR_ToIR_lowerLet___closed__11; -x_793 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_793, 0, x_791); -lean_ctor_set(x_793, 1, x_792); -x_794 = l_Lean_MessageData_ofFormat(x_793); -x_795 = l_Lean_throwNamedError___at___00Lean_IR_ToIR_lowerLet_spec__8___redArg(x_786, x_794, x_4, x_5); +if (lean_is_exclusive(x_481)) { + lean_ctor_release(x_481, 0); + x_515 = x_481; +} else { + lean_dec_ref(x_481); + x_515 = lean_box(0); +} +x_516 = l_Lean_IR_ToIR_lowerLet___closed__7; +x_517 = l_Lean_IR_ToIR_lowerLet___closed__9; +x_518 = 1; +x_519 = l_Lean_Name_toString(x_154, x_518); +if (lean_is_scalar(x_515)) { + x_520 = lean_alloc_ctor(3, 1, 0); +} else { + x_520 = x_515; + lean_ctor_set_tag(x_520, 3); +} +lean_ctor_set(x_520, 0, x_519); +x_521 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_521, 0, x_517); +lean_ctor_set(x_521, 1, x_520); +x_522 = l_Lean_IR_ToIR_lowerLet___closed__11; +x_523 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_523, 0, x_521); +lean_ctor_set(x_523, 1, x_522); +x_524 = l_Lean_MessageData_ofFormat(x_523); +x_525 = l_Lean_throwNamedError___at___00Lean_IR_ToIR_lowerLet_spec__8___redArg(x_516, x_524, x_4, x_5); lean_dec(x_5); lean_dec_ref(x_4); -return x_795; +return x_525; } case 6: { -lean_object* x_796; lean_object* x_797; lean_object* x_798; lean_object* x_799; lean_object* x_800; lean_object* x_801; -lean_inc(x_630); +lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; +lean_inc_ref(x_16); +lean_inc(x_15); lean_dec_ref(x_1); -x_796 = lean_ctor_get(x_750, 0); -lean_inc_ref(x_796); -if (lean_is_exclusive(x_750)) { - lean_ctor_release(x_750, 0); - x_797 = x_750; -} else { - lean_dec_ref(x_750); - x_797 = lean_box(0); -} -x_798 = lean_ctor_get(x_796, 1); -lean_inc(x_798); -x_799 = lean_ctor_get(x_796, 2); -lean_inc(x_799); -x_800 = lean_ctor_get(x_796, 3); -lean_inc(x_800); -lean_dec_ref(x_796); +x_526 = lean_ctor_get(x_481, 0); +lean_inc_ref(x_526); +if (lean_is_exclusive(x_481)) { + lean_ctor_release(x_481, 0); + x_527 = x_481; +} else { + lean_dec_ref(x_481); + x_527 = lean_box(0); +} +x_528 = lean_ctor_get(x_526, 1); +lean_inc(x_528); +x_529 = lean_ctor_get(x_526, 2); +lean_inc(x_529); +x_530 = lean_ctor_get(x_526, 3); +lean_inc(x_530); +lean_dec_ref(x_526); lean_inc(x_5); lean_inc_ref(x_4); -lean_inc(x_798); -x_801 = l_Lean_IR_hasTrivialStructure_x3f(x_798, x_4, x_5); -if (lean_obj_tag(x_801) == 0) -{ -lean_object* x_802; -x_802 = lean_ctor_get(x_801, 0); -lean_inc(x_802); -lean_dec_ref(x_801); -if (lean_obj_tag(x_802) == 1) -{ -lean_object* x_803; lean_object* x_804; lean_object* x_805; lean_object* x_806; lean_object* x_807; lean_object* x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; lean_object* x_812; lean_object* x_813; lean_object* x_814; lean_object* x_815; lean_object* x_816; lean_object* x_817; lean_object* x_818; -lean_dec(x_800); -lean_dec(x_799); -lean_dec(x_798); -lean_dec(x_797); -lean_dec(x_751); -lean_dec(x_732); -lean_dec(x_727); -x_803 = lean_ctor_get(x_802, 0); -lean_inc(x_803); -lean_dec_ref(x_802); -x_804 = lean_st_ref_take(x_3); -x_805 = lean_ctor_get(x_803, 1); -lean_inc(x_805); -x_806 = lean_ctor_get(x_803, 2); -lean_inc(x_806); -lean_dec(x_803); -x_807 = lean_ctor_get(x_804, 0); -lean_inc_ref(x_807); -x_808 = lean_ctor_get(x_804, 1); -lean_inc_ref(x_808); -x_809 = lean_ctor_get(x_804, 2); -lean_inc(x_809); -x_810 = lean_ctor_get(x_804, 3); -lean_inc_ref(x_810); -if (lean_is_exclusive(x_804)) { - lean_ctor_release(x_804, 0); - lean_ctor_release(x_804, 1); - lean_ctor_release(x_804, 2); - lean_ctor_release(x_804, 3); - x_811 = x_804; -} else { - lean_dec_ref(x_804); - x_811 = lean_box(0); -} -x_812 = lean_box(0); -x_813 = lean_nat_add(x_805, x_806); -lean_dec(x_806); -lean_dec(x_805); -x_814 = lean_array_get(x_812, x_728, x_813); -lean_dec(x_813); -lean_dec_ref(x_728); -x_815 = l_Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_ToIR_lowerParam_spec__0___redArg(x_810, x_630, x_814); -if (lean_is_scalar(x_811)) { - x_816 = lean_alloc_ctor(0, 4, 0); -} else { - x_816 = x_811; -} -lean_ctor_set(x_816, 0, x_807); -lean_ctor_set(x_816, 1, x_808); -lean_ctor_set(x_816, 2, x_809); -lean_ctor_set(x_816, 3, x_815); -x_817 = lean_st_ref_set(x_3, x_816); -x_818 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); -return x_818; -} -else -{ -lean_object* x_819; -lean_dec(x_802); -lean_dec_ref(x_728); +lean_inc(x_528); +x_531 = l_Lean_IR_hasTrivialStructure_x3f(x_528, x_4, x_5); +if (lean_obj_tag(x_531) == 0) +{ +lean_object* x_532; +x_532 = lean_ctor_get(x_531, 0); +lean_inc(x_532); +lean_dec_ref(x_531); +if (lean_obj_tag(x_532) == 1) +{ +lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; +lean_dec(x_530); +lean_dec(x_529); +lean_dec(x_528); +lean_dec(x_527); +lean_dec(x_159); +lean_dec(x_154); +lean_dec(x_19); +lean_dec_ref(x_16); +x_533 = lean_ctor_get(x_532, 0); +lean_inc(x_533); +lean_dec_ref(x_532); +x_534 = lean_st_ref_take(x_3); +x_535 = lean_ctor_get(x_533, 1); +lean_inc(x_535); +x_536 = lean_ctor_get(x_533, 2); +lean_inc(x_536); +lean_dec(x_533); +x_537 = lean_ctor_get(x_534, 0); +lean_inc_ref(x_537); +x_538 = lean_ctor_get(x_534, 1); +lean_inc_ref(x_538); +x_539 = lean_ctor_get(x_534, 2); +lean_inc(x_539); +x_540 = lean_ctor_get(x_534, 3); +lean_inc_ref(x_540); +if (lean_is_exclusive(x_534)) { + lean_ctor_release(x_534, 0); + lean_ctor_release(x_534, 1); + lean_ctor_release(x_534, 2); + lean_ctor_release(x_534, 3); + x_541 = x_534; +} else { + lean_dec_ref(x_534); + x_541 = lean_box(0); +} +x_542 = lean_box(0); +x_543 = lean_nat_add(x_535, x_536); +lean_dec(x_536); +lean_dec(x_535); +x_544 = lean_array_get(x_542, x_155, x_543); +lean_dec(x_543); +lean_dec_ref(x_155); +x_545 = l_Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_ToIR_lowerParam_spec__0___redArg(x_540, x_15, x_544); +if (lean_is_scalar(x_541)) { + x_546 = lean_alloc_ctor(0, 4, 0); +} else { + x_546 = x_541; +} +lean_ctor_set(x_546, 0, x_537); +lean_ctor_set(x_546, 1, x_538); +lean_ctor_set(x_546, 2, x_539); +lean_ctor_set(x_546, 3, x_545); +x_547 = lean_st_ref_set(x_3, x_546); +x_548 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); +return x_548; +} +else +{ +lean_object* x_549; +lean_dec(x_532); +lean_dec_ref(x_155); lean_inc(x_5); lean_inc_ref(x_4); -x_819 = l_Lean_IR_nameToIRType(x_798, x_4, x_5); -if (lean_obj_tag(x_819) == 0) -{ -lean_object* x_820; uint8_t x_821; -x_820 = lean_ctor_get(x_819, 0); -lean_inc(x_820); -lean_dec_ref(x_819); -x_821 = l_Lean_IR_IRType_isScalar(x_820); -if (x_821 == 0) -{ -lean_object* x_822; -lean_dec(x_820); -lean_dec(x_799); -lean_dec(x_797); -lean_dec(x_751); +lean_inc(x_528); +x_549 = l_Lean_IR_nameToIRType(x_528, x_4, x_5); +if (lean_obj_tag(x_549) == 0) +{ +lean_object* x_550; uint8_t x_551; +x_550 = lean_ctor_get(x_549, 0); +lean_inc(x_550); +lean_dec_ref(x_549); +x_551 = l_Lean_IR_IRType_isScalar(x_550); +if (x_551 == 0) +{ +lean_object* x_552; +lean_dec(x_550); +lean_dec(x_529); +lean_dec(x_527); lean_inc(x_5); lean_inc_ref(x_4); -x_822 = l_Lean_IR_getCtorLayout(x_727, x_4, x_5); -if (lean_obj_tag(x_822) == 0) -{ -lean_object* x_823; lean_object* x_824; lean_object* x_828; lean_object* x_829; lean_object* x_830; lean_object* x_831; lean_object* x_832; lean_object* x_833; lean_object* x_834; uint8_t x_835; -x_823 = lean_ctor_get(x_822, 0); -lean_inc(x_823); -if (lean_is_exclusive(x_822)) { - lean_ctor_release(x_822, 0); - x_824 = x_822; -} else { - lean_dec_ref(x_822); - x_824 = lean_box(0); -} -x_828 = lean_ctor_get(x_823, 0); -lean_inc_ref(x_828); -x_829 = lean_ctor_get(x_823, 1); -lean_inc_ref(x_829); -if (lean_is_exclusive(x_823)) { - lean_ctor_release(x_823, 0); - lean_ctor_release(x_823, 1); - x_830 = x_823; -} else { - lean_dec_ref(x_823); - x_830 = lean_box(0); -} -x_831 = lean_array_get_size(x_732); -x_832 = l_Array_extract___redArg(x_732, x_800, x_831); -lean_dec(x_732); -x_833 = lean_array_get_size(x_832); -x_834 = lean_array_get_size(x_829); -x_835 = lean_nat_dec_eq(x_833, x_834); -if (x_835 == 0) -{ -lean_dec_ref(x_832); -lean_dec(x_830); -lean_dec_ref(x_829); -lean_dec_ref(x_828); -lean_dec(x_630); +x_552 = l_Lean_IR_getCtorLayout(x_154, x_4, x_5); +if (lean_obj_tag(x_552) == 0) +{ +lean_object* x_553; lean_object* x_554; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; uint8_t x_565; +x_553 = lean_ctor_get(x_552, 0); +lean_inc(x_553); +if (lean_is_exclusive(x_552)) { + lean_ctor_release(x_552, 0); + x_554 = x_552; +} else { + lean_dec_ref(x_552); + x_554 = lean_box(0); +} +x_558 = lean_ctor_get(x_553, 0); +lean_inc_ref(x_558); +x_559 = lean_ctor_get(x_553, 1); +lean_inc_ref(x_559); +if (lean_is_exclusive(x_553)) { + lean_ctor_release(x_553, 0); + lean_ctor_release(x_553, 1); + x_560 = x_553; +} else { + lean_dec_ref(x_553); + x_560 = lean_box(0); +} +x_561 = lean_array_get_size(x_159); +x_562 = l_Array_extract___redArg(x_159, x_530, x_561); +lean_dec(x_159); +x_563 = lean_array_get_size(x_562); +x_564 = lean_array_get_size(x_559); +x_565 = lean_nat_dec_eq(x_563, x_564); +if (x_565 == 0) +{ +lean_dec_ref(x_562); +lean_dec(x_560); +lean_dec_ref(x_559); +lean_dec_ref(x_558); +lean_dec(x_528); +lean_dec(x_19); +lean_dec_ref(x_16); +lean_dec(x_15); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -goto block_827; +goto block_557; +} +else +{ +if (x_551 == 0) +{ +lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; +lean_dec(x_554); +x_566 = l_Lean_IR_ToIR_getFVarValue___redArg___closed__0; +x_567 = lean_unsigned_to_nat(0u); +x_568 = l_Lean_IR_ToIR_lowerLet___closed__12; +x_569 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_ToIR_lowerLet_spec__9___redArg(x_564, x_559, x_566, x_562, x_567, x_568); +if (lean_obj_tag(x_569) == 0) +{ +lean_object* x_570; lean_object* x_571; +x_570 = lean_ctor_get(x_569, 0); +lean_inc(x_570); +lean_dec_ref(x_569); +x_571 = l_Lean_IR_ToIR_bindVar___redArg(x_15, x_3); +if (lean_obj_tag(x_571) == 0) +{ +lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_585; lean_object* x_586; uint8_t x_587; +x_572 = lean_ctor_get(x_571, 0); +lean_inc(x_572); +lean_dec_ref(x_571); +x_585 = lean_st_ref_get(x_5); +x_586 = lean_ctor_get(x_585, 0); +lean_inc_ref(x_586); +lean_dec(x_585); +x_587 = l_Lean_IR_UnboxResult_hasUnboxAttr(x_586, x_528); +if (x_587 == 0) +{ +lean_object* x_588; +lean_dec_ref(x_16); +x_588 = l_Lean_IR_CtorInfo_type(x_558); +x_573 = x_588; +x_574 = x_3; +x_575 = x_4; +x_576 = x_5; +x_577 = lean_box(0); +goto block_584; } else { -if (x_821 == 0) +lean_object* x_589; +lean_inc(x_5); +lean_inc_ref(x_4); +x_589 = l_Lean_IR_toIRType(x_16, x_4, x_5); +if (lean_obj_tag(x_589) == 0) { -lean_object* x_836; lean_object* x_837; lean_object* x_838; lean_object* x_839; -lean_dec(x_824); -x_836 = l_Lean_IR_ToIR_getFVarValue___redArg___closed__0; -x_837 = lean_unsigned_to_nat(0u); -x_838 = l_Lean_IR_ToIR_lowerLet___closed__12; -x_839 = l_WellFounded_opaqueFix_u2083___at___00Lean_IR_ToIR_lowerLet_spec__9___redArg(x_834, x_829, x_836, x_832, x_837, x_838); -if (lean_obj_tag(x_839) == 0) +lean_object* x_590; +x_590 = lean_ctor_get(x_589, 0); +lean_inc(x_590); +lean_dec_ref(x_589); +x_573 = x_590; +x_574 = x_3; +x_575 = x_4; +x_576 = x_5; +x_577 = lean_box(0); +goto block_584; +} +else { -lean_object* x_840; lean_object* x_841; -x_840 = lean_ctor_get(x_839, 0); -lean_inc(x_840); -lean_dec_ref(x_839); -x_841 = l_Lean_IR_ToIR_bindVar___redArg(x_630, x_3); -if (lean_obj_tag(x_841) == 0) +lean_object* x_591; lean_object* x_592; lean_object* x_593; +lean_dec(x_572); +lean_dec(x_570); +lean_dec_ref(x_562); +lean_dec(x_560); +lean_dec_ref(x_559); +lean_dec_ref(x_558); +lean_dec(x_19); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec(x_3); +lean_dec_ref(x_2); +x_591 = lean_ctor_get(x_589, 0); +lean_inc(x_591); +if (lean_is_exclusive(x_589)) { + lean_ctor_release(x_589, 0); + x_592 = x_589; +} else { + lean_dec_ref(x_589); + x_592 = lean_box(0); +} +if (lean_is_scalar(x_592)) { + x_593 = lean_alloc_ctor(1, 1, 0); +} else { + x_593 = x_592; +} +lean_ctor_set(x_593, 0, x_591); +return x_593; +} +} +block_584: { -lean_object* x_842; lean_object* x_843; -x_842 = lean_ctor_get(x_841, 0); -lean_inc(x_842); -lean_dec_ref(x_841); -lean_inc(x_842); -x_843 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_lowerNonObjectFields___redArg(x_2, x_828, x_829, x_832, x_842, x_3, x_4, x_5); -lean_dec_ref(x_832); -lean_dec_ref(x_829); -if (lean_obj_tag(x_843) == 0) +lean_object* x_578; +lean_inc(x_572); +lean_inc_ref(x_558); +x_578 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_lowerNonObjectFields___redArg(x_2, x_558, x_559, x_562, x_572, x_574, x_575, x_576); +lean_dec_ref(x_562); +lean_dec_ref(x_559); +if (lean_obj_tag(x_578) == 0) { -lean_object* x_844; lean_object* x_845; lean_object* x_846; lean_object* x_847; lean_object* x_848; lean_object* x_849; -x_844 = lean_ctor_get(x_843, 0); -lean_inc(x_844); -if (lean_is_exclusive(x_843)) { - lean_ctor_release(x_843, 0); - x_845 = x_843; +lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; +x_579 = lean_ctor_get(x_578, 0); +lean_inc(x_579); +if (lean_is_exclusive(x_578)) { + lean_ctor_release(x_578, 0); + x_580 = x_578; +} else { + lean_dec_ref(x_578); + x_580 = lean_box(0); +} +if (lean_is_scalar(x_560)) { + x_581 = lean_alloc_ctor(0, 2, 0); } else { - lean_dec_ref(x_843); - x_845 = lean_box(0); + x_581 = x_560; } -x_846 = l_Lean_IR_CtorInfo_type(x_828); -if (lean_is_scalar(x_830)) { - x_847 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_581, 0, x_558); +lean_ctor_set(x_581, 1, x_570); +if (lean_is_scalar(x_19)) { + x_582 = lean_alloc_ctor(0, 4, 0); } else { - x_847 = x_830; + x_582 = x_19; } -lean_ctor_set(x_847, 0, x_828); -lean_ctor_set(x_847, 1, x_840); -x_848 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_848, 0, x_842); -lean_ctor_set(x_848, 1, x_846); -lean_ctor_set(x_848, 2, x_847); -lean_ctor_set(x_848, 3, x_844); -if (lean_is_scalar(x_845)) { - x_849 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_582, 0, x_572); +lean_ctor_set(x_582, 1, x_573); +lean_ctor_set(x_582, 2, x_581); +lean_ctor_set(x_582, 3, x_579); +if (lean_is_scalar(x_580)) { + x_583 = lean_alloc_ctor(0, 1, 0); } else { - x_849 = x_845; + x_583 = x_580; } -lean_ctor_set(x_849, 0, x_848); -return x_849; +lean_ctor_set(x_583, 0, x_582); +return x_583; } else { -lean_dec(x_842); -lean_dec(x_840); -lean_dec(x_830); -lean_dec_ref(x_828); -return x_843; +lean_dec(x_573); +lean_dec(x_572); +lean_dec(x_570); +lean_dec(x_560); +lean_dec_ref(x_558); +lean_dec(x_19); +return x_578; +} } } else { -lean_object* x_850; lean_object* x_851; lean_object* x_852; -lean_dec(x_840); -lean_dec_ref(x_832); -lean_dec(x_830); -lean_dec_ref(x_829); -lean_dec_ref(x_828); +lean_object* x_594; lean_object* x_595; lean_object* x_596; +lean_dec(x_570); +lean_dec_ref(x_562); +lean_dec(x_560); +lean_dec_ref(x_559); +lean_dec_ref(x_558); +lean_dec(x_528); +lean_dec(x_19); +lean_dec_ref(x_16); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -x_850 = lean_ctor_get(x_841, 0); -lean_inc(x_850); -if (lean_is_exclusive(x_841)) { - lean_ctor_release(x_841, 0); - x_851 = x_841; +x_594 = lean_ctor_get(x_571, 0); +lean_inc(x_594); +if (lean_is_exclusive(x_571)) { + lean_ctor_release(x_571, 0); + x_595 = x_571; } else { - lean_dec_ref(x_841); - x_851 = lean_box(0); + lean_dec_ref(x_571); + x_595 = lean_box(0); } -if (lean_is_scalar(x_851)) { - x_852 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_595)) { + x_596 = lean_alloc_ctor(1, 1, 0); } else { - x_852 = x_851; + x_596 = x_595; } -lean_ctor_set(x_852, 0, x_850); -return x_852; +lean_ctor_set(x_596, 0, x_594); +return x_596; } } else { -lean_object* x_853; lean_object* x_854; lean_object* x_855; -lean_dec_ref(x_832); -lean_dec(x_830); -lean_dec_ref(x_829); -lean_dec_ref(x_828); -lean_dec(x_630); +lean_object* x_597; lean_object* x_598; lean_object* x_599; +lean_dec_ref(x_562); +lean_dec(x_560); +lean_dec_ref(x_559); +lean_dec_ref(x_558); +lean_dec(x_528); +lean_dec(x_19); +lean_dec_ref(x_16); +lean_dec(x_15); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -x_853 = lean_ctor_get(x_839, 0); -lean_inc(x_853); -if (lean_is_exclusive(x_839)) { - lean_ctor_release(x_839, 0); - x_854 = x_839; +x_597 = lean_ctor_get(x_569, 0); +lean_inc(x_597); +if (lean_is_exclusive(x_569)) { + lean_ctor_release(x_569, 0); + x_598 = x_569; } else { - lean_dec_ref(x_839); - x_854 = lean_box(0); + lean_dec_ref(x_569); + x_598 = lean_box(0); } -if (lean_is_scalar(x_854)) { - x_855 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_598)) { + x_599 = lean_alloc_ctor(1, 1, 0); } else { - x_855 = x_854; + x_599 = x_598; } -lean_ctor_set(x_855, 0, x_853); -return x_855; +lean_ctor_set(x_599, 0, x_597); +return x_599; } } else { -lean_dec_ref(x_832); -lean_dec(x_830); -lean_dec_ref(x_829); -lean_dec_ref(x_828); -lean_dec(x_630); +lean_dec_ref(x_562); +lean_dec(x_560); +lean_dec_ref(x_559); +lean_dec_ref(x_558); +lean_dec(x_528); +lean_dec(x_19); +lean_dec_ref(x_16); +lean_dec(x_15); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -goto block_827; +goto block_557; } } -block_827: +block_557: { -lean_object* x_825; lean_object* x_826; -x_825 = lean_box(12); -if (lean_is_scalar(x_824)) { - x_826 = lean_alloc_ctor(0, 1, 0); +lean_object* x_555; lean_object* x_556; +x_555 = lean_box(12); +if (lean_is_scalar(x_554)) { + x_556 = lean_alloc_ctor(0, 1, 0); } else { - x_826 = x_824; + x_556 = x_554; } -lean_ctor_set(x_826, 0, x_825); -return x_826; +lean_ctor_set(x_556, 0, x_555); +return x_556; } } else { -lean_object* x_856; lean_object* x_857; lean_object* x_858; -lean_dec(x_800); -lean_dec(x_732); -lean_dec(x_630); +lean_object* x_600; lean_object* x_601; lean_object* x_602; +lean_dec(x_530); +lean_dec(x_528); +lean_dec(x_159); +lean_dec(x_19); +lean_dec_ref(x_16); +lean_dec(x_15); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -x_856 = lean_ctor_get(x_822, 0); -lean_inc(x_856); -if (lean_is_exclusive(x_822)) { - lean_ctor_release(x_822, 0); - x_857 = x_822; +x_600 = lean_ctor_get(x_552, 0); +lean_inc(x_600); +if (lean_is_exclusive(x_552)) { + lean_ctor_release(x_552, 0); + x_601 = x_552; } else { - lean_dec_ref(x_822); - x_857 = lean_box(0); + lean_dec_ref(x_552); + x_601 = lean_box(0); } -if (lean_is_scalar(x_857)) { - x_858 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_601)) { + x_602 = lean_alloc_ctor(1, 1, 0); } else { - x_858 = x_857; + x_602 = x_601; } -lean_ctor_set(x_858, 0, x_856); -return x_858; +lean_ctor_set(x_602, 0, x_600); +return x_602; } } else { -lean_object* x_859; -lean_dec(x_800); -lean_dec(x_732); -lean_dec(x_727); -x_859 = l_Lean_IR_ToIR_bindVar___redArg(x_630, x_3); -if (lean_obj_tag(x_859) == 0) -{ -lean_object* x_860; lean_object* x_861; -x_860 = lean_ctor_get(x_859, 0); -lean_inc(x_860); -lean_dec_ref(x_859); -x_861 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); -if (lean_obj_tag(x_861) == 0) -{ -lean_object* x_862; lean_object* x_863; lean_object* x_864; lean_object* x_865; lean_object* x_866; lean_object* x_867; -x_862 = lean_ctor_get(x_861, 0); -lean_inc(x_862); -if (lean_is_exclusive(x_861)) { - lean_ctor_release(x_861, 0); - x_863 = x_861; -} else { - lean_dec_ref(x_861); - x_863 = lean_box(0); -} -if (lean_is_scalar(x_797)) { - x_864 = lean_alloc_ctor(0, 1, 0); -} else { - x_864 = x_797; - lean_ctor_set_tag(x_864, 0); -} -lean_ctor_set(x_864, 0, x_799); -if (lean_is_scalar(x_751)) { - x_865 = lean_alloc_ctor(11, 1, 0); +lean_object* x_603; +lean_dec(x_530); +lean_dec(x_528); +lean_dec(x_159); +lean_dec(x_154); +lean_dec_ref(x_16); +x_603 = l_Lean_IR_ToIR_bindVar___redArg(x_15, x_3); +if (lean_obj_tag(x_603) == 0) +{ +lean_object* x_604; lean_object* x_605; +x_604 = lean_ctor_get(x_603, 0); +lean_inc(x_604); +lean_dec_ref(x_603); +x_605 = l_Lean_IR_ToIR_lowerCode(x_2, x_3, x_4, x_5); +if (lean_obj_tag(x_605) == 0) +{ +lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; +x_606 = lean_ctor_get(x_605, 0); +lean_inc(x_606); +if (lean_is_exclusive(x_605)) { + lean_ctor_release(x_605, 0); + x_607 = x_605; +} else { + lean_dec_ref(x_605); + x_607 = lean_box(0); +} +if (lean_is_scalar(x_527)) { + x_608 = lean_alloc_ctor(0, 1, 0); +} else { + x_608 = x_527; + lean_ctor_set_tag(x_608, 0); +} +lean_ctor_set(x_608, 0, x_529); +x_609 = lean_alloc_ctor(11, 1, 0); +lean_ctor_set(x_609, 0, x_608); +if (lean_is_scalar(x_19)) { + x_610 = lean_alloc_ctor(0, 4, 0); } else { - x_865 = x_751; - lean_ctor_set_tag(x_865, 11); + x_610 = x_19; } -lean_ctor_set(x_865, 0, x_864); -x_866 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_866, 0, x_860); -lean_ctor_set(x_866, 1, x_820); -lean_ctor_set(x_866, 2, x_865); -lean_ctor_set(x_866, 3, x_862); -if (lean_is_scalar(x_863)) { - x_867 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_610, 0, x_604); +lean_ctor_set(x_610, 1, x_550); +lean_ctor_set(x_610, 2, x_609); +lean_ctor_set(x_610, 3, x_606); +if (lean_is_scalar(x_607)) { + x_611 = lean_alloc_ctor(0, 1, 0); } else { - x_867 = x_863; + x_611 = x_607; } -lean_ctor_set(x_867, 0, x_866); -return x_867; +lean_ctor_set(x_611, 0, x_610); +return x_611; } else { -lean_dec(x_860); -lean_dec(x_820); -lean_dec(x_799); -lean_dec(x_797); -lean_dec(x_751); -return x_861; +lean_dec(x_604); +lean_dec(x_550); +lean_dec(x_529); +lean_dec(x_527); +lean_dec(x_19); +return x_605; } } else { -lean_object* x_868; lean_object* x_869; lean_object* x_870; -lean_dec(x_820); -lean_dec(x_799); -lean_dec(x_797); -lean_dec(x_751); +lean_object* x_612; lean_object* x_613; lean_object* x_614; +lean_dec(x_550); +lean_dec(x_529); +lean_dec(x_527); +lean_dec(x_19); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -x_868 = lean_ctor_get(x_859, 0); -lean_inc(x_868); -if (lean_is_exclusive(x_859)) { - lean_ctor_release(x_859, 0); - x_869 = x_859; +x_612 = lean_ctor_get(x_603, 0); +lean_inc(x_612); +if (lean_is_exclusive(x_603)) { + lean_ctor_release(x_603, 0); + x_613 = x_603; } else { - lean_dec_ref(x_859); - x_869 = lean_box(0); + lean_dec_ref(x_603); + x_613 = lean_box(0); } -if (lean_is_scalar(x_869)) { - x_870 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_613)) { + x_614 = lean_alloc_ctor(1, 1, 0); } else { - x_870 = x_869; + x_614 = x_613; } -lean_ctor_set(x_870, 0, x_868); -return x_870; +lean_ctor_set(x_614, 0, x_612); +return x_614; } } } else { -lean_object* x_871; lean_object* x_872; lean_object* x_873; -lean_dec(x_800); -lean_dec(x_799); -lean_dec(x_797); -lean_dec(x_751); -lean_dec(x_732); -lean_dec(x_727); -lean_dec(x_630); +lean_object* x_615; lean_object* x_616; lean_object* x_617; +lean_dec(x_530); +lean_dec(x_529); +lean_dec(x_528); +lean_dec(x_527); +lean_dec(x_159); +lean_dec(x_154); +lean_dec(x_19); +lean_dec_ref(x_16); +lean_dec(x_15); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -x_871 = lean_ctor_get(x_819, 0); -lean_inc(x_871); -if (lean_is_exclusive(x_819)) { - lean_ctor_release(x_819, 0); - x_872 = x_819; +x_615 = lean_ctor_get(x_549, 0); +lean_inc(x_615); +if (lean_is_exclusive(x_549)) { + lean_ctor_release(x_549, 0); + x_616 = x_549; } else { - lean_dec_ref(x_819); - x_872 = lean_box(0); + lean_dec_ref(x_549); + x_616 = lean_box(0); } -if (lean_is_scalar(x_872)) { - x_873 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_616)) { + x_617 = lean_alloc_ctor(1, 1, 0); } else { - x_873 = x_872; + x_617 = x_616; } -lean_ctor_set(x_873, 0, x_871); -return x_873; +lean_ctor_set(x_617, 0, x_615); +return x_617; } } } else { -lean_object* x_874; lean_object* x_875; lean_object* x_876; -lean_dec(x_800); -lean_dec(x_799); -lean_dec(x_798); -lean_dec(x_797); -lean_dec(x_751); -lean_dec(x_732); -lean_dec_ref(x_728); -lean_dec(x_727); -lean_dec(x_630); +lean_object* x_618; lean_object* x_619; lean_object* x_620; +lean_dec(x_530); +lean_dec(x_529); +lean_dec(x_528); +lean_dec(x_527); +lean_dec(x_159); +lean_dec_ref(x_155); +lean_dec(x_154); +lean_dec(x_19); +lean_dec_ref(x_16); +lean_dec(x_15); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -x_874 = lean_ctor_get(x_801, 0); -lean_inc(x_874); -if (lean_is_exclusive(x_801)) { - lean_ctor_release(x_801, 0); - x_875 = x_801; +x_618 = lean_ctor_get(x_531, 0); +lean_inc(x_618); +if (lean_is_exclusive(x_531)) { + lean_ctor_release(x_531, 0); + x_619 = x_531; } else { - lean_dec_ref(x_801); - x_875 = lean_box(0); + lean_dec_ref(x_531); + x_619 = lean_box(0); } -if (lean_is_scalar(x_875)) { - x_876 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_619)) { + x_620 = lean_alloc_ctor(1, 1, 0); } else { - x_876 = x_875; + x_620 = x_619; } -lean_ctor_set(x_876, 0, x_874); -return x_876; +lean_ctor_set(x_620, 0, x_618); +return x_620; } } case 7: { -lean_object* x_877; lean_object* x_878; uint8_t x_879; lean_object* x_880; 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_dec(x_751); -lean_dec(x_732); -lean_dec_ref(x_728); +lean_object* x_621; lean_object* x_622; uint8_t x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; +lean_dec(x_159); +lean_dec_ref(x_155); +lean_dec(x_19); lean_dec(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -if (lean_is_exclusive(x_750)) { - lean_ctor_release(x_750, 0); - x_877 = x_750; -} else { - lean_dec_ref(x_750); - x_877 = lean_box(0); -} -x_878 = l_Lean_IR_ToIR_lowerLet___closed__14; -x_879 = 1; -x_880 = l_Lean_Name_toString(x_727, x_879); -if (lean_is_scalar(x_877)) { - x_881 = lean_alloc_ctor(3, 1, 0); -} else { - x_881 = x_877; - lean_ctor_set_tag(x_881, 3); -} -lean_ctor_set(x_881, 0, x_880); -x_882 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_882, 0, x_878); -lean_ctor_set(x_882, 1, x_881); -x_883 = l_Lean_IR_ToIR_lowerLet___closed__16; -x_884 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_884, 0, x_882); -lean_ctor_set(x_884, 1, x_883); -x_885 = l_Lean_MessageData_ofFormat(x_884); -x_886 = l_Lean_throwError___at___00Lean_IR_ToIR_lowerLet_spec__10___redArg(x_885, x_4, x_5); +if (lean_is_exclusive(x_481)) { + lean_ctor_release(x_481, 0); + x_621 = x_481; +} else { + lean_dec_ref(x_481); + x_621 = lean_box(0); +} +x_622 = l_Lean_IR_ToIR_lowerLet___closed__14; +x_623 = 1; +x_624 = l_Lean_Name_toString(x_154, x_623); +if (lean_is_scalar(x_621)) { + x_625 = lean_alloc_ctor(3, 1, 0); +} else { + x_625 = x_621; + lean_ctor_set_tag(x_625, 3); +} +lean_ctor_set(x_625, 0, x_624); +x_626 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_626, 0, x_622); +lean_ctor_set(x_626, 1, x_625); +x_627 = l_Lean_IR_ToIR_lowerLet___closed__16; +x_628 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_628, 0, x_626); +lean_ctor_set(x_628, 1, x_627); +x_629 = l_Lean_MessageData_ofFormat(x_628); +x_630 = l_Lean_throwError___at___00Lean_IR_ToIR_lowerLet_spec__10___redArg(x_629, x_4, x_5); lean_dec(x_5); lean_dec_ref(x_4); -return x_886; +return x_630; } default: { -lean_object* x_887; -lean_dec(x_751); -lean_dec(x_750); -lean_dec_ref(x_728); -x_887 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_mkFap(x_1, x_2, x_727, x_732, x_3, x_4, x_5); -return x_887; +lean_object* x_631; +lean_dec(x_481); +lean_dec_ref(x_155); +lean_dec(x_19); +x_631 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_mkFap(x_1, x_2, x_154, x_159, x_3, x_4, x_5); +return x_631; +} } } } @@ -9764,184 +8347,182 @@ return x_887; } else { -lean_object* x_888; lean_object* x_889; lean_object* x_890; -lean_dec(x_732); -lean_dec_ref(x_728); -lean_dec(x_727); +uint8_t x_632; +lean_dec(x_159); +lean_dec_ref(x_155); +lean_dec(x_154); +lean_dec(x_19); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_888 = lean_ctor_get(x_739, 0); -lean_inc(x_888); -if (lean_is_exclusive(x_739)) { - lean_ctor_release(x_739, 0); - x_889 = x_739; -} else { - lean_dec_ref(x_739); - x_889 = lean_box(0); +x_632 = !lean_is_exclusive(x_166); +if (x_632 == 0) +{ +return x_166; } -if (lean_is_scalar(x_889)) { - x_890 = lean_alloc_ctor(1, 1, 0); -} else { - x_890 = x_889; +else +{ +lean_object* x_633; lean_object* x_634; +x_633 = lean_ctor_get(x_166, 0); +lean_inc(x_633); +lean_dec(x_166); +x_634 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_634, 0, x_633); +return x_634; } -lean_ctor_set(x_890, 0, x_888); -return x_890; } } } else { -lean_object* x_891; lean_object* x_892; lean_object* x_893; -lean_dec(x_732); -lean_dec_ref(x_728); -lean_dec(x_727); +uint8_t x_635; +lean_dec(x_159); +lean_dec_ref(x_155); +lean_dec(x_154); +lean_dec(x_19); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_891 = lean_ctor_get(x_733, 0); -lean_inc(x_891); -if (lean_is_exclusive(x_733)) { - lean_ctor_release(x_733, 0); - x_892 = x_733; -} else { - lean_dec_ref(x_733); - x_892 = lean_box(0); +x_635 = !lean_is_exclusive(x_160); +if (x_635 == 0) +{ +return x_160; } -if (lean_is_scalar(x_892)) { - x_893 = lean_alloc_ctor(1, 1, 0); -} else { - x_893 = x_892; +else +{ +lean_object* x_636; lean_object* x_637; +x_636 = lean_ctor_get(x_160, 0); +lean_inc(x_636); +lean_dec(x_160); +x_637 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_637, 0, x_636); +return x_637; } -lean_ctor_set(x_893, 0, x_891); -return x_893; } } else { -lean_object* x_894; lean_object* x_895; lean_object* x_896; -lean_dec_ref(x_728); -lean_dec(x_727); +uint8_t x_638; +lean_dec_ref(x_155); +lean_dec(x_154); +lean_dec(x_19); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_894 = lean_ctor_get(x_731, 0); -lean_inc(x_894); -if (lean_is_exclusive(x_731)) { - lean_ctor_release(x_731, 0); - x_895 = x_731; -} else { - lean_dec_ref(x_731); - x_895 = lean_box(0); +x_638 = !lean_is_exclusive(x_158); +if (x_638 == 0) +{ +return x_158; } -if (lean_is_scalar(x_895)) { - x_896 = lean_alloc_ctor(1, 1, 0); -} else { - x_896 = x_895; +else +{ +lean_object* x_639; lean_object* x_640; +x_639 = lean_ctor_get(x_158, 0); +lean_inc(x_639); +lean_dec(x_158); +x_640 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_640, 0, x_639); +return x_640; } -lean_ctor_set(x_896, 0, x_894); -return x_896; } } default: { -lean_object* x_897; lean_object* x_898; lean_object* x_899; -x_897 = lean_ctor_get(x_634, 0); -lean_inc(x_897); -x_898 = lean_ctor_get(x_634, 1); -lean_inc_ref(x_898); -lean_dec_ref(x_634); -x_899 = l_Lean_IR_ToIR_getFVarValue___redArg(x_897, x_3); -if (lean_obj_tag(x_899) == 0) -{ -lean_object* x_900; -x_900 = lean_ctor_get(x_899, 0); -lean_inc(x_900); -lean_dec_ref(x_899); -if (lean_obj_tag(x_900) == 0) -{ -lean_object* x_901; size_t x_902; size_t x_903; lean_object* x_904; -x_901 = lean_ctor_get(x_900, 0); -lean_inc(x_901); -lean_dec_ref(x_900); -x_902 = lean_array_size(x_898); -x_903 = 0; -x_904 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ToIR_lowerCode_spec__3___redArg(x_902, x_903, x_898, x_3); -if (lean_obj_tag(x_904) == 0) -{ -lean_object* x_905; lean_object* x_906; -x_905 = lean_ctor_get(x_904, 0); -lean_inc(x_905); -lean_dec_ref(x_904); -x_906 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_mkAp(x_1, x_2, x_901, x_905, x_3, x_4, x_5); -return x_906; -} -else -{ -lean_object* x_907; lean_object* x_908; lean_object* x_909; -lean_dec(x_901); +lean_object* x_641; lean_object* x_642; lean_object* x_643; +lean_dec(x_19); +x_641 = lean_ctor_get(x_21, 0); +lean_inc(x_641); +x_642 = lean_ctor_get(x_21, 1); +lean_inc_ref(x_642); +lean_dec_ref(x_21); +x_643 = l_Lean_IR_ToIR_getFVarValue___redArg(x_641, x_3); +if (lean_obj_tag(x_643) == 0) +{ +lean_object* x_644; +x_644 = lean_ctor_get(x_643, 0); +lean_inc(x_644); +lean_dec_ref(x_643); +if (lean_obj_tag(x_644) == 0) +{ +lean_object* x_645; size_t x_646; size_t x_647; lean_object* x_648; +x_645 = lean_ctor_get(x_644, 0); +lean_inc(x_645); +lean_dec_ref(x_644); +x_646 = lean_array_size(x_642); +x_647 = 0; +x_648 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ToIR_lowerCode_spec__3___redArg(x_646, x_647, x_642, x_3); +if (lean_obj_tag(x_648) == 0) +{ +lean_object* x_649; lean_object* x_650; +x_649 = lean_ctor_get(x_648, 0); +lean_inc(x_649); +lean_dec_ref(x_648); +x_650 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_mkAp(x_1, x_2, x_645, x_649, x_3, x_4, x_5); +return x_650; +} +else +{ +uint8_t x_651; +lean_dec(x_645); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_907 = lean_ctor_get(x_904, 0); -lean_inc(x_907); -if (lean_is_exclusive(x_904)) { - lean_ctor_release(x_904, 0); - x_908 = x_904; -} else { - lean_dec_ref(x_904); - x_908 = lean_box(0); +x_651 = !lean_is_exclusive(x_648); +if (x_651 == 0) +{ +return x_648; } -if (lean_is_scalar(x_908)) { - x_909 = lean_alloc_ctor(1, 1, 0); -} else { - x_909 = x_908; +else +{ +lean_object* x_652; lean_object* x_653; +x_652 = lean_ctor_get(x_648, 0); +lean_inc(x_652); +lean_dec(x_648); +x_653 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_653, 0, x_652); +return x_653; } -lean_ctor_set(x_909, 0, x_907); -return x_909; } } else { -lean_object* x_910; -lean_dec_ref(x_898); -x_910 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_mkErased___redArg(x_1, x_2, x_3, x_4, x_5); -return x_910; +lean_object* x_654; +lean_dec_ref(x_642); +x_654 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_mkErased___redArg(x_1, x_2, x_3, x_4, x_5); +return x_654; } } else { -lean_object* x_911; lean_object* x_912; lean_object* x_913; -lean_dec_ref(x_898); +uint8_t x_655; +lean_dec_ref(x_642); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_911 = lean_ctor_get(x_899, 0); -lean_inc(x_911); -if (lean_is_exclusive(x_899)) { - lean_ctor_release(x_899, 0); - x_912 = x_899; -} else { - lean_dec_ref(x_899); - x_912 = lean_box(0); -} -if (lean_is_scalar(x_912)) { - x_913 = lean_alloc_ctor(1, 1, 0); -} else { - x_913 = x_912; +x_655 = !lean_is_exclusive(x_643); +if (x_655 == 0) +{ +return x_643; } -lean_ctor_set(x_913, 0, x_911); -return x_913; +else +{ +lean_object* x_656; lean_object* x_657; +x_656 = lean_ctor_get(x_643, 0); +lean_inc(x_656); +lean_dec(x_643); +x_657 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_657, 0, x_656); +return x_657; } } } @@ -9977,7 +8558,7 @@ static lean_object* _init_l_Lean_IR_ToIR_lowerCode___closed__2() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_IR_ToIR_lowerCode___closed__1; x_2 = lean_unsigned_to_nat(15u); -x_3 = lean_unsigned_to_nat(166u); +x_3 = lean_unsigned_to_nat(168u); x_4 = l_Lean_IR_ToIR_lowerCode___closed__0; x_5 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerAlt_loop___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -10059,7 +8640,7 @@ static lean_object* _init_l_Lean_IR_ToIR_lowerCode___closed__4() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_IR_ToIR_lowerCode___closed__3; x_2 = lean_unsigned_to_nat(6u); -x_3 = lean_unsigned_to_nat(145u); +x_3 = lean_unsigned_to_nat(147u); x_4 = l_Lean_IR_ToIR_lowerCode___closed__0; x_5 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerAlt_loop___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -10088,7 +8669,7 @@ static lean_object* _init_l_Lean_IR_ToIR_lowerCode___closed__7() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_IR_ToIR_lowerCode___closed__6; x_2 = lean_unsigned_to_nat(6u); -x_3 = lean_unsigned_to_nat(147u); +x_3 = lean_unsigned_to_nat(149u); x_4 = l_Lean_IR_ToIR_lowerCode___closed__0; x_5 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerAlt_loop___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -10109,7 +8690,7 @@ static lean_object* _init_l_Lean_IR_ToIR_lowerCode___closed__9() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_IR_ToIR_lowerCode___closed__8; x_2 = lean_unsigned_to_nat(6u); -x_3 = lean_unsigned_to_nat(148u); +x_3 = lean_unsigned_to_nat(150u); x_4 = l_Lean_IR_ToIR_lowerCode___closed__0; x_5 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerAlt_loop___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -10130,7 +8711,7 @@ static lean_object* _init_l_Lean_IR_ToIR_lowerCode___closed__11() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_IR_ToIR_lowerCode___closed__10; x_2 = lean_unsigned_to_nat(49u); -x_3 = lean_unsigned_to_nat(146u); +x_3 = lean_unsigned_to_nat(148u); x_4 = l_Lean_IR_ToIR_lowerCode___closed__0; x_5 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerAlt_loop___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -10159,7 +8740,7 @@ static lean_object* _init_l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowe 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_IR_ToIR_0__Lean_IR_ToIR_lowerAlt_loop___closed__2; x_2 = lean_unsigned_to_nat(18u); -x_3 = lean_unsigned_to_nat(324u); +x_3 = lean_unsigned_to_nat(331u); x_4 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerAlt_loop___closed__1; x_5 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerAlt_loop___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -10755,7 +9336,7 @@ static lean_object* _init_l_Lean_IR_ToIR_lowerCode___closed__12() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_IR_ToIR_lowerCode___closed__10; x_2 = lean_unsigned_to_nat(55u); -x_3 = lean_unsigned_to_nat(158u); +x_3 = lean_unsigned_to_nat(160u); x_4 = l_Lean_IR_ToIR_lowerCode___closed__0; x_5 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerAlt_loop___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -11913,6 +10494,7 @@ if (x_12 == 0) lean_object* x_13; lean_dec(x_6); lean_dec(x_5); +lean_dec_ref(x_2); x_13 = l_Lean_IR_ToIR_lowerCode(x_1, x_7, x_8, x_9); return x_13; } @@ -11944,6 +10526,7 @@ x_22 = lean_unsigned_to_nat(1u); x_23 = lean_nat_add(x_6, x_22); lean_dec(x_6); lean_inc(x_5); +lean_inc_ref(x_2); x_24 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_lowerNonObjectFields_loop(x_1, x_2, x_3, x_4, x_5, x_23, x_7, x_8, x_9); if (lean_obj_tag(x_24) == 0) { @@ -11951,124 +10534,193 @@ uint8_t x_25; x_25 = !lean_is_exclusive(x_24); if (x_25 == 0) { -lean_object* x_26; lean_object* x_27; -x_26 = lean_ctor_get(x_24, 0); +uint8_t x_26; +x_26 = !lean_is_exclusive(x_2); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_27 = lean_ctor_get(x_24, 0); +x_28 = lean_ctor_get(x_2, 4); +lean_dec(x_28); +x_29 = lean_ctor_get(x_2, 3); +lean_dec(x_29); +x_30 = lean_ctor_get(x_2, 2); +lean_dec(x_30); +x_31 = lean_ctor_get(x_2, 0); +lean_dec(x_31); lean_inc(x_15); lean_inc(x_21); -x_27 = lean_alloc_ctor(4, 4, 0); -lean_ctor_set(x_27, 0, x_5); -lean_ctor_set(x_27, 1, x_21); -lean_ctor_set(x_27, 2, x_15); -lean_ctor_set(x_27, 3, x_26); -lean_ctor_set(x_24, 0, x_27); +lean_ctor_set_tag(x_2, 4); +lean_ctor_set(x_2, 4, x_27); +lean_ctor_set(x_2, 3, x_15); +lean_ctor_set(x_2, 2, x_21); +lean_ctor_set(x_2, 0, x_5); +lean_ctor_set(x_24, 0, x_2); return x_24; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_24, 0); -lean_inc(x_28); -lean_dec(x_24); +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_24, 0); +x_33 = lean_ctor_get(x_2, 1); +lean_inc(x_33); +lean_dec(x_2); lean_inc(x_15); lean_inc(x_21); -x_29 = lean_alloc_ctor(4, 4, 0); -lean_ctor_set(x_29, 0, x_5); -lean_ctor_set(x_29, 1, x_21); -lean_ctor_set(x_29, 2, x_15); -lean_ctor_set(x_29, 3, x_28); -x_30 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_30, 0, x_29); -return x_30; +x_34 = lean_alloc_ctor(4, 5, 0); +lean_ctor_set(x_34, 0, x_5); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_21); +lean_ctor_set(x_34, 3, x_15); +lean_ctor_set(x_34, 4, x_32); +lean_ctor_set(x_24, 0, x_34); +return x_24; +} +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_35 = lean_ctor_get(x_24, 0); +lean_inc(x_35); +lean_dec(x_24); +x_36 = lean_ctor_get(x_2, 1); +lean_inc(x_36); +if (lean_is_exclusive(x_2)) { + lean_ctor_release(x_2, 0); + lean_ctor_release(x_2, 1); + lean_ctor_release(x_2, 2); + lean_ctor_release(x_2, 3); + lean_ctor_release(x_2, 4); + x_37 = x_2; +} else { + lean_dec_ref(x_2); + x_37 = lean_box(0); +} +lean_inc(x_15); +lean_inc(x_21); +if (lean_is_scalar(x_37)) { + x_38 = lean_alloc_ctor(4, 5, 0); +} else { + x_38 = x_37; + lean_ctor_set_tag(x_38, 4); +} +lean_ctor_set(x_38, 0, x_5); +lean_ctor_set(x_38, 1, x_36); +lean_ctor_set(x_38, 2, x_21); +lean_ctor_set(x_38, 3, x_15); +lean_ctor_set(x_38, 4, x_35); +x_39 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_39, 0, x_38); +return x_39; } } else { lean_dec(x_5); +lean_dec_ref(x_2); return x_24; } } case 3: { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_31 = lean_ctor_get(x_17, 1); -x_32 = lean_ctor_get(x_17, 2); -x_33 = lean_unsigned_to_nat(1u); -x_34 = lean_nat_add(x_6, x_33); +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_40 = lean_ctor_get(x_17, 1); +x_41 = lean_ctor_get(x_17, 2); +x_42 = lean_unsigned_to_nat(1u); +x_43 = lean_nat_add(x_6, x_42); lean_dec(x_6); lean_inc(x_5); -x_35 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_lowerNonObjectFields_loop(x_1, x_2, x_3, x_4, x_5, x_34, x_7, x_8, x_9); -if (lean_obj_tag(x_35) == 0) +lean_inc_ref(x_2); +x_44 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_lowerNonObjectFields_loop(x_1, x_2, x_3, x_4, x_5, x_43, x_7, x_8, x_9); +if (lean_obj_tag(x_44) == 0) { -uint8_t x_36; -x_36 = !lean_is_exclusive(x_35); -if (x_36 == 0) +uint8_t x_45; +x_45 = !lean_is_exclusive(x_44); +if (x_45 == 0) { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_37 = lean_ctor_get(x_35, 0); -x_38 = lean_ctor_get(x_2, 2); -x_39 = lean_ctor_get(x_2, 3); -x_40 = lean_nat_add(x_38, x_39); -lean_inc(x_32); +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_46 = lean_ctor_get(x_44, 0); +x_47 = lean_ctor_get(x_2, 1); +lean_inc(x_47); +x_48 = lean_ctor_get(x_2, 2); +lean_inc(x_48); +x_49 = lean_ctor_get(x_2, 3); +lean_inc(x_49); +lean_dec_ref(x_2); +x_50 = lean_nat_add(x_48, x_49); +lean_dec(x_49); +lean_dec(x_48); +lean_inc(x_41); lean_inc(x_15); -lean_inc(x_31); -x_41 = lean_alloc_ctor(5, 6, 0); -lean_ctor_set(x_41, 0, x_5); -lean_ctor_set(x_41, 1, x_40); -lean_ctor_set(x_41, 2, x_31); -lean_ctor_set(x_41, 3, x_15); -lean_ctor_set(x_41, 4, x_32); -lean_ctor_set(x_41, 5, x_37); -lean_ctor_set(x_35, 0, x_41); -return x_35; +lean_inc(x_40); +x_51 = lean_alloc_ctor(5, 7, 0); +lean_ctor_set(x_51, 0, x_5); +lean_ctor_set(x_51, 1, x_47); +lean_ctor_set(x_51, 2, x_50); +lean_ctor_set(x_51, 3, x_40); +lean_ctor_set(x_51, 4, x_15); +lean_ctor_set(x_51, 5, x_41); +lean_ctor_set(x_51, 6, x_46); +lean_ctor_set(x_44, 0, x_51); +return x_44; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_42 = lean_ctor_get(x_35, 0); -lean_inc(x_42); -lean_dec(x_35); -x_43 = lean_ctor_get(x_2, 2); -x_44 = lean_ctor_get(x_2, 3); -x_45 = lean_nat_add(x_43, x_44); -lean_inc(x_32); +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_52 = lean_ctor_get(x_44, 0); +lean_inc(x_52); +lean_dec(x_44); +x_53 = lean_ctor_get(x_2, 1); +lean_inc(x_53); +x_54 = lean_ctor_get(x_2, 2); +lean_inc(x_54); +x_55 = lean_ctor_get(x_2, 3); +lean_inc(x_55); +lean_dec_ref(x_2); +x_56 = lean_nat_add(x_54, x_55); +lean_dec(x_55); +lean_dec(x_54); +lean_inc(x_41); lean_inc(x_15); -lean_inc(x_31); -x_46 = lean_alloc_ctor(5, 6, 0); -lean_ctor_set(x_46, 0, x_5); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_31); -lean_ctor_set(x_46, 3, x_15); -lean_ctor_set(x_46, 4, x_32); -lean_ctor_set(x_46, 5, x_42); -x_47 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_47, 0, x_46); -return x_47; +lean_inc(x_40); +x_57 = lean_alloc_ctor(5, 7, 0); +lean_ctor_set(x_57, 0, x_5); +lean_ctor_set(x_57, 1, x_53); +lean_ctor_set(x_57, 2, x_56); +lean_ctor_set(x_57, 3, x_40); +lean_ctor_set(x_57, 4, x_15); +lean_ctor_set(x_57, 5, x_41); +lean_ctor_set(x_57, 6, x_52); +x_58 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_58, 0, x_57); +return x_58; } } else { lean_dec(x_5); -return x_35; +lean_dec_ref(x_2); +return x_44; } } default: { -lean_object* x_48; lean_object* x_49; -x_48 = lean_unsigned_to_nat(1u); -x_49 = lean_nat_add(x_6, x_48); +lean_object* x_59; lean_object* x_60; +x_59 = lean_unsigned_to_nat(1u); +x_60 = lean_nat_add(x_6, x_59); lean_dec(x_6); -x_6 = x_49; +x_6 = x_60; goto _start; } } } else { -lean_object* x_51; lean_object* x_52; -x_51 = lean_unsigned_to_nat(1u); -x_52 = lean_nat_add(x_6, x_51); +lean_object* x_62; lean_object* x_63; +x_62 = lean_unsigned_to_nat(1u); +x_63 = lean_nat_add(x_6, x_62); lean_dec(x_6); -x_6 = x_52; +x_6 = x_63; goto _start; } } @@ -12171,7 +10823,6 @@ lean_object* x_11; x_11 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_lowerNonObjectFields(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec_ref(x_4); lean_dec_ref(x_3); -lean_dec_ref(x_2); return x_11; } } @@ -12190,7 +10841,6 @@ lean_object* x_10; x_10 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_lowerNonObjectFields___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec_ref(x_4); lean_dec_ref(x_3); -lean_dec_ref(x_2); return x_10; } } @@ -12357,7 +11007,6 @@ lean_object* x_11; x_11 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_lowerNonObjectFields_loop(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec_ref(x_4); lean_dec_ref(x_3); -lean_dec_ref(x_2); return x_11; } } @@ -12440,7 +11089,7 @@ static lean_object* _init_l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowe 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_IR_ToIR_0__Lean_IR_ToIR_lowerResultType_resultTypeForArity___closed__1; x_2 = lean_unsigned_to_nat(11u); -x_3 = lean_unsigned_to_nat(341u); +x_3 = lean_unsigned_to_nat(348u); x_4 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerResultType_resultTypeForArity___closed__0; x_5 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerAlt_loop___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -12676,320 +11325,324 @@ return x_2; LEAN_EXPORT lean_object* l_Lean_IR_ToIR_lowerDecl(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; size_t x_10; size_t x_11; lean_object* x_12; -x_6 = lean_ctor_get(x_1, 0); -lean_inc(x_6); -x_7 = lean_ctor_get(x_1, 2); -lean_inc_ref(x_7); -x_8 = lean_ctor_get(x_1, 3); -lean_inc_ref(x_8); -x_9 = lean_ctor_get(x_1, 4); -lean_inc_ref(x_9); +lean_object* x_6; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; size_t x_21; size_t x_22; lean_object* x_23; +x_10 = lean_ctor_get(x_1, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_1, 2); +lean_inc_ref(x_11); +x_12 = lean_ctor_get(x_1, 3); +lean_inc_ref(x_12); +x_13 = lean_ctor_get(x_1, 4); +lean_inc_ref(x_13); lean_dec_ref(x_1); -x_10 = lean_array_size(x_8); -x_11 = 0; +x_21 = lean_array_size(x_12); +x_22 = 0; lean_inc(x_4); lean_inc_ref(x_3); -lean_inc_ref(x_8); -x_12 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ToIR_lowerCode_spec__2(x_10, x_11, x_8, x_2, x_3, x_4); -if (lean_obj_tag(x_12) == 0) +lean_inc_ref(x_12); +x_23 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ToIR_lowerCode_spec__2(x_21, x_22, x_12, x_2, x_3, x_4); +if (lean_obj_tag(x_23) == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec_ref(x_23); +x_25 = lean_array_get_size(x_12); lean_dec_ref(x_12); -x_14 = lean_array_get_size(x_8); -lean_dec_ref(x_8); lean_inc(x_4); lean_inc_ref(x_3); -x_15 = l_Lean_IR_ToIR_lowerResultType___redArg(x_7, x_14, x_3, x_4); -lean_dec_ref(x_7); -if (lean_obj_tag(x_15) == 0) +x_26 = l_Lean_IR_ToIR_lowerResultType___redArg(x_11, x_25, x_3, x_4); +lean_dec_ref(x_11); +if (lean_obj_tag(x_26) == 0) { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -if (lean_is_exclusive(x_15)) { - lean_ctor_release(x_15, 0); - x_17 = x_15; +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +if (lean_is_exclusive(x_26)) { + lean_ctor_release(x_26, 0); + x_28 = x_26; } else { - lean_dec_ref(x_15); - x_17 = lean_box(0); + lean_dec_ref(x_26); + x_28 = lean_box(0); } -x_18 = lean_st_ref_get(x_4); -x_19 = lean_ctor_get(x_18, 0); -lean_inc_ref(x_19); -lean_dec(x_18); -x_20 = l_Lean_IR_ToIR_lowerDecl___closed__0; -lean_inc(x_6); -x_21 = l_Lean_TagAttribute_hasTag(x_20, x_19, x_6); -if (lean_obj_tag(x_9) == 0) +x_29 = lean_st_ref_get(x_4); +x_30 = lean_ctor_get(x_29, 0); +lean_inc_ref(x_30); +lean_dec(x_29); +x_31 = l_Lean_IR_ToIR_lowerDecl___closed__0; +lean_inc(x_10); +x_32 = l_Lean_TagAttribute_hasTag(x_31, x_30, x_10); +if (lean_obj_tag(x_13) == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -lean_dec(x_17); -x_22 = lean_ctor_get(x_9, 0); -lean_inc_ref(x_22); -if (lean_is_exclusive(x_9)) { - lean_ctor_release(x_9, 0); - x_23 = x_9; +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +lean_dec(x_28); +x_33 = lean_ctor_get(x_13, 0); +lean_inc_ref(x_33); +if (lean_is_exclusive(x_13)) { + lean_ctor_release(x_13, 0); + x_34 = x_13; } else { - lean_dec_ref(x_9); - x_23 = lean_box(0); + lean_dec_ref(x_13); + x_34 = lean_box(0); } -if (x_21 == 0) +if (x_32 == 0) { -x_24 = x_2; -x_25 = x_3; -x_26 = x_4; -x_27 = lean_box(0); -goto block_42; +x_35 = x_2; +x_36 = x_3; +x_37 = x_4; +x_38 = lean_box(0); +goto block_53; } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; -lean_dec(x_23); -lean_dec_ref(x_22); -lean_dec(x_16); -lean_dec(x_13); +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; +lean_dec(x_34); +lean_dec_ref(x_33); +lean_dec(x_27); +lean_dec(x_24); lean_dec(x_2); -x_43 = l_Lean_IR_ToIR_lowerDecl___closed__2; -x_44 = l_Lean_MessageData_ofName(x_6); -x_45 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_45, 0, x_43); -lean_ctor_set(x_45, 1, x_44); -x_46 = l_Lean_IR_ToIR_lowerDecl___closed__4; -x_47 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_46); -x_48 = l_Lean_throwError___at___00Lean_IR_ToIR_lowerLet_spec__10___redArg(x_47, x_3, x_4); +x_54 = l_Lean_IR_ToIR_lowerDecl___closed__2; +x_55 = l_Lean_MessageData_ofName(x_10); +x_56 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_55); +x_57 = l_Lean_IR_ToIR_lowerDecl___closed__4; +x_58 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_58, 0, x_56); +lean_ctor_set(x_58, 1, x_57); +x_59 = l_Lean_throwError___at___00Lean_IR_ToIR_lowerLet_spec__10___redArg(x_58, x_3, x_4); lean_dec(x_4); lean_dec_ref(x_3); -x_49 = !lean_is_exclusive(x_48); -if (x_49 == 0) +x_60 = !lean_is_exclusive(x_59); +if (x_60 == 0) { -return x_48; +return x_59; } else { -lean_object* x_50; lean_object* x_51; -x_50 = lean_ctor_get(x_48, 0); -lean_inc(x_50); -lean_dec(x_48); -x_51 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_51, 0, x_50); -return x_51; +lean_object* x_61; lean_object* x_62; +x_61 = lean_ctor_get(x_59, 0); +lean_inc(x_61); +lean_dec(x_59); +x_62 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_62, 0, x_61); +return x_62; } } -block_42: +block_53: { -lean_object* x_28; -x_28 = l_Lean_IR_ToIR_lowerCode(x_22, x_24, x_25, x_26); -if (lean_obj_tag(x_28) == 0) +lean_object* x_39; +x_39 = l_Lean_IR_ToIR_lowerCode(x_33, x_35, x_36, x_37); +if (lean_obj_tag(x_39) == 0) { -uint8_t x_29; -x_29 = !lean_is_exclusive(x_28); -if (x_29 == 0) -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_30 = lean_ctor_get(x_28, 0); -x_31 = lean_box(0); -x_32 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_32, 0, x_6); -lean_ctor_set(x_32, 1, x_13); -lean_ctor_set(x_32, 2, x_16); -lean_ctor_set(x_32, 3, x_30); -lean_ctor_set(x_32, 4, x_31); -if (lean_is_scalar(x_23)) { - x_33 = lean_alloc_ctor(1, 1, 0); -} else { - x_33 = x_23; - lean_ctor_set_tag(x_33, 1); -} -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_28, 0, x_33); -return x_28; +uint8_t x_40; +x_40 = !lean_is_exclusive(x_39); +if (x_40 == 0) +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_41 = lean_ctor_get(x_39, 0); +x_42 = lean_box(0); +x_43 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_43, 0, x_10); +lean_ctor_set(x_43, 1, x_24); +lean_ctor_set(x_43, 2, x_27); +lean_ctor_set(x_43, 3, x_41); +lean_ctor_set(x_43, 4, x_42); +if (lean_is_scalar(x_34)) { + x_44 = lean_alloc_ctor(1, 1, 0); +} else { + x_44 = x_34; + lean_ctor_set_tag(x_44, 1); +} +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_39, 0, x_44); +return x_39; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_34 = lean_ctor_get(x_28, 0); -lean_inc(x_34); -lean_dec(x_28); -x_35 = lean_box(0); -x_36 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_36, 0, x_6); -lean_ctor_set(x_36, 1, x_13); -lean_ctor_set(x_36, 2, x_16); -lean_ctor_set(x_36, 3, x_34); -lean_ctor_set(x_36, 4, x_35); -if (lean_is_scalar(x_23)) { - x_37 = lean_alloc_ctor(1, 1, 0); -} else { - x_37 = x_23; - lean_ctor_set_tag(x_37, 1); +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_45 = lean_ctor_get(x_39, 0); +lean_inc(x_45); +lean_dec(x_39); +x_46 = lean_box(0); +x_47 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_47, 0, x_10); +lean_ctor_set(x_47, 1, x_24); +lean_ctor_set(x_47, 2, x_27); +lean_ctor_set(x_47, 3, x_45); +lean_ctor_set(x_47, 4, x_46); +if (lean_is_scalar(x_34)) { + x_48 = lean_alloc_ctor(1, 1, 0); +} else { + x_48 = x_34; + lean_ctor_set_tag(x_48, 1); } -lean_ctor_set(x_37, 0, x_36); -x_38 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_38, 0, x_37); -return x_38; +lean_ctor_set(x_48, 0, x_47); +x_49 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_49, 0, x_48); +return x_49; } } else { -uint8_t x_39; -lean_dec(x_23); -lean_dec(x_16); -lean_dec(x_13); -lean_dec(x_6); -x_39 = !lean_is_exclusive(x_28); -if (x_39 == 0) +uint8_t x_50; +lean_dec(x_34); +lean_dec(x_27); +lean_dec(x_24); +lean_dec(x_10); +x_50 = !lean_is_exclusive(x_39); +if (x_50 == 0) { -return x_28; +return x_39; } else { -lean_object* x_40; lean_object* x_41; -x_40 = lean_ctor_get(x_28, 0); -lean_inc(x_40); -lean_dec(x_28); -x_41 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_41, 0, x_40); -return x_41; +lean_object* x_51; lean_object* x_52; +x_51 = lean_ctor_get(x_39, 0); +lean_inc(x_51); +lean_dec(x_39); +x_52 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_52, 0, x_51); +return x_52; } } } } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_dec(x_2); -x_52 = lean_ctor_get(x_9, 0); -lean_inc(x_52); -if (lean_is_exclusive(x_9)) { - lean_ctor_release(x_9, 0); - x_53 = x_9; +x_63 = lean_ctor_get(x_13, 0); +lean_inc(x_63); +if (lean_is_exclusive(x_13)) { + lean_ctor_release(x_13, 0); + x_64 = x_13; } else { - lean_dec_ref(x_9); - x_53 = lean_box(0); + lean_dec_ref(x_13); + x_64 = lean_box(0); } -if (x_21 == 0) +if (x_32 == 0) { lean_dec_ref(x_3); -x_54 = x_16; -x_55 = x_4; -x_56 = lean_box(0); -goto block_68; +x_65 = x_27; +x_66 = x_4; +x_67 = lean_box(0); +goto block_76; } else { -uint8_t x_69; -x_69 = l_Lean_IR_IRType_isScalar(x_16); -if (x_69 == 0) +uint8_t x_77; +x_77 = l_Lean_IR_IRType_isScalar(x_27); +if (x_77 == 0) { -lean_object* x_70; -lean_dec(x_16); +lean_object* x_78; +lean_dec(x_27); lean_dec_ref(x_3); -x_70 = lean_box(12); -x_54 = x_70; -x_55 = x_4; -x_56 = lean_box(0); -goto block_68; +x_78 = lean_box(12); +x_65 = x_78; +x_66 = x_4; +x_67 = lean_box(0); +goto block_76; } else { -lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; uint8_t x_80; -lean_dec(x_53); -lean_dec(x_52); -lean_dec(x_17); -lean_dec(x_13); -x_71 = l_Lean_IR_ToIR_lowerDecl___closed__6; -x_72 = l_Lean_MessageData_ofName(x_6); -x_73 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_73, 0, x_71); -lean_ctor_set(x_73, 1, x_72); -x_74 = l_Lean_IR_ToIR_lowerDecl___closed__8; -x_75 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_75, 0, x_73); -lean_ctor_set(x_75, 1, x_74); -x_76 = l_Lean_IR_instToFormatIRType___private__1(x_16); -x_77 = l_Lean_MessageData_ofFormat(x_76); -x_78 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_78, 0, x_75); -lean_ctor_set(x_78, 1, x_77); -x_79 = l_Lean_throwError___at___00Lean_IR_ToIR_lowerLet_spec__10___redArg(x_78, x_3, x_4); +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; uint8_t x_88; +lean_dec(x_64); +lean_dec(x_63); +lean_dec(x_28); +lean_dec(x_24); +x_79 = l_Lean_IR_ToIR_lowerDecl___closed__6; +x_80 = l_Lean_MessageData_ofName(x_10); +x_81 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_81, 0, x_79); +lean_ctor_set(x_81, 1, x_80); +x_82 = l_Lean_IR_ToIR_lowerDecl___closed__8; +x_83 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_83, 0, x_81); +lean_ctor_set(x_83, 1, x_82); +x_84 = l_Lean_IR_instToFormatIRType___private__1(x_27); +x_85 = l_Lean_MessageData_ofFormat(x_84); +x_86 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_86, 0, x_83); +lean_ctor_set(x_86, 1, x_85); +x_87 = l_Lean_throwError___at___00Lean_IR_ToIR_lowerLet_spec__10___redArg(x_86, x_3, x_4); lean_dec(x_4); lean_dec_ref(x_3); -x_80 = !lean_is_exclusive(x_79); -if (x_80 == 0) +x_88 = !lean_is_exclusive(x_87); +if (x_88 == 0) { -return x_79; +return x_87; } else { -lean_object* x_81; lean_object* x_82; -x_81 = lean_ctor_get(x_79, 0); -lean_inc(x_81); -lean_dec(x_79); -x_82 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_82, 0, x_81); -return x_82; +lean_object* x_89; lean_object* x_90; +x_89 = lean_ctor_get(x_87, 0); +lean_inc(x_89); +lean_dec(x_87); +x_90 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_90, 0, x_89); +return x_90; } } } -block_68: +block_76: { -uint8_t x_57; -x_57 = l_List_isEmpty___redArg(x_52); -if (x_57 == 0) +uint8_t x_68; +x_68 = l_List_isEmpty___redArg(x_63); +if (x_68 == 0) { -lean_object* x_58; lean_object* x_59; lean_object* x_60; -lean_dec(x_55); -x_58 = lean_alloc_ctor(1, 4, 0); -lean_ctor_set(x_58, 0, x_6); -lean_ctor_set(x_58, 1, x_13); -lean_ctor_set(x_58, 2, x_54); -lean_ctor_set(x_58, 3, x_52); -if (lean_is_scalar(x_53)) { - x_59 = lean_alloc_ctor(1, 1, 0); +lean_object* x_69; lean_object* x_70; lean_object* x_71; +lean_dec(x_66); +x_69 = lean_alloc_ctor(1, 4, 0); +lean_ctor_set(x_69, 0, x_10); +lean_ctor_set(x_69, 1, x_24); +lean_ctor_set(x_69, 2, x_65); +lean_ctor_set(x_69, 3, x_63); +if (lean_is_scalar(x_64)) { + x_70 = lean_alloc_ctor(1, 1, 0); } else { - x_59 = x_53; + x_70 = x_64; } -lean_ctor_set(x_59, 0, x_58); -if (lean_is_scalar(x_17)) { - x_60 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_70, 0, x_69); +if (lean_is_scalar(x_28)) { + x_71 = lean_alloc_ctor(0, 1, 0); } else { - x_60 = x_17; + x_71 = x_28; } -lean_ctor_set(x_60, 0, x_59); -return x_60; +lean_ctor_set(x_71, 0, x_70); +return x_71; } else { -lean_object* x_61; lean_object* x_62; uint8_t x_63; -lean_dec(x_53); -lean_dec(x_52); -lean_dec(x_17); -x_61 = l_Lean_IR_mkDummyExternDecl(x_6, x_13, x_54); -x_62 = l_Lean_IR_ToIR_addDecl___redArg(x_61, x_55); -lean_dec(x_55); -x_63 = !lean_is_exclusive(x_62); -if (x_63 == 0) -{ -lean_object* x_64; lean_object* x_65; -x_64 = lean_ctor_get(x_62, 0); +lean_object* x_72; lean_object* x_73; uint8_t x_74; lean_dec(x_64); -x_65 = lean_box(0); -lean_ctor_set(x_62, 0, x_65); -return x_62; +lean_dec(x_63); +lean_dec(x_28); +lean_inc(x_65); +lean_inc(x_24); +lean_inc(x_10); +x_72 = l_Lean_IR_mkDummyExternDecl(x_10, x_24, x_65); +x_73 = l_Lean_IR_ToIR_addDecl___redArg(x_72, x_66); +lean_dec_ref(x_73); +x_74 = l_Array_isEmpty___redArg(x_24); +lean_dec(x_24); +if (x_74 == 0) +{ +x_14 = x_65; +x_15 = x_66; +x_16 = lean_box(0); +x_17 = x_74; +goto block_20; } else { -lean_object* x_66; lean_object* x_67; -lean_dec(x_62); -x_66 = lean_box(0); -x_67 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_67, 0, x_66); -return x_67; +uint8_t x_75; +x_75 = l_Lean_IR_IRType_isStruct(x_65); +x_14 = x_65; +x_15 = x_66; +x_16 = lean_box(0); +x_17 = x_75; +goto block_20; } } } @@ -12997,54 +11650,83 @@ return x_67; } else { -uint8_t x_83; -lean_dec(x_13); -lean_dec_ref(x_9); -lean_dec(x_6); +uint8_t x_91; +lean_dec(x_24); +lean_dec_ref(x_13); +lean_dec(x_10); lean_dec(x_4); lean_dec_ref(x_3); lean_dec(x_2); -x_83 = !lean_is_exclusive(x_15); -if (x_83 == 0) +x_91 = !lean_is_exclusive(x_26); +if (x_91 == 0) { -return x_15; +return x_26; } else { -lean_object* x_84; lean_object* x_85; -x_84 = lean_ctor_get(x_15, 0); -lean_inc(x_84); -lean_dec(x_15); -x_85 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_85, 0, x_84); -return x_85; +lean_object* x_92; lean_object* x_93; +x_92 = lean_ctor_get(x_26, 0); +lean_inc(x_92); +lean_dec(x_26); +x_93 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_93, 0, x_92); +return x_93; } } } else { -uint8_t x_86; -lean_dec_ref(x_9); -lean_dec_ref(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); +uint8_t x_94; +lean_dec_ref(x_13); +lean_dec_ref(x_12); +lean_dec_ref(x_11); +lean_dec(x_10); lean_dec(x_4); lean_dec_ref(x_3); lean_dec(x_2); -x_86 = !lean_is_exclusive(x_12); -if (x_86 == 0) +x_94 = !lean_is_exclusive(x_23); +if (x_94 == 0) { -return x_12; +return x_23; } else { -lean_object* x_87; lean_object* x_88; -x_87 = lean_ctor_get(x_12, 0); -lean_inc(x_87); -lean_dec(x_12); -x_88 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_88, 0, x_87); -return x_88; +lean_object* x_95; lean_object* x_96; +x_95 = lean_ctor_get(x_23, 0); +lean_inc(x_95); +lean_dec(x_23); +x_96 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_96, 0, x_95); +return x_96; +} +} +block_9: +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_box(0); +x_8 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_8, 0, x_7); +return x_8; +} +block_20: +{ +if (x_17 == 0) +{ +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_10); +x_6 = lean_box(0); +goto block_9; +} +else +{ +lean_object* x_18; lean_object* x_19; +x_18 = l_Lean_IR_ExplicitBoxing_boxedConstDecl(x_10, x_14); +x_19 = l_Lean_IR_ToIR_addDecl___redArg(x_18, x_15); +lean_dec(x_15); +lean_dec_ref(x_19); +x_6 = lean_box(0); +goto block_9; } } } @@ -13181,6 +11863,8 @@ return x_10; } lean_object* initialize_Lean_Compiler_IR_CompilerM(uint8_t builtin); lean_object* initialize_Lean_Compiler_IR_ToIRType(uint8_t builtin); +lean_object* initialize_Lean_Compiler_IR_UnboxResult(uint8_t builtin); +lean_object* initialize_Lean_Compiler_IR_Boxing(uint8_t builtin); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Compiler_IR_ToIR(uint8_t builtin) { lean_object * res; @@ -13192,6 +11876,12 @@ lean_dec_ref(res); res = initialize_Lean_Compiler_IR_ToIRType(builtin); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Compiler_IR_UnboxResult(builtin); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Compiler_IR_Boxing(builtin); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); l_Lean_IR_ToIR_initFn___closed__0_00___x40_Lean_Compiler_IR_ToIR_1721792695____hygCtx___hyg_2_ = _init_l_Lean_IR_ToIR_initFn___closed__0_00___x40_Lean_Compiler_IR_ToIR_1721792695____hygCtx___hyg_2_(); lean_mark_persistent(l_Lean_IR_ToIR_initFn___closed__0_00___x40_Lean_Compiler_IR_ToIR_1721792695____hygCtx___hyg_2_); l_Lean_IR_ToIR_initFn___closed__1_00___x40_Lean_Compiler_IR_ToIR_1721792695____hygCtx___hyg_2_ = _init_l_Lean_IR_ToIR_initFn___closed__1_00___x40_Lean_Compiler_IR_ToIR_1721792695____hygCtx___hyg_2_(); diff --git a/stage0/stdlib/Lean/Compiler/IR/ToIRType.c b/stage0/stdlib/Lean/Compiler/IR/ToIRType.c index e101f346fdd0..748f6789f4f4 100644 --- a/stage0/stdlib/Lean/Compiler/IR/ToIRType.c +++ b/stage0/stdlib/Lean/Compiler/IR/ToIRType.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Compiler.IR.ToIRType -// Imports: public import Lean.Compiler.IR.Format public import Lean.Compiler.LCNF.MonoTypes +// Imports: public import Lean.Compiler.IR.Format public import Lean.Compiler.LCNF.MonoTypes import Lean.Compiler.IR.UnboxResult #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -15,48 +15,67 @@ extern "C" { #endif LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___lam__0(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__11; LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_format(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_nameToIRType___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__7; -LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_toIRType_spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at___00Lean_PersistentHashMap_find_x3f___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__1_spec__3(lean_object*, lean_object*, size_t, lean_object*); static lean_object* l_Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0___redArg___closed__0; lean_object* l_List_lengthTR___redArg(lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00Lean_IR_unboxedIRType_spec__10___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); +static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__8; LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_toIRType_visitApp___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_right(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_scalar_elim(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__10; +LEAN_EXPORT lean_object* l_Lean_IR_unboxedIRType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_getCtorLayout___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_hasTrivialStructure_x3f(lean_object*, lean_object*, lean_object*); -static lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___closed__0; +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__20___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__9; LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__6; +static lean_object* l_Lean_IR_unboxedIRType___closed__3; +static lean_object* l_Lean_getConstInfoInduct___at___00Lean_IR_unboxedIRType_spec__9___closed__1; static lean_object* l_Lean_IR_CtorFieldInfo_format___closed__2; +lean_object* l_Lean_Environment_header(lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00Lean_IR_unboxedIRType_spec__10___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_le(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_IR_instInhabitedCtorFieldInfo; LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_usize_elim(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_hasTrivialStructure_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_getConstInfoInduct___at___00Lean_IR_unboxedIRType_spec__9___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2_(); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__1___redArg(lean_object*, lean_object*); +static lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg___closed__0; size_t lean_uint64_to_usize(uint64_t); static lean_object* l_Lean_IR_toIRType___closed__0; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__13___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Meta_instInhabitedConfigWithKey___private__1; +lean_object* l_Lean_Environment_findConstVal_x3f(lean_object*, lean_object*, uint8_t); static lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__2; lean_object* l_Lean_Expr_sort___override(lean_object*); static lean_object* l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__0___closed__0; lean_object* lean_array_push(lean_object*, lean_object*); static lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__6; size_t lean_usize_mul(size_t, size_t); -static lean_object* l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___closed__0; +static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__5; +static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__13; static lean_object* l_Lean_IR_CtorFieldInfo_format___closed__4; +static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__9; +extern lean_object* l_Lean_unknownIdentifierMessageTag; static lean_object* l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg___closed__1; +uint8_t l_Lean_Name_isAnonymous(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_895574536____hygCtx___hyg_2____boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__4___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at___00List_foldl___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__3_spec__6___redArg(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_replaceRef(lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); lean_object* l_Lean_registerEnvExtension___redArg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -64,19 +83,22 @@ LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCto static lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__3; lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_hasTrivialStructure_x3f___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__8; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__0_spec__1___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_fvarId_x21(lean_object*); static lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__1; lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*, uint8_t); -LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_hasTrivialStructure_x3f___lam__0___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_erased_elim(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget_borrowed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___lam__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAuxAux(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_instInhabitedCtorLayout_default___closed__1; +lean_object* l_Lean_stringToMessageData(lean_object*); +lean_object* l_Lean_MessageData_note(lean_object*); static lean_object* l_Lean_IR_toIRType___closed__4; extern lean_object* l_Lean_IR_instInhabitedCtorInfo_default; uint8_t lean_string_dec_eq(lean_object*, lean_object*); @@ -84,45 +106,55 @@ LEAN_EXPORT lean_object* l_Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_330 LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_toIRType___closed__1; +LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_toIRType_spec__6(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_CtorFieldInfo_format___closed__3; +static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__3; LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_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_List_foldl___at___00List_foldl___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__3_spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__5(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___lam__0(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_isAnyProducingType___boxed(lean_object*); static lean_object* l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg___closed__2; static lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__7; LEAN_EXPORT lean_object* l_Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_895574536____hygCtx___hyg_2_(); static lean_object* l_Lean_IR_CtorFieldInfo_instToFormat___closed__0; +LEAN_EXPORT lean_object* l_Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Nat_reprFast(lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingImp(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg___lam__0(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_unboxedIRType___closed__4; LEAN_EXPORT lean_object* l_panic___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__2(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_toLCNFType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__5; +LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0(lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_scalar_elim___redArg(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_isAnyProducingType___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_void_elim(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__2___redArg(lean_object*, lean_object*); static lean_object* l_Lean_IR_CtorFieldInfo_format___closed__11; static lean_object* l_Lean_IR_hasTrivialStructure_x3f___lam__0___closed__1; static lean_object* l_Lean_IR_CtorFieldInfo_format___closed__9; +static lean_object* l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__0; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__0___redArg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0___redArg___lam__0___closed__0; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0___redArg___boxed(lean_object*, lean_object*); lean_object* l_Array_empty(lean_object*); lean_object* l_Lean_EnvExtension_modifyState___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__6; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_erased_elim___redArg(lean_object*, lean_object*); static lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2___redArg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(lean_object*, uint8_t, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__0_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_instantiateForall(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MessageData_ofFormat(lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_unboxedIRType___closed__0; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at___00Lean_PersistentHashMap_findAux___at___00Lean_PersistentHashMap_find_x3f___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__1_spec__3_spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__5; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__0(lean_object*, size_t, size_t, lean_object*, lean_object*); @@ -133,97 +165,125 @@ LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_ctorElim___boxed(lean_object*, static lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__9; LEAN_EXPORT lean_object* l_Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3390564948____hygCtx___hyg_2_(); lean_object* lean_st_ref_get(lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static size_t l_Lean_PersistentHashMap_findAux___at___00Lean_PersistentHashMap_find_x3f___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__1_spec__3___redArg___closed__0; +static lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg___closed__1; +static lean_object* l_Lean_IR_unboxedIRType___closed__5; lean_object* lean_st_mk_ref(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__4___redArg(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg___closed__2; static lean_object* l_Lean_IR_CtorFieldInfo_format___closed__12; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_getParamTypes(lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ctorLayoutExt; static lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__2; static lean_object* l_Lean_IR_instInhabitedCtorLayout_default___closed__0; +LEAN_EXPORT lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__4; LEAN_EXPORT lean_object* l_Lean_IR_instInhabitedCtorLayout; static lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__0; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at___00Lean_PersistentHashMap_findAux___at___00Lean_PersistentHashMap_find_x3f___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__1_spec__3_spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_getConstInfoInduct___at___00Lean_IR_unboxedIRType_spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedExpr; +LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__4(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_irTypeForEnum___boxed(lean_object*); static lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__8; -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_whnfD(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__10; +static lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___closed__2; LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__3___redArg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__4; LEAN_EXPORT lean_object* l_Lean_IR_toIRType(lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); static lean_object* l_List_foldl___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__3___redArg___closed__0; lean_object* l_Lean_Core_instInhabitedCoreM___lam__0___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__0_spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__0_spec__1_spec__3___redArg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__0_spec__1_spec__4(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__13___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Array_extract___redArg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_CtorFieldInfo_format___closed__1; static lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___lam__0___closed__0; lean_object* l_Lean_Compiler_LCNF_getOtherDeclBaseType(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Subarray_toArray___redArg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__5___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg___closed__3; +static lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___closed__1; lean_object* l___private_Init_Data_List_Impl_0__List_takeTR_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_CtorFieldInfo_format___closed__10; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_hasTrivialStructure_x3f___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_get_borrowed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0___boxed(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isErased(lean_object*); lean_object* lean_usize_to_nat(size_t); -LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_toIRType_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MessageData_ofConstName(lean_object*, uint8_t); +static lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg___closed__2; +LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11___redArg(lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Environment_contains(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___00Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__0_spec__1_spec__3_spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static size_t l_Lean_PersistentHashMap_findAux___at___00Lean_PersistentHashMap_find_x3f___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__1_spec__3___redArg___closed__1; static lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__0; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3(size_t, size_t, lean_object*, lean_object*); +static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__7; +static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__12; static lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__5; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppNumArgs(lean_object*); static lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__1; +LEAN_EXPORT lean_object* l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_getCtorLayout(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_trivialStructureInfoExt; LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_ctorIdx___boxed(lean_object*); static lean_object* l_Lean_IR_CtorFieldInfo_format___closed__7; +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__20___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__0_spec__1_spec__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg___closed__0; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0___redArg___lam__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__5___redArg(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_replaceFVars(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_instInhabited(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache(lean_object*, lean_object*, lean_object*); static lean_object* l_List_foldl___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__3___redArg___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___00Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__0_spec__1_spec__3_spec__6___redArg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_toMonoType(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__1___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__2; static lean_object* l_Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0___redArg___closed__1; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkCollisionNode___redArg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Irrelevant_hasTrivialStructure_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_beq___boxed(lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___closed__3; lean_object* l___private_Lean_Environment_0__Lean_EnvExtension_getStateUnsafe___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00Lean_IR_unboxedIRType_spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19___closed__0; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__0_spec__1___redArg(lean_object*, size_t, size_t, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg___closed__1; LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_instToFormat; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_nameToIRType(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_EnvironmentHeader_moduleNames(lean_object*); static lean_object* l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__0_spec__1___redArg___closed__0; -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(lean_object*, uint8_t, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__20(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__6; LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_ctorIdx(lean_object*); +lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); uint64_t l_Lean_Name_hash___override(lean_object*); -lean_object* l_Lean_Meta_instInhabitedMetaM___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_toSubarray___redArg(lean_object*, lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_object_elim(lean_object*, lean_object*, lean_object*, lean_object*); @@ -233,32 +293,43 @@ LEAN_EXPORT lean_object* l_Lean_IR_irTypeForEnum(lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at___00Lean_PersistentHashMap_find_x3f___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__1_spec__3___redArg___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_getConstInfoInduct___at___00Lean_IR_unboxedIRType_spec__9(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__3; +static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0___redArg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0___redArg___lam__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_irTypeExt; +uint8_t l_Lean_isPrivateName(lean_object*); +LEAN_EXPORT lean_object* l_Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__11; lean_object* l_Lean_IR_instToFormatIRType___private__1(lean_object*); static lean_object* l_Lean_IR_instInhabitedCtorLayout_default___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__1___redArg___boxed(lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); static lean_object* l_Lean_IR_CtorFieldInfo_format___closed__13; -static lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__2; +lean_object* l_Lean_Environment_getModuleIdxFor_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_usize_elim___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__0_spec__1_spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__1___boxed(lean_object*, lean_object*, lean_object*); size_t lean_usize_sub(size_t, size_t); -LEAN_EXPORT lean_object* l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___redArg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_hasTrivialStructure_x3f___lam__0___closed__2; +LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__0_spec__1_spec__4___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11(lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); static lean_object* l_List_foldl___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__3___redArg___closed__2; static lean_object* l_Lean_IR_nameToIRType___closed__0; lean_object* l_mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___closed__0; static lean_object* l_Lean_IR_CtorFieldInfo_format___closed__8; static lean_object* l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg___closed__3; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at___00Lean_PersistentHashMap_findAux___at___00Lean_PersistentHashMap_find_x3f___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__1_spec__3_spec__7___redArg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -266,20 +337,30 @@ lean_object* lean_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__0(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_toIRType_visitApp(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_toIRType_spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00Lean_IR_unboxedIRType_spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*); lean_object* l_Lean_FVarId_getType___redArg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__3; LEAN_EXPORT uint8_t l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_isAnyProducingType(lean_object*); size_t lean_usize_shift_left(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15___redArg(lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__4; +static lean_object* l_Lean_IR_unboxedIRType___closed__1; LEAN_EXPORT lean_object* l_Lean_IR_toIRType___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__1; static lean_object* l_Lean_IR_CtorFieldInfo_format___closed__6; -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__4(size_t, size_t, lean_object*, lean_object*); +static lean_object* l_Lean_IR_unboxedIRType___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__2; lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__13___boxed(lean_object**); uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_getCtorLayout___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); @@ -289,6 +370,8 @@ LEAN_EXPORT lean_object* l_Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_339 lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__0_spec__1_spec__4___redArg(size_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Environment_setExporting(lean_object*, uint8_t); static lean_object* l_Lean_IR_toIRType___closed__2; static lean_object* l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_ctorElim___redArg(lean_object*, lean_object*); @@ -296,22 +379,32 @@ LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_ctorElim(lean_object*, lean_obj LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at___00Lean_PersistentHashMap_find_x3f___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__1_spec__3___redArg(lean_object*, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_unboxedIRType(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_hasTrivialStructure_x3f___closed__0; LEAN_EXPORT lean_object* l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__0(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MessageData_ofName(lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__1___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_toIRType___closed__3; +LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19(lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_IR_UnboxResult_hasUnboxAttr(lean_object*, lean_object*); static lean_object* l_Lean_IR_CtorFieldInfo_format___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_instInhabitedCtorLayout_default; +static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__0; lean_object* l_Lean_Name_hash___override___boxed(lean_object*); static lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__3; +lean_object* l_Lean_isInductiveCore_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at___00Lean_PersistentHashMap_find_x3f___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__1_spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_void_elim___redArg(lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); +LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instInhabitedCtorFieldInfo_default; +static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__4; size_t lean_usize_land(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__20___redArg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_CtorFieldInfo_format___closed__5; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at___00Lean_PersistentHashMap_findAux___at___00Lean_PersistentHashMap_find_x3f___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__1_spec__3_spec__7___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_foldl___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__3___redArg___closed__3; +LEAN_EXPORT lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_irTypeForEnum(lean_object* x_1) { _start: { @@ -1711,6147 +1804,7776 @@ x_2 = l_Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3390564948____hygCtx__ return x_2; } } -static lean_object* _init_l_Lean_IR_hasTrivialStructure_x3f___closed__0() { +LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_ctorIdx(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = l_Lean_IR_trivialStructureInfoExt; -return x_1; +switch (lean_obj_tag(x_1)) { +case 0: +{ +lean_object* x_2; +x_2 = lean_unsigned_to_nat(0u); +return x_2; } +case 1: +{ +lean_object* x_3; +x_3 = lean_unsigned_to_nat(1u); +return x_3; } -static lean_object* _init_l_Lean_IR_hasTrivialStructure_x3f___lam__0___closed__0() { -_start: +case 2: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Subtype", 7, 7); -return x_1; +lean_object* x_4; +x_4 = lean_unsigned_to_nat(2u); +return x_4; } +case 3: +{ +lean_object* x_5; +x_5 = lean_unsigned_to_nat(3u); +return x_5; } -static lean_object* _init_l_Lean_IR_hasTrivialStructure_x3f___lam__0___closed__1() { -_start: +default: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Void", 4, 4); -return x_1; +lean_object* x_6; +x_6 = lean_unsigned_to_nat(4u); +return x_6; } } -static lean_object* _init_l_Lean_IR_hasTrivialStructure_x3f___lam__0___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("nonemptyType", 12, 12); -return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_hasTrivialStructure_x3f___lam__0(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_IR_CtorFieldInfo_ctorIdx___boxed(lean_object* x_1) { _start: { -lean_object* x_7; -lean_inc(x_5); -lean_inc_ref(x_4); -lean_inc(x_3); -lean_inc_ref(x_2); -lean_inc_ref(x_1); -x_7 = l_Lean_Meta_isProp(x_1, x_2, x_3, x_4, x_5); -if (lean_obj_tag(x_7) == 0) -{ -lean_object* x_8; uint8_t x_9; -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_unbox(x_8); -lean_dec(x_8); -if (x_9 == 0) -{ -lean_object* x_10; -lean_dec_ref(x_7); -lean_inc(x_5); -lean_inc_ref(x_4); -lean_inc(x_3); -lean_inc_ref(x_2); -lean_inc_ref(x_1); -x_10 = l_Lean_Meta_isTypeFormerType(x_1, x_2, x_3, x_4, x_5); -if (lean_obj_tag(x_10) == 0) -{ -lean_object* x_11; uint8_t x_12; -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_unbox(x_11); -if (x_12 == 0) -{ -lean_object* x_13; -lean_dec_ref(x_10); -x_13 = l_Lean_Meta_whnfD(x_1, x_2, x_3, x_4, x_5); -if (lean_obj_tag(x_13) == 0) -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -lean_object* x_15; -x_15 = lean_ctor_get(x_13, 0); -if (lean_obj_tag(x_15) == 11) -{ -lean_object* x_16; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -if (lean_obj_tag(x_16) == 1) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -if (lean_obj_tag(x_17) == 0) -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_18 = lean_ctor_get(x_15, 1); -lean_inc(x_18); -x_19 = lean_ctor_get(x_15, 2); -lean_inc_ref(x_19); -lean_dec_ref(x_15); -x_20 = lean_ctor_get(x_16, 1); -lean_inc_ref(x_20); -lean_dec_ref(x_16); -x_21 = l_Lean_IR_hasTrivialStructure_x3f___lam__0___closed__0; -x_22 = lean_string_dec_eq(x_20, x_21); -lean_dec_ref(x_20); -if (x_22 == 0) -{ -lean_dec_ref(x_19); -lean_dec(x_18); -lean_ctor_set(x_13, 0, x_11); -return x_13; +lean_object* x_2; +x_2 = l_Lean_IR_CtorFieldInfo_ctorIdx(x_1); +lean_dec(x_1); +return x_2; } -else -{ -lean_object* x_23; uint8_t x_24; -x_23 = lean_unsigned_to_nat(0u); -x_24 = lean_nat_dec_eq(x_18, x_23); -lean_dec(x_18); -if (x_24 == 0) -{ -lean_dec_ref(x_19); -lean_ctor_set(x_13, 0, x_11); -return x_13; } -else -{ -if (lean_obj_tag(x_19) == 5) -{ -lean_object* x_25; -x_25 = lean_ctor_get(x_19, 0); -lean_inc_ref(x_25); -lean_dec_ref(x_19); -if (lean_obj_tag(x_25) == 4) -{ -lean_object* x_26; -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -if (lean_obj_tag(x_26) == 1) -{ -lean_object* x_27; -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -if (lean_obj_tag(x_27) == 1) -{ -lean_object* x_28; -x_28 = lean_ctor_get(x_27, 0); -if (lean_obj_tag(x_28) == 0) +LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_ctorElim___redArg(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; -x_29 = lean_ctor_get(x_25, 1); -lean_inc(x_29); -lean_dec_ref(x_25); -x_30 = lean_ctor_get(x_26, 1); -lean_inc_ref(x_30); -lean_dec_ref(x_26); -x_31 = lean_ctor_get(x_27, 1); -lean_inc_ref(x_31); -lean_dec_ref(x_27); -x_32 = l_Lean_IR_hasTrivialStructure_x3f___lam__0___closed__1; -x_33 = lean_string_dec_eq(x_31, x_32); -lean_dec_ref(x_31); -if (x_33 == 0) +switch (lean_obj_tag(x_1)) { +case 1: { -lean_dec_ref(x_30); -lean_dec(x_29); -lean_ctor_set(x_13, 0, x_11); -return x_13; +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +lean_dec_ref(x_1); +x_5 = lean_apply_2(x_2, x_3, x_4); +return x_5; } -else -{ -lean_object* x_34; uint8_t x_35; -x_34 = l_Lean_IR_hasTrivialStructure_x3f___lam__0___closed__2; -x_35 = lean_string_dec_eq(x_30, x_34); -lean_dec_ref(x_30); -if (x_35 == 0) +case 2: { -lean_dec(x_29); -lean_ctor_set(x_13, 0, x_11); -return x_13; +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_1, 0); +lean_inc(x_6); +lean_dec_ref(x_1); +x_7 = lean_apply_1(x_2, x_6); +return x_7; } -else -{ -if (lean_obj_tag(x_29) == 0) +case 3: { -lean_object* x_36; -lean_dec(x_11); -x_36 = lean_box(x_35); -lean_ctor_set(x_13, 0, x_36); -return x_13; +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_1, 1); +lean_inc(x_9); +x_10 = lean_ctor_get(x_1, 2); +lean_inc(x_10); +lean_dec_ref(x_1); +x_11 = lean_apply_3(x_2, x_8, x_9, x_10); +return x_11; } -else +default: { -lean_dec(x_29); -lean_ctor_set(x_13, 0, x_11); -return x_13; +lean_dec(x_1); +return x_2; } } } } -else +LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_ctorElim(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_ctor_set(x_13, 0, x_11); -return x_13; +lean_object* x_6; +x_6 = l_Lean_IR_CtorFieldInfo_ctorElim___redArg(x_3, x_5); +return x_6; } } -else +LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_ctorElim___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -lean_dec(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_ctor_set(x_13, 0, x_11); -return x_13; +lean_object* x_6; +x_6 = l_Lean_IR_CtorFieldInfo_ctorElim(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_2); +return x_6; } } -else +LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_erased_elim(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_dec(x_26); -lean_dec_ref(x_25); -lean_ctor_set(x_13, 0, x_11); -return x_13; +lean_object* x_5; +x_5 = l_Lean_IR_CtorFieldInfo_ctorElim___redArg(x_2, x_4); +return x_5; } } -else +LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_erased_elim___redArg(lean_object* x_1, lean_object* x_2) { +_start: { -lean_dec_ref(x_25); -lean_ctor_set(x_13, 0, x_11); -return x_13; +lean_object* x_3; +x_3 = l_Lean_IR_CtorFieldInfo_ctorElim___redArg(x_1, x_2); +return x_3; } } -else +LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_object_elim(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_dec_ref(x_19); -lean_ctor_set(x_13, 0, x_11); -return x_13; -} -} +lean_object* x_5; +x_5 = l_Lean_IR_CtorFieldInfo_ctorElim___redArg(x_2, x_4); +return x_5; } } -else +LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_object_elim___redArg(lean_object* x_1, lean_object* x_2) { +_start: { -lean_dec_ref(x_16); -lean_dec_ref(x_15); -lean_ctor_set(x_13, 0, x_11); -return x_13; +lean_object* x_3; +x_3 = l_Lean_IR_CtorFieldInfo_ctorElim___redArg(x_1, x_2); +return x_3; } } -else +LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_usize_elim(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_dec(x_16); -lean_dec_ref(x_15); -lean_ctor_set(x_13, 0, x_11); -return x_13; +lean_object* x_5; +x_5 = l_Lean_IR_CtorFieldInfo_ctorElim___redArg(x_2, x_4); +return x_5; } } -else +LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_usize_elim___redArg(lean_object* x_1, lean_object* x_2) { +_start: { -lean_dec(x_15); -lean_ctor_set(x_13, 0, x_11); -return x_13; +lean_object* x_3; +x_3 = l_Lean_IR_CtorFieldInfo_ctorElim___redArg(x_1, x_2); +return x_3; } } -else -{ -lean_object* x_37; -x_37 = lean_ctor_get(x_13, 0); -lean_inc(x_37); -lean_dec(x_13); -if (lean_obj_tag(x_37) == 11) -{ -lean_object* x_38; -x_38 = lean_ctor_get(x_37, 0); -lean_inc(x_38); -if (lean_obj_tag(x_38) == 1) -{ -lean_object* x_39; -x_39 = lean_ctor_get(x_38, 0); -if (lean_obj_tag(x_39) == 0) -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; -x_40 = lean_ctor_get(x_37, 1); -lean_inc(x_40); -x_41 = lean_ctor_get(x_37, 2); -lean_inc_ref(x_41); -lean_dec_ref(x_37); -x_42 = lean_ctor_get(x_38, 1); -lean_inc_ref(x_42); -lean_dec_ref(x_38); -x_43 = l_Lean_IR_hasTrivialStructure_x3f___lam__0___closed__0; -x_44 = lean_string_dec_eq(x_42, x_43); -lean_dec_ref(x_42); -if (x_44 == 0) +LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_scalar_elim(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_45; -lean_dec_ref(x_41); -lean_dec(x_40); -x_45 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_45, 0, x_11); -return x_45; +lean_object* x_5; +x_5 = l_Lean_IR_CtorFieldInfo_ctorElim___redArg(x_2, x_4); +return x_5; } -else -{ -lean_object* x_46; uint8_t x_47; -x_46 = lean_unsigned_to_nat(0u); -x_47 = lean_nat_dec_eq(x_40, x_46); -lean_dec(x_40); -if (x_47 == 0) -{ -lean_object* x_48; -lean_dec_ref(x_41); -x_48 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_48, 0, x_11); -return x_48; } -else -{ -if (lean_obj_tag(x_41) == 5) -{ -lean_object* x_49; -x_49 = lean_ctor_get(x_41, 0); -lean_inc_ref(x_49); -lean_dec_ref(x_41); -if (lean_obj_tag(x_49) == 4) -{ -lean_object* x_50; -x_50 = lean_ctor_get(x_49, 0); -lean_inc(x_50); -if (lean_obj_tag(x_50) == 1) -{ -lean_object* x_51; -x_51 = lean_ctor_get(x_50, 0); -lean_inc(x_51); -if (lean_obj_tag(x_51) == 1) -{ -lean_object* x_52; -x_52 = lean_ctor_get(x_51, 0); -if (lean_obj_tag(x_52) == 0) -{ -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; -x_53 = lean_ctor_get(x_49, 1); -lean_inc(x_53); -lean_dec_ref(x_49); -x_54 = lean_ctor_get(x_50, 1); -lean_inc_ref(x_54); -lean_dec_ref(x_50); -x_55 = lean_ctor_get(x_51, 1); -lean_inc_ref(x_55); -lean_dec_ref(x_51); -x_56 = l_Lean_IR_hasTrivialStructure_x3f___lam__0___closed__1; -x_57 = lean_string_dec_eq(x_55, x_56); -lean_dec_ref(x_55); -if (x_57 == 0) +LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_scalar_elim___redArg(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_58; -lean_dec_ref(x_54); -lean_dec(x_53); -x_58 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_58, 0, x_11); -return x_58; +lean_object* x_3; +x_3 = l_Lean_IR_CtorFieldInfo_ctorElim___redArg(x_1, x_2); +return x_3; } -else -{ -lean_object* x_59; uint8_t x_60; -x_59 = l_Lean_IR_hasTrivialStructure_x3f___lam__0___closed__2; -x_60 = lean_string_dec_eq(x_54, x_59); -lean_dec_ref(x_54); -if (x_60 == 0) -{ -lean_object* x_61; -lean_dec(x_53); -x_61 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_61, 0, x_11); -return x_61; } -else -{ -if (lean_obj_tag(x_53) == 0) +LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_void_elim(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_62; lean_object* x_63; -lean_dec(x_11); -x_62 = lean_box(x_60); -x_63 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_63, 0, x_62); -return x_63; +lean_object* x_5; +x_5 = l_Lean_IR_CtorFieldInfo_ctorElim___redArg(x_2, x_4); +return x_5; } -else +} +LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_void_elim___redArg(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_64; -lean_dec(x_53); -x_64 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_64, 0, x_11); -return x_64; +lean_object* x_3; +x_3 = l_Lean_IR_CtorFieldInfo_ctorElim___redArg(x_1, x_2); +return x_3; } } +static lean_object* _init_l_Lean_IR_instInhabitedCtorFieldInfo_default() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; } } -else +static lean_object* _init_l_Lean_IR_instInhabitedCtorFieldInfo() { +_start: { -lean_object* x_65; -lean_dec_ref(x_51); -lean_dec_ref(x_50); -lean_dec_ref(x_49); -x_65 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_65, 0, x_11); -return x_65; +lean_object* x_1; +x_1 = lean_box(0); +return x_1; } } -else +static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__0() { +_start: { -lean_object* x_66; -lean_dec(x_51); -lean_dec_ref(x_50); -lean_dec_ref(x_49); -x_66 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_66, 0, x_11); -return x_66; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("◾", 3, 1); +return x_1; } } -else +static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__1() { +_start: { -lean_object* x_67; -lean_dec(x_50); -lean_dec_ref(x_49); -x_67 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_67, 0, x_11); -return x_67; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_IR_CtorFieldInfo_format___closed__0; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } -else +static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__2() { +_start: { -lean_object* x_68; -lean_dec_ref(x_49); -x_68 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_68, 0, x_11); -return x_68; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("obj@", 4, 4); +return x_1; } } -else -{ -lean_object* x_69; -lean_dec_ref(x_41); -x_69 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_69, 0, x_11); -return x_69; -} -} -} -} -else -{ -lean_object* x_70; -lean_dec_ref(x_38); -lean_dec_ref(x_37); -x_70 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_70, 0, x_11); -return x_70; -} -} -else -{ -lean_object* x_71; -lean_dec(x_38); -lean_dec_ref(x_37); -x_71 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_71, 0, x_11); -return x_71; -} -} -else +static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__3() { +_start: { -lean_object* x_72; -lean_dec(x_37); -x_72 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_72, 0, x_11); -return x_72; -} -} +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_IR_CtorFieldInfo_format___closed__2; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } -else -{ -uint8_t x_73; -lean_dec(x_11); -x_73 = !lean_is_exclusive(x_13); -if (x_73 == 0) -{ -return x_13; } -else +static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__4() { +_start: { -lean_object* x_74; lean_object* x_75; -x_74 = lean_ctor_get(x_13, 0); -lean_inc(x_74); -lean_dec(x_13); -x_75 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_75, 0, x_74); -return x_75; -} +lean_object* x_1; +x_1 = lean_mk_string_unchecked(":", 1, 1); +return x_1; } } -else +static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__5() { +_start: { -lean_dec(x_11); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -return x_10; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_IR_CtorFieldInfo_format___closed__4; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } -else +static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__6() { +_start: { -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -return x_10; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("usize@", 6, 6); +return x_1; } } -else +static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__7() { +_start: { -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -return x_7; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_IR_CtorFieldInfo_format___closed__6; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } -else +static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__8() { +_start: { -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -return x_7; -} +lean_object* x_1; +x_1 = lean_mk_string_unchecked("scalar#", 7, 7); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_hasTrivialStructure_x3f___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__9() { _start: { -lean_object* x_7; -x_7 = l_Lean_IR_hasTrivialStructure_x3f___lam__0(x_1, x_2, x_3, x_4, x_5); -return x_7; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_IR_CtorFieldInfo_format___closed__8; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_IR_hasTrivialStructure_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__10() { _start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = lean_alloc_closure((void*)(l_Lean_IR_hasTrivialStructure_x3f___lam__0___boxed), 6, 0); -x_6 = l_Lean_IR_hasTrivialStructure_x3f___closed__0; -x_7 = l_Lean_Compiler_LCNF_Irrelevant_hasTrivialStructure_x3f(x_6, x_5, x_1, x_2, x_3); -return x_7; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("@", 1, 1); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_hasTrivialStructure_x3f___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__11() { _start: { -lean_object* x_5; -x_5 = l_Lean_IR_hasTrivialStructure_x3f(x_1, x_2, x_3); -return x_5; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_IR_CtorFieldInfo_format___closed__10; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2___redArg___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__12() { _start: { -lean_object* x_9; -x_9 = lean_apply_7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, lean_box(0)); -return x_9; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("void", 4, 4); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2___redArg___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__13() { _start: { -lean_object* x_9; -x_9 = l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2___redArg___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -return x_9; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_IR_CtorFieldInfo_format___closed__12; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2___redArg(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_format(lean_object* x_1) { _start: { -lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_alloc_closure((void*)(l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2___redArg___lam__0___boxed), 8, 1); -lean_closure_set(x_9, 0, x_2); -x_10 = 0; -x_11 = lean_box(0); -x_12 = l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAuxAux(lean_box(0), x_10, x_11, x_1, x_9, x_3, x_10, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_12) == 0) +switch (lean_obj_tag(x_1)) { +case 0: { -uint8_t x_13; -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) +lean_object* x_2; +x_2 = l_Lean_IR_CtorFieldInfo_format___closed__1; +return x_2; +} +case 1: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) { +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_4 = lean_ctor_get(x_1, 0); +x_5 = lean_ctor_get(x_1, 1); +x_6 = l_Lean_IR_CtorFieldInfo_format___closed__3; +x_7 = l_Nat_reprFast(x_4); +x_8 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set_tag(x_1, 5); +lean_ctor_set(x_1, 1, x_8); +lean_ctor_set(x_1, 0, x_6); +x_9 = l_Lean_IR_CtorFieldInfo_format___closed__5; +x_10 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_10, 0, x_1); +lean_ctor_set(x_10, 1, x_9); +x_11 = l_Lean_IR_instToFormatIRType___private__1(x_5); +x_12 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set(x_12, 1, x_11); return x_12; } else { -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_12, 0); +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_13 = lean_ctor_get(x_1, 0); +x_14 = lean_ctor_get(x_1, 1); lean_inc(x_14); -lean_dec(x_12); -x_15 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_15, 0, x_14); -return x_15; +lean_inc(x_13); +lean_dec(x_1); +x_15 = l_Lean_IR_CtorFieldInfo_format___closed__3; +x_16 = l_Nat_reprFast(x_13); +x_17 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_17, 0, x_16); +x_18 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_18, 0, x_15); +lean_ctor_set(x_18, 1, x_17); +x_19 = l_Lean_IR_CtorFieldInfo_format___closed__5; +x_20 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +x_21 = l_Lean_IR_instToFormatIRType___private__1(x_14); +x_22 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; } } -else +case 2: { -uint8_t x_16; -x_16 = !lean_is_exclusive(x_12); -if (x_16 == 0) +uint8_t x_23; +x_23 = !lean_is_exclusive(x_1); +if (x_23 == 0) { -return x_12; +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_24 = lean_ctor_get(x_1, 0); +x_25 = l_Lean_IR_CtorFieldInfo_format___closed__7; +x_26 = l_Nat_reprFast(x_24); +lean_ctor_set_tag(x_1, 3); +lean_ctor_set(x_1, 0, x_26); +x_27 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_1); +return x_27; } else { -lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_12, 0); -lean_inc(x_17); -lean_dec(x_12); -x_18 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_18, 0, x_17); -return x_18; -} +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_28 = lean_ctor_get(x_1, 0); +lean_inc(x_28); +lean_dec(x_1); +x_29 = l_Lean_IR_CtorFieldInfo_format___closed__7; +x_30 = l_Nat_reprFast(x_28); +x_31 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_31, 0, x_30); +x_32 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_32, 0, x_29); +lean_ctor_set(x_32, 1, x_31); +return x_32; } } +case 3: +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_33 = lean_ctor_get(x_1, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_1, 1); +lean_inc(x_34); +x_35 = lean_ctor_get(x_1, 2); +lean_inc(x_35); +lean_dec_ref(x_1); +x_36 = l_Lean_IR_CtorFieldInfo_format___closed__9; +x_37 = l_Nat_reprFast(x_33); +x_38 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_38, 0, x_37); +x_39 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_39, 0, x_36); +lean_ctor_set(x_39, 1, x_38); +x_40 = l_Lean_IR_CtorFieldInfo_format___closed__11; +x_41 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +x_42 = l_Nat_reprFast(x_34); +x_43 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_43, 0, x_42); +x_44 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_44, 0, x_41); +lean_ctor_set(x_44, 1, x_43); +x_45 = l_Lean_IR_CtorFieldInfo_format___closed__5; +x_46 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +x_47 = l_Lean_IR_instToFormatIRType___private__1(x_35); +x_48 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_48, 0, x_46); +lean_ctor_set(x_48, 1, x_47); +return x_48; } -LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: +default: { -lean_object* x_10; -x_10 = l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_10; +lean_object* x_49; +x_49 = l_Lean_IR_CtorFieldInfo_format___closed__13; +return x_49; } } -static lean_object* _init_l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__2() { +} +} +static lean_object* _init_l_Lean_IR_CtorFieldInfo_instToFormat___closed__0() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("unreachable code has been reached", 33, 33); +x_1 = lean_alloc_closure((void*)(l_Lean_IR_CtorFieldInfo_format), 1, 0); return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__1() { +static lean_object* _init_l_Lean_IR_CtorFieldInfo_instToFormat() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.IR.nameToIRType.fillCache", 30, 30); +x_1 = l_Lean_IR_CtorFieldInfo_instToFormat___closed__0; return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__0() { +static lean_object* _init_l_Lean_IR_instInhabitedCtorLayout_default___closed__0() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Compiler.IR.ToIRType", 25, 25); +x_1 = l_Lean_IR_instInhabitedCtorInfo_default; return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__3() { +static lean_object* _init_l_Lean_IR_instInhabitedCtorLayout_default___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__2; -x_2 = lean_unsigned_to_nat(62u); -x_3 = lean_unsigned_to_nat(81u); -x_4 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__1; -x_5 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__0; -x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); -return x_6; +lean_object* x_1; +x_1 = l_Array_empty(lean_box(0)); +return x_1; } } -static lean_object* _init_l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__0___closed__0() { +static lean_object* _init_l_Lean_IR_instInhabitedCtorLayout_default___closed__2() { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Core_instInhabitedCoreM___lam__0___boxed), 3, 0); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_IR_instInhabitedCtorLayout_default___closed__1; +x_2 = l_Lean_IR_instInhabitedCtorLayout_default___closed__0; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_IR_instInhabitedCtorLayout_default() { _start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__0___closed__0; -x_6 = lean_panic_fn(x_5, x_1); -x_7 = lean_apply_3(x_6, x_2, x_3, lean_box(0)); -return x_7; +lean_object* x_1; +x_1 = l_Lean_IR_instInhabitedCtorLayout_default___closed__2; +return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__4() { +static lean_object* _init_l_Lean_IR_instInhabitedCtorLayout() { _start: { lean_object* x_1; -x_1 = l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_box(0), lean_box(0)); +x_1 = l_Lean_IR_instInhabitedCtorLayout_default; return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__5() { +LEAN_EXPORT lean_object* l_Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_895574536____hygCtx___hyg_2_() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__4; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_IR_instInhabitedCtorLayout_default; +x_3 = l_Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0___redArg(x_2); +return x_3; } } -static lean_object* _init_l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__6() { +LEAN_EXPORT lean_object* l_Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_895574536____hygCtx___hyg_2____boxed(lean_object* x_1) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__5; -x_2 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_2, 0, x_1); -lean_ctor_set(x_2, 1, x_1); -lean_ctor_set(x_2, 2, x_1); -lean_ctor_set(x_2, 3, x_1); -lean_ctor_set(x_2, 4, x_1); -lean_ctor_set(x_2, 5, x_1); +lean_object* x_2; +x_2 = l_Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_895574536____hygCtx___hyg_2_(); return x_2; } } -static lean_object* _init_l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__7() { +static lean_object* _init_l_Lean_IR_hasTrivialStructure_x3f___closed__0() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(32u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; +lean_object* x_1; +x_1 = l_Lean_IR_trivialStructureInfoExt; +return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__8() { +static lean_object* _init_l_Lean_IR_hasTrivialStructure_x3f___lam__0___closed__0() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__7; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Subtype", 7, 7); +return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__9() { +static lean_object* _init_l_Lean_IR_hasTrivialStructure_x3f___lam__0___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__5; -x_2 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_2, 0, x_1); -lean_ctor_set(x_2, 1, x_1); -lean_ctor_set(x_2, 2, x_1); -lean_ctor_set(x_2, 3, x_1); -lean_ctor_set(x_2, 4, x_1); -return x_2; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Void", 4, 4); +return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___lam__0___closed__0() { +static lean_object* _init_l_Lean_IR_hasTrivialStructure_x3f___lam__0___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = lean_box(0); -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("nonemptyType", 12, 12); +return x_1; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__1___redArg(lean_object* x_1, lean_object* x_2, lean_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_IR_hasTrivialStructure_x3f___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -uint8_t x_10; -x_10 = !lean_is_exclusive(x_3); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_11 = lean_ctor_get(x_3, 0); -x_12 = lean_ctor_get(x_3, 1); -x_13 = lean_ctor_get(x_3, 2); -x_14 = lean_nat_dec_lt(x_12, x_13); -if (x_14 == 0) +lean_object* x_7; +lean_inc(x_5); +lean_inc_ref(x_4); +lean_inc(x_3); +lean_inc_ref(x_2); +lean_inc_ref(x_1); +x_7 = l_Lean_Meta_isProp(x_1, x_2, x_3, x_4, x_5); +if (lean_obj_tag(x_7) == 0) { -lean_object* x_15; -lean_free_object(x_3); -lean_dec(x_13); -lean_dec(x_12); -lean_dec_ref(x_11); +lean_object* x_8; uint8_t x_9; +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_unbox(x_8); lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec_ref(x_2); -x_15 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_15, 0, x_4); -return x_15; -} -else +if (x_9 == 0) { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -lean_dec_ref(x_4); -x_16 = lean_array_fget_borrowed(x_11, x_12); -x_17 = l_Lean_Expr_fvarId_x21(x_16); -lean_inc_ref(x_5); -x_18 = l_Lean_FVarId_getType___redArg(x_17, x_5, x_7, x_8); -if (lean_obj_tag(x_18) == 0) +lean_object* x_10; +lean_dec_ref(x_7); +lean_inc(x_5); +lean_inc_ref(x_4); +lean_inc(x_3); +lean_inc_ref(x_2); +lean_inc_ref(x_1); +x_10 = l_Lean_Meta_isTypeFormerType(x_1, x_2, x_3, x_4, x_5); +if (lean_obj_tag(x_10) == 0) { -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -lean_dec_ref(x_18); -lean_inc(x_8); -lean_inc_ref(x_7); -lean_inc(x_6); -lean_inc_ref(x_5); -x_20 = l_Lean_Compiler_LCNF_toLCNFType(x_19, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_20) == 0) +lean_object* x_11; uint8_t x_12; +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_unbox(x_11); +if (x_12 == 0) { -lean_object* x_21; lean_object* x_22; -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -lean_dec_ref(x_20); -lean_inc(x_8); -lean_inc_ref(x_7); -x_22 = l_Lean_Compiler_LCNF_toMonoType(x_21, x_7, x_8); -if (lean_obj_tag(x_22) == 0) +lean_object* x_13; +lean_dec_ref(x_10); +x_13 = l_Lean_Meta_whnfD(x_1, x_2, x_3, x_4, x_5); +if (lean_obj_tag(x_13) == 0) { -uint8_t x_23; -x_23 = !lean_is_exclusive(x_22); -if (x_23 == 0) +uint8_t x_14; +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) { -lean_object* x_24; uint8_t x_25; -x_24 = lean_ctor_get(x_22, 0); -x_25 = l_Lean_Expr_isErased(x_24); -lean_dec(x_24); -if (x_25 == 0) +lean_object* x_15; +x_15 = lean_ctor_get(x_13, 0); +if (lean_obj_tag(x_15) == 11) { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -lean_free_object(x_3); -lean_dec(x_13); -lean_dec(x_12); -lean_dec_ref(x_11); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec_ref(x_2); -x_26 = lean_box(x_14); -x_27 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_27, 0, x_26); -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_28, 1, x_1); -lean_ctor_set(x_22, 0, x_28); -return x_22; +lean_object* x_16; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +if (lean_obj_tag(x_16) == 1) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_18 = lean_ctor_get(x_15, 1); +lean_inc(x_18); +x_19 = lean_ctor_get(x_15, 2); +lean_inc_ref(x_19); +lean_dec_ref(x_15); +x_20 = lean_ctor_get(x_16, 1); +lean_inc_ref(x_20); +lean_dec_ref(x_16); +x_21 = l_Lean_IR_hasTrivialStructure_x3f___lam__0___closed__0; +x_22 = lean_string_dec_eq(x_20, x_21); +lean_dec_ref(x_20); +if (x_22 == 0) +{ +lean_dec_ref(x_19); +lean_dec(x_18); +lean_ctor_set(x_13, 0, x_11); +return x_13; } else { -lean_object* x_29; lean_object* x_30; -lean_free_object(x_22); -x_29 = lean_unsigned_to_nat(1u); -x_30 = lean_nat_add(x_12, x_29); -lean_dec(x_12); -lean_ctor_set(x_3, 1, x_30); -lean_inc_ref(x_2); +lean_object* x_23; uint8_t x_24; +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_nat_dec_eq(x_18, x_23); +lean_dec(x_18); +if (x_24 == 0) { -lean_object* _tmp_3 = x_2; -x_4 = _tmp_3; -} -goto _start; -} +lean_dec_ref(x_19); +lean_ctor_set(x_13, 0, x_11); +return x_13; } else { -lean_object* x_32; uint8_t x_33; -x_32 = lean_ctor_get(x_22, 0); -lean_inc(x_32); -lean_dec(x_22); -x_33 = l_Lean_Expr_isErased(x_32); -lean_dec(x_32); +if (lean_obj_tag(x_19) == 5) +{ +lean_object* x_25; +x_25 = lean_ctor_get(x_19, 0); +lean_inc_ref(x_25); +lean_dec_ref(x_19); +if (lean_obj_tag(x_25) == 4) +{ +lean_object* x_26; +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +if (lean_obj_tag(x_26) == 1) +{ +lean_object* x_27; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +if (lean_obj_tag(x_27) == 1) +{ +lean_object* x_28; +x_28 = lean_ctor_get(x_27, 0); +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_29 = lean_ctor_get(x_25, 1); +lean_inc(x_29); +lean_dec_ref(x_25); +x_30 = lean_ctor_get(x_26, 1); +lean_inc_ref(x_30); +lean_dec_ref(x_26); +x_31 = lean_ctor_get(x_27, 1); +lean_inc_ref(x_31); +lean_dec_ref(x_27); +x_32 = l_Lean_IR_hasTrivialStructure_x3f___lam__0___closed__1; +x_33 = lean_string_dec_eq(x_31, x_32); +lean_dec_ref(x_31); if (x_33 == 0) { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -lean_free_object(x_3); -lean_dec(x_13); -lean_dec(x_12); -lean_dec_ref(x_11); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec_ref(x_2); -x_34 = lean_box(x_14); -x_35 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_35, 0, x_34); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_1); -x_37 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_37, 0, x_36); -return x_37; +lean_dec_ref(x_30); +lean_dec(x_29); +lean_ctor_set(x_13, 0, x_11); +return x_13; } else { -lean_object* x_38; lean_object* x_39; -x_38 = lean_unsigned_to_nat(1u); -x_39 = lean_nat_add(x_12, x_38); -lean_dec(x_12); -lean_ctor_set(x_3, 1, x_39); -lean_inc_ref(x_2); +lean_object* x_34; uint8_t x_35; +x_34 = l_Lean_IR_hasTrivialStructure_x3f___lam__0___closed__2; +x_35 = lean_string_dec_eq(x_30, x_34); +lean_dec_ref(x_30); +if (x_35 == 0) { -lean_object* _tmp_3 = x_2; -x_4 = _tmp_3; +lean_dec(x_29); +lean_ctor_set(x_13, 0, x_11); +return x_13; +} +else +{ +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_36; +lean_dec(x_11); +x_36 = lean_box(x_35); +lean_ctor_set(x_13, 0, x_36); +return x_13; +} +else +{ +lean_dec(x_29); +lean_ctor_set(x_13, 0, x_11); +return x_13; } -goto _start; } } } else { -uint8_t x_41; -lean_free_object(x_3); -lean_dec(x_13); -lean_dec(x_12); -lean_dec_ref(x_11); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec_ref(x_2); -x_41 = !lean_is_exclusive(x_22); -if (x_41 == 0) -{ -return x_22; +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_ctor_set(x_13, 0, x_11); +return x_13; +} } else { -lean_object* x_42; lean_object* x_43; -x_42 = lean_ctor_get(x_22, 0); -lean_inc(x_42); -lean_dec(x_22); -x_43 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_43, 0, x_42); -return x_43; -} +lean_dec(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_ctor_set(x_13, 0, x_11); +return x_13; } } else { -uint8_t x_44; -lean_free_object(x_3); -lean_dec(x_13); -lean_dec(x_12); -lean_dec_ref(x_11); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec_ref(x_2); -x_44 = !lean_is_exclusive(x_20); -if (x_44 == 0) +lean_dec(x_26); +lean_dec_ref(x_25); +lean_ctor_set(x_13, 0, x_11); +return x_13; +} +} +else { -return x_20; +lean_dec_ref(x_25); +lean_ctor_set(x_13, 0, x_11); +return x_13; +} } else { -lean_object* x_45; lean_object* x_46; -x_45 = lean_ctor_get(x_20, 0); -lean_inc(x_45); -lean_dec(x_20); -x_46 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_46, 0, x_45); -return x_46; +lean_dec_ref(x_19); +lean_ctor_set(x_13, 0, x_11); +return x_13; +} } } } else { -uint8_t x_47; -lean_free_object(x_3); -lean_dec(x_13); -lean_dec(x_12); -lean_dec_ref(x_11); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec_ref(x_2); -x_47 = !lean_is_exclusive(x_18); -if (x_47 == 0) -{ -return x_18; +lean_dec_ref(x_16); +lean_dec_ref(x_15); +lean_ctor_set(x_13, 0, x_11); +return x_13; +} } else { -lean_object* x_48; lean_object* x_49; -x_48 = lean_ctor_get(x_18, 0); -lean_inc(x_48); -lean_dec(x_18); -x_49 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_49, 0, x_48); -return x_49; +lean_dec(x_16); +lean_dec_ref(x_15); +lean_ctor_set(x_13, 0, x_11); +return x_13; } } +else +{ +lean_dec(x_15); +lean_ctor_set(x_13, 0, x_11); +return x_13; } } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; -x_50 = lean_ctor_get(x_3, 0); -x_51 = lean_ctor_get(x_3, 1); -x_52 = lean_ctor_get(x_3, 2); -lean_inc(x_52); -lean_inc(x_51); -lean_inc(x_50); -lean_dec(x_3); -x_53 = lean_nat_dec_lt(x_51, x_52); -if (x_53 == 0) +lean_object* x_37; +x_37 = lean_ctor_get(x_13, 0); +lean_inc(x_37); +lean_dec(x_13); +if (lean_obj_tag(x_37) == 11) { -lean_object* x_54; -lean_dec(x_52); -lean_dec(x_51); -lean_dec_ref(x_50); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec_ref(x_2); -x_54 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_54, 0, x_4); -return x_54; +lean_object* x_38; +x_38 = lean_ctor_get(x_37, 0); +lean_inc(x_38); +if (lean_obj_tag(x_38) == 1) +{ +lean_object* x_39; +x_39 = lean_ctor_get(x_38, 0); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_40 = lean_ctor_get(x_37, 1); +lean_inc(x_40); +x_41 = lean_ctor_get(x_37, 2); +lean_inc_ref(x_41); +lean_dec_ref(x_37); +x_42 = lean_ctor_get(x_38, 1); +lean_inc_ref(x_42); +lean_dec_ref(x_38); +x_43 = l_Lean_IR_hasTrivialStructure_x3f___lam__0___closed__0; +x_44 = lean_string_dec_eq(x_42, x_43); +lean_dec_ref(x_42); +if (x_44 == 0) +{ +lean_object* x_45; +lean_dec_ref(x_41); +lean_dec(x_40); +x_45 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_45, 0, x_11); +return x_45; } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; -lean_dec_ref(x_4); -x_55 = lean_array_fget_borrowed(x_50, x_51); -x_56 = l_Lean_Expr_fvarId_x21(x_55); -lean_inc_ref(x_5); -x_57 = l_Lean_FVarId_getType___redArg(x_56, x_5, x_7, x_8); -if (lean_obj_tag(x_57) == 0) +lean_object* x_46; uint8_t x_47; +x_46 = lean_unsigned_to_nat(0u); +x_47 = lean_nat_dec_eq(x_40, x_46); +lean_dec(x_40); +if (x_47 == 0) { -lean_object* x_58; lean_object* x_59; -x_58 = lean_ctor_get(x_57, 0); -lean_inc(x_58); -lean_dec_ref(x_57); -lean_inc(x_8); -lean_inc_ref(x_7); -lean_inc(x_6); -lean_inc_ref(x_5); -x_59 = l_Lean_Compiler_LCNF_toLCNFType(x_58, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_59) == 0) +lean_object* x_48; +lean_dec_ref(x_41); +x_48 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_48, 0, x_11); +return x_48; +} +else { -lean_object* x_60; lean_object* x_61; -x_60 = lean_ctor_get(x_59, 0); -lean_inc(x_60); -lean_dec_ref(x_59); -lean_inc(x_8); -lean_inc_ref(x_7); -x_61 = l_Lean_Compiler_LCNF_toMonoType(x_60, x_7, x_8); -if (lean_obj_tag(x_61) == 0) +if (lean_obj_tag(x_41) == 5) { -lean_object* x_62; lean_object* x_63; uint8_t x_64; -x_62 = lean_ctor_get(x_61, 0); -lean_inc(x_62); -if (lean_is_exclusive(x_61)) { - lean_ctor_release(x_61, 0); - x_63 = x_61; -} else { - lean_dec_ref(x_61); - x_63 = lean_box(0); -} -x_64 = l_Lean_Expr_isErased(x_62); -lean_dec(x_62); -if (x_64 == 0) +lean_object* x_49; +x_49 = lean_ctor_get(x_41, 0); +lean_inc_ref(x_49); +lean_dec_ref(x_41); +if (lean_obj_tag(x_49) == 4) { -lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; -lean_dec(x_52); -lean_dec(x_51); +lean_object* x_50; +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +if (lean_obj_tag(x_50) == 1) +{ +lean_object* x_51; +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +if (lean_obj_tag(x_51) == 1) +{ +lean_object* x_52; +x_52 = lean_ctor_get(x_51, 0); +if (lean_obj_tag(x_52) == 0) +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; +x_53 = lean_ctor_get(x_49, 1); +lean_inc(x_53); +lean_dec_ref(x_49); +x_54 = lean_ctor_get(x_50, 1); +lean_inc_ref(x_54); lean_dec_ref(x_50); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec_ref(x_2); -x_65 = lean_box(x_53); -x_66 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_66, 0, x_65); -x_67 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_67, 0, x_66); -lean_ctor_set(x_67, 1, x_1); -if (lean_is_scalar(x_63)) { - x_68 = lean_alloc_ctor(0, 1, 0); -} else { - x_68 = x_63; +x_55 = lean_ctor_get(x_51, 1); +lean_inc_ref(x_55); +lean_dec_ref(x_51); +x_56 = l_Lean_IR_hasTrivialStructure_x3f___lam__0___closed__1; +x_57 = lean_string_dec_eq(x_55, x_56); +lean_dec_ref(x_55); +if (x_57 == 0) +{ +lean_object* x_58; +lean_dec_ref(x_54); +lean_dec(x_53); +x_58 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_58, 0, x_11); +return x_58; } -lean_ctor_set(x_68, 0, x_67); -return x_68; +else +{ +lean_object* x_59; uint8_t x_60; +x_59 = l_Lean_IR_hasTrivialStructure_x3f___lam__0___closed__2; +x_60 = lean_string_dec_eq(x_54, x_59); +lean_dec_ref(x_54); +if (x_60 == 0) +{ +lean_object* x_61; +lean_dec(x_53); +x_61 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_61, 0, x_11); +return x_61; } else { -lean_object* x_69; lean_object* x_70; lean_object* x_71; -lean_dec(x_63); -x_69 = lean_unsigned_to_nat(1u); -x_70 = lean_nat_add(x_51, x_69); -lean_dec(x_51); -x_71 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_71, 0, x_50); -lean_ctor_set(x_71, 1, x_70); -lean_ctor_set(x_71, 2, x_52); -lean_inc_ref(x_2); +if (lean_obj_tag(x_53) == 0) { -lean_object* _tmp_2 = x_71; -lean_object* _tmp_3 = x_2; -x_3 = _tmp_2; -x_4 = _tmp_3; +lean_object* x_62; lean_object* x_63; +lean_dec(x_11); +x_62 = lean_box(x_60); +x_63 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_63, 0, x_62); +return x_63; +} +else +{ +lean_object* x_64; +lean_dec(x_53); +x_64 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_64, 0, x_11); +return x_64; +} } -goto _start; } } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; -lean_dec(x_52); -lean_dec(x_51); +lean_object* x_65; +lean_dec_ref(x_51); lean_dec_ref(x_50); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec_ref(x_2); -x_73 = lean_ctor_get(x_61, 0); -lean_inc(x_73); -if (lean_is_exclusive(x_61)) { - lean_ctor_release(x_61, 0); - x_74 = x_61; -} else { - lean_dec_ref(x_61); - x_74 = lean_box(0); -} -if (lean_is_scalar(x_74)) { - x_75 = lean_alloc_ctor(1, 1, 0); -} else { - x_75 = x_74; -} -lean_ctor_set(x_75, 0, x_73); -return x_75; +lean_dec_ref(x_49); +x_65 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_65, 0, x_11); +return x_65; } } else { -lean_object* x_76; lean_object* x_77; lean_object* x_78; -lean_dec(x_52); +lean_object* x_66; lean_dec(x_51); lean_dec_ref(x_50); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec_ref(x_2); -x_76 = lean_ctor_get(x_59, 0); -lean_inc(x_76); -if (lean_is_exclusive(x_59)) { - lean_ctor_release(x_59, 0); - x_77 = x_59; -} else { - lean_dec_ref(x_59); - x_77 = lean_box(0); +lean_dec_ref(x_49); +x_66 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_66, 0, x_11); +return x_66; } -if (lean_is_scalar(x_77)) { - x_78 = lean_alloc_ctor(1, 1, 0); -} else { - x_78 = x_77; } -lean_ctor_set(x_78, 0, x_76); -return x_78; +else +{ +lean_object* x_67; +lean_dec(x_50); +lean_dec_ref(x_49); +x_67 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_67, 0, x_11); +return x_67; } } else { -lean_object* x_79; lean_object* x_80; lean_object* x_81; -lean_dec(x_52); -lean_dec(x_51); -lean_dec_ref(x_50); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec_ref(x_2); -x_79 = lean_ctor_get(x_57, 0); -lean_inc(x_79); -if (lean_is_exclusive(x_57)) { - lean_ctor_release(x_57, 0); - x_80 = x_57; -} else { - lean_dec_ref(x_57); - x_80 = lean_box(0); -} -if (lean_is_scalar(x_80)) { - x_81 = lean_alloc_ctor(1, 1, 0); -} else { - x_81 = x_80; +lean_object* x_68; +lean_dec_ref(x_49); +x_68 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_68, 0, x_11); +return x_68; } -lean_ctor_set(x_81, 0, x_79); -return x_81; } +else +{ +lean_object* x_69; +lean_dec_ref(x_41); +x_69 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_69, 0, x_11); +return x_69; } } } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___lam__0(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_10; lean_object* x_11; lean_object* x_31; lean_object* x_32; uint8_t x_33; -x_31 = lean_unsigned_to_nat(0u); -x_32 = lean_array_get_size(x_3); -x_33 = lean_nat_dec_le(x_2, x_31); -if (x_33 == 0) +else { -x_10 = x_2; -x_11 = x_32; -goto block_30; +lean_object* x_70; +lean_dec_ref(x_38); +lean_dec_ref(x_37); +x_70 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_70, 0, x_11); +return x_70; +} } else { -lean_dec(x_2); -x_10 = x_31; -x_11 = x_32; -goto block_30; +lean_object* x_71; +lean_dec(x_38); +lean_dec_ref(x_37); +x_71 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_71, 0, x_11); +return x_71; } -block_30: -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_12 = l_Array_toSubarray___redArg(x_3, x_10, x_11); -x_13 = lean_box(0); -x_14 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___lam__0___closed__0; -x_15 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__1___redArg(x_13, x_14, x_12, x_14, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_15) == 0) -{ -uint8_t x_16; -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_15, 0); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -lean_dec(x_17); -if (lean_obj_tag(x_18) == 0) -{ -lean_object* x_19; -x_19 = lean_box(x_1); -lean_ctor_set(x_15, 0, x_19); -return x_15; } else { -lean_object* x_20; -x_20 = lean_ctor_get(x_18, 0); -lean_inc(x_20); -lean_dec_ref(x_18); -lean_ctor_set(x_15, 0, x_20); -return x_15; +lean_object* x_72; +lean_dec(x_37); +x_72 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_72, 0, x_11); +return x_72; +} } } else { -lean_object* x_21; lean_object* x_22; -x_21 = lean_ctor_get(x_15, 0); -lean_inc(x_21); -lean_dec(x_15); -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -lean_dec(x_21); -if (lean_obj_tag(x_22) == 0) +uint8_t x_73; +lean_dec(x_11); +x_73 = !lean_is_exclusive(x_13); +if (x_73 == 0) { -lean_object* x_23; lean_object* x_24; -x_23 = lean_box(x_1); -x_24 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_24, 0, x_23); -return x_24; +return x_13; } else { -lean_object* x_25; lean_object* x_26; -x_25 = lean_ctor_get(x_22, 0); -lean_inc(x_25); -lean_dec_ref(x_22); -x_26 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_26, 0, x_25); -return x_26; +lean_object* x_74; lean_object* x_75; +x_74 = lean_ctor_get(x_13, 0); +lean_inc(x_74); +lean_dec(x_13); +x_75 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_75, 0, x_74); +return x_75; } } } else { -uint8_t x_27; -x_27 = !lean_is_exclusive(x_15); -if (x_27 == 0) -{ -return x_15; +lean_dec(x_11); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_10; +} } else { -lean_object* x_28; lean_object* x_29; -x_28 = lean_ctor_get(x_15, 0); -lean_inc(x_28); -lean_dec(x_15); -x_29 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_29, 0, x_28); -return x_29; -} +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_10; } } +else +{ +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_7; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: +else { -uint8_t x_10; lean_object* x_11; -x_10 = lean_unbox(x_1); -x_11 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___lam__0(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_5); lean_dec_ref(x_4); -return x_11; +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_7; } } -static lean_object* _init_l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__10() { +} +LEAN_EXPORT lean_object* l_Lean_IR_hasTrivialStructure_x3f___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_1; -x_1 = l_Lean_Meta_instInhabitedConfigWithKey___private__1; -return x_1; +lean_object* x_7; +x_7 = l_Lean_IR_hasTrivialStructure_x3f___lam__0(x_1, x_2, x_3, x_4, x_5); +return x_7; } } -static lean_object* _init_l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__11() { +LEAN_EXPORT lean_object* l_Lean_IR_hasTrivialStructure_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _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_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_alloc_closure((void*)(l_Lean_IR_hasTrivialStructure_x3f___lam__0___boxed), 6, 0); +x_6 = l_Lean_IR_hasTrivialStructure_x3f___closed__0; +x_7 = l_Lean_Compiler_LCNF_Irrelevant_hasTrivialStructure_x3f(x_6, x_5, x_1, x_2, x_3); +return x_7; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg(lean_object* x_1, lean_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_IR_hasTrivialStructure_x3f___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -if (lean_obj_tag(x_4) == 0) +lean_object* x_5; +x_5 = l_Lean_IR_hasTrivialStructure_x3f(x_1, x_2, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2___redArg___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: { lean_object* x_9; -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec_ref(x_1); -x_9 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_9, 0, x_5); +x_9 = lean_apply_7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, lean_box(0)); return x_9; } -else -{ -lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_29; lean_object* x_30; -x_10 = lean_ctor_get(x_4, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_4, 1); -lean_inc(x_11); -lean_dec_ref(x_4); -x_29 = 0; -lean_inc_ref(x_1); -x_30 = l_Lean_Environment_find_x3f(x_1, x_10, x_29); -if (lean_obj_tag(x_30) == 1) +} +LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2___redArg___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -lean_object* x_31; -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -lean_dec_ref(x_30); -if (lean_obj_tag(x_31) == 6) +lean_object* x_9; +x_9 = l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2___redArg___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2___redArg(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: { -lean_object* x_32; size_t x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; -x_32 = lean_ctor_get(x_31, 0); -lean_inc_ref(x_32); -lean_dec_ref(x_31); -x_33 = 5; -x_34 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__5; -lean_inc_n(x_3, 3); -x_35 = lean_alloc_ctor(0, 9, 0); -lean_ctor_set(x_35, 0, x_3); -lean_ctor_set(x_35, 1, x_3); -lean_ctor_set(x_35, 2, x_3); -lean_ctor_set(x_35, 3, x_34); -lean_ctor_set(x_35, 4, x_34); -lean_ctor_set(x_35, 5, x_34); -lean_ctor_set(x_35, 6, x_34); -lean_ctor_set(x_35, 7, x_34); -lean_ctor_set(x_35, 8, x_34); -x_36 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__6; -x_37 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__7; -x_38 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__8; -lean_inc_n(x_3, 2); -x_39 = lean_alloc_ctor(0, 4, sizeof(size_t)*1); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_37); -lean_ctor_set(x_39, 2, x_3); -lean_ctor_set(x_39, 3, x_3); -lean_ctor_set_usize(x_39, 4, x_33); -x_40 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__9; -lean_inc_ref(x_39); -lean_inc(x_2); -x_41 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_41, 0, x_35); -lean_ctor_set(x_41, 1, x_36); -lean_ctor_set(x_41, 2, x_2); -lean_ctor_set(x_41, 3, x_39); -lean_ctor_set(x_41, 4, x_40); -x_42 = lean_st_mk_ref(x_41); -x_43 = lean_ctor_get(x_32, 0); -lean_inc_ref(x_43); -x_44 = lean_ctor_get(x_32, 3); -lean_inc(x_44); -lean_dec_ref(x_32); -x_45 = !lean_is_exclusive(x_43); -if (x_45 == 0) +lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_alloc_closure((void*)(l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2___redArg___lam__0___boxed), 8, 1); +lean_closure_set(x_9, 0, x_2); +x_10 = 0; +x_11 = lean_box(0); +x_12 = l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAuxAux(lean_box(0), x_10, x_11, x_1, x_9, x_3, x_10, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_12) == 0) { -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; lean_object* x_56; lean_object* x_57; -x_46 = lean_ctor_get(x_43, 2); -x_47 = lean_ctor_get(x_43, 1); -lean_dec(x_47); -x_48 = lean_ctor_get(x_43, 0); -lean_dec(x_48); -x_49 = lean_box(x_29); -x_50 = lean_alloc_closure((void*)(l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___lam__0___boxed), 9, 2); -lean_closure_set(x_50, 0, x_49); -lean_closure_set(x_50, 1, x_44); -x_51 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__10; -x_52 = lean_box(1); -lean_ctor_set(x_43, 2, x_52); -lean_ctor_set(x_43, 1, x_39); -lean_ctor_set(x_43, 0, x_34); -x_53 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__11; -x_54 = lean_box(0); -x_55 = 1; -lean_inc(x_3); -lean_inc(x_2); -x_56 = lean_alloc_ctor(0, 7, 4); -lean_ctor_set(x_56, 0, x_51); -lean_ctor_set(x_56, 1, x_2); -lean_ctor_set(x_56, 2, x_43); -lean_ctor_set(x_56, 3, x_53); -lean_ctor_set(x_56, 4, x_54); -lean_ctor_set(x_56, 5, x_3); -lean_ctor_set(x_56, 6, x_54); -lean_ctor_set_uint8(x_56, sizeof(void*)*7, x_29); -lean_ctor_set_uint8(x_56, sizeof(void*)*7 + 1, x_29); -lean_ctor_set_uint8(x_56, sizeof(void*)*7 + 2, x_29); -lean_ctor_set_uint8(x_56, sizeof(void*)*7 + 3, x_55); -lean_inc(x_7); -lean_inc_ref(x_6); -lean_inc(x_42); -x_57 = l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2___redArg(x_46, x_50, x_29, x_56, x_42, x_6, x_7); -if (lean_obj_tag(x_57) == 0) +uint8_t x_13; +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) { -lean_object* x_58; lean_object* x_59; uint8_t x_60; -x_58 = lean_ctor_get(x_57, 0); -lean_inc(x_58); -lean_dec_ref(x_57); -x_59 = lean_st_ref_get(x_42); -lean_dec(x_42); -lean_dec(x_59); -x_60 = lean_unbox(x_58); -lean_dec(x_58); -x_12 = x_60; -x_13 = lean_box(0); -goto block_18; +return x_12; } else { -lean_dec(x_42); -if (lean_obj_tag(x_57) == 0) -{ -lean_object* x_61; uint8_t x_62; -x_61 = lean_ctor_get(x_57, 0); -lean_inc(x_61); -lean_dec_ref(x_57); -x_62 = lean_unbox(x_61); -lean_dec(x_61); -x_12 = x_62; -x_13 = lean_box(0); -goto block_18; +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_12, 0); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_15, 0, x_14); +return x_15; +} } else { -uint8_t x_63; -lean_dec(x_11); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec_ref(x_1); -x_63 = !lean_is_exclusive(x_57); -if (x_63 == 0) +uint8_t x_16; +x_16 = !lean_is_exclusive(x_12); +if (x_16 == 0) { -return x_57; +return x_12; } else { -lean_object* x_64; lean_object* x_65; -x_64 = lean_ctor_get(x_57, 0); -lean_inc(x_64); -lean_dec(x_57); -x_65 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_65, 0, x_64); -return x_65; +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_12, 0); +lean_inc(x_17); +lean_dec(x_12); +x_18 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_18, 0, x_17); +return x_18; } } } } -else -{ -lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; uint8_t x_74; lean_object* x_75; lean_object* x_76; -x_66 = lean_ctor_get(x_43, 2); -lean_inc(x_66); -lean_dec(x_43); -x_67 = lean_box(x_29); -x_68 = lean_alloc_closure((void*)(l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___lam__0___boxed), 9, 2); -lean_closure_set(x_68, 0, x_67); -lean_closure_set(x_68, 1, x_44); -x_69 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__10; -x_70 = lean_box(1); -x_71 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_71, 0, x_34); -lean_ctor_set(x_71, 1, x_39); -lean_ctor_set(x_71, 2, x_70); -x_72 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__11; -x_73 = lean_box(0); -x_74 = 1; -lean_inc(x_3); -lean_inc(x_2); -x_75 = lean_alloc_ctor(0, 7, 4); -lean_ctor_set(x_75, 0, x_69); -lean_ctor_set(x_75, 1, x_2); -lean_ctor_set(x_75, 2, x_71); -lean_ctor_set(x_75, 3, x_72); -lean_ctor_set(x_75, 4, x_73); -lean_ctor_set(x_75, 5, x_3); -lean_ctor_set(x_75, 6, x_73); -lean_ctor_set_uint8(x_75, sizeof(void*)*7, x_29); -lean_ctor_set_uint8(x_75, sizeof(void*)*7 + 1, x_29); -lean_ctor_set_uint8(x_75, sizeof(void*)*7 + 2, x_29); -lean_ctor_set_uint8(x_75, sizeof(void*)*7 + 3, x_74); -lean_inc(x_7); -lean_inc_ref(x_6); -lean_inc(x_42); -x_76 = l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2___redArg(x_66, x_68, x_29, x_75, x_42, x_6, x_7); -if (lean_obj_tag(x_76) == 0) +LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -lean_object* x_77; lean_object* x_78; uint8_t x_79; -x_77 = lean_ctor_get(x_76, 0); -lean_inc(x_77); -lean_dec_ref(x_76); -x_78 = lean_st_ref_get(x_42); -lean_dec(x_42); -lean_dec(x_78); -x_79 = lean_unbox(x_77); -lean_dec(x_77); -x_12 = x_79; -x_13 = lean_box(0); -goto block_18; +lean_object* x_10; +x_10 = l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_10; } -else +} +static lean_object* _init_l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__2() { +_start: { -lean_dec(x_42); -if (lean_obj_tag(x_76) == 0) -{ -lean_object* x_80; uint8_t x_81; -x_80 = lean_ctor_get(x_76, 0); -lean_inc(x_80); -lean_dec_ref(x_76); -x_81 = lean_unbox(x_80); -lean_dec(x_80); -x_12 = x_81; -x_13 = lean_box(0); -goto block_18; -} -else -{ -lean_object* x_82; lean_object* x_83; lean_object* x_84; -lean_dec(x_11); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec_ref(x_1); -x_82 = lean_ctor_get(x_76, 0); -lean_inc(x_82); -if (lean_is_exclusive(x_76)) { - lean_ctor_release(x_76, 0); - x_83 = x_76; -} else { - lean_dec_ref(x_76); - x_83 = lean_box(0); -} -if (lean_is_scalar(x_83)) { - x_84 = lean_alloc_ctor(1, 1, 0); -} else { - x_84 = x_83; -} -lean_ctor_set(x_84, 0, x_82); -return x_84; -} -} -} -} -else -{ -lean_dec(x_31); -lean_inc(x_7); -lean_inc_ref(x_6); -x_19 = x_6; -x_20 = x_7; -x_21 = lean_box(0); -goto block_28; -} -} -else -{ -lean_dec(x_30); -lean_inc(x_7); -lean_inc_ref(x_6); -x_19 = x_6; -x_20 = x_7; -x_21 = lean_box(0); -goto block_28; -} -block_18: -{ -if (x_12 == 0) -{ -lean_object* x_14; lean_object* x_15; -x_14 = lean_unsigned_to_nat(1u); -x_15 = lean_nat_add(x_5, x_14); -lean_dec(x_5); -x_4 = x_11; -x_5 = x_15; -goto _start; -} -else -{ -x_4 = x_11; -goto _start; -} -} -block_28: -{ -lean_object* x_22; lean_object* x_23; -x_22 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__3; -x_23 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__0(x_22, x_19, x_20); -if (lean_obj_tag(x_23) == 0) -{ -lean_dec_ref(x_23); -x_4 = x_11; -goto _start; -} -else -{ -uint8_t x_25; -lean_dec(x_11); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec_ref(x_1); -x_25 = !lean_is_exclusive(x_23); -if (x_25 == 0) -{ -return x_23; -} -else -{ -lean_object* x_26; lean_object* x_27; -x_26 = lean_ctor_get(x_23, 0); -lean_inc(x_26); -lean_dec(x_23); -x_27 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_27, 0, x_26); -return x_27; -} -} -} -} +lean_object* x_1; +x_1 = lean_mk_string_unchecked("unreachable code has been reached", 33, 33); +return x_1; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21, lean_object* x_22, lean_object* x_23, lean_object* x_24, lean_object* x_25, lean_object* x_26, lean_object* x_27) { +static lean_object* _init_l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__1() { _start: { -lean_object* x_29; -x_29 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg(x_1, x_2, x_9, x_23, x_24, x_26, x_27); -return x_29; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.IR.nameToIRType.fillCache", 30, 30); +return x_1; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +static lean_object* _init_l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__0() { _start: { -lean_object* x_13; -x_13 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__1___redArg(x_1, x_2, x_5, x_6, x_8, x_9, x_10, x_11); -return x_13; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Compiler.IR.ToIRType", 25, 25); +return x_1; } } -static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__0() { +static lean_object* _init_l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__3() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("UInt8", 5, 5); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__2; +x_2 = lean_unsigned_to_nat(62u); +x_3 = lean_unsigned_to_nat(111u); +x_4 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__1; +x_5 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__0; +x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); +return x_6; } } -static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__1() { +static lean_object* _init_l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__0___closed__0() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("UInt16", 6, 6); +x_1 = lean_alloc_closure((void*)(l_Lean_Core_instInhabitedCoreM___lam__0___boxed), 3, 0); return x_1; } } -static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__2() { +LEAN_EXPORT lean_object* l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("UInt32", 6, 6); -return x_1; +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__0___closed__0; +x_6 = lean_panic_fn(x_5, x_1); +x_7 = lean_apply_3(x_6, x_2, x_3, lean_box(0)); +return x_7; } } -static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__3() { +static lean_object* _init_l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("UInt64", 6, 6); +x_1 = l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_box(0), lean_box(0)); return x_1; } } -static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__4() { +static lean_object* _init_l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__5() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("USize", 5, 5); -return x_1; +lean_object* x_1; lean_object* x_2; +x_1 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__4; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } -static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__5() { +static lean_object* _init_l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__6() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Float", 5, 5); -return x_1; +lean_object* x_1; lean_object* x_2; +x_1 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__5; +x_2 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +lean_ctor_set(x_2, 2, x_1); +lean_ctor_set(x_2, 3, x_1); +lean_ctor_set(x_2, 4, x_1); +lean_ctor_set(x_2, 5, x_1); +return x_2; } } -static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__6() { +static lean_object* _init_l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__7() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Float32", 7, 7); -return x_1; +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(32u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; } } -static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__7() { +static lean_object* _init_l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__8() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("lcErased", 8, 8); -return x_1; +lean_object* x_1; lean_object* x_2; +x_1 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__7; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } -static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__8() { +static lean_object* _init_l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__9() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Int", 3, 3); -return x_1; +lean_object* x_1; lean_object* x_2; +x_1 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__5; +x_2 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +lean_ctor_set(x_2, 2, x_1); +lean_ctor_set(x_2, 3, x_1); +lean_ctor_set(x_2, 4, x_1); +return x_2; } } -static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__9() { +static lean_object* _init_l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___lam__0___closed__0() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("lcVoid", 6, 6); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = lean_box(0); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__1___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_5; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_box(1); -if (lean_obj_tag(x_1) == 1) +uint8_t x_10; +x_10 = !lean_is_exclusive(x_3); +if (x_10 == 0) { -lean_object* x_44; -x_44 = lean_ctor_get(x_1, 0); -if (lean_obj_tag(x_44) == 0) +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); +x_14 = lean_nat_dec_lt(x_12, x_13); +if (x_14 == 0) { -lean_object* x_45; lean_object* x_46; uint8_t x_47; -x_45 = lean_ctor_get(x_1, 1); -x_46 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__0; -x_47 = lean_string_dec_eq(x_45, x_46); -if (x_47 == 0) -{ -lean_object* x_48; uint8_t x_49; -x_48 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__1; -x_49 = lean_string_dec_eq(x_45, x_48); -if (x_49 == 0) -{ -lean_object* x_50; uint8_t x_51; -x_50 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__2; -x_51 = lean_string_dec_eq(x_45, x_50); -if (x_51 == 0) -{ -lean_object* x_52; uint8_t x_53; -x_52 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__3; -x_53 = lean_string_dec_eq(x_45, x_52); -if (x_53 == 0) -{ -lean_object* x_54; uint8_t x_55; -x_54 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__4; -x_55 = lean_string_dec_eq(x_45, x_54); -if (x_55 == 0) +lean_object* x_15; +lean_free_object(x_3); +lean_dec(x_13); +lean_dec(x_12); +lean_dec_ref(x_11); +lean_dec(x_8); +lean_dec_ref(x_7); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec_ref(x_2); +x_15 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_15, 0, x_4); +return x_15; +} +else { -lean_object* x_56; uint8_t x_57; -x_56 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__5; -x_57 = lean_string_dec_eq(x_45, x_56); -if (x_57 == 0) +lean_object* x_16; lean_object* x_17; lean_object* x_18; +lean_dec_ref(x_4); +x_16 = lean_array_fget_borrowed(x_11, x_12); +x_17 = l_Lean_Expr_fvarId_x21(x_16); +lean_inc_ref(x_5); +x_18 = l_Lean_FVarId_getType___redArg(x_17, x_5, x_7, x_8); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_58; uint8_t x_59; -x_58 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__6; -x_59 = lean_string_dec_eq(x_45, x_58); -if (x_59 == 0) +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +lean_dec_ref(x_18); +lean_inc(x_8); +lean_inc_ref(x_7); +lean_inc(x_6); +lean_inc_ref(x_5); +x_20 = l_Lean_Compiler_LCNF_toLCNFType(x_19, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_60; uint8_t x_61; -x_60 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__7; -x_61 = lean_string_dec_eq(x_45, x_60); -if (x_61 == 0) +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +lean_dec_ref(x_20); +lean_inc(x_8); +lean_inc_ref(x_7); +x_22 = l_Lean_Compiler_LCNF_toMonoType(x_21, x_7, x_8); +if (lean_obj_tag(x_22) == 0) { -lean_object* x_62; uint8_t x_63; -x_62 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__8; -x_63 = lean_string_dec_eq(x_45, x_62); -if (x_63 == 0) +uint8_t x_23; +x_23 = !lean_is_exclusive(x_22); +if (x_23 == 0) { -lean_object* x_64; uint8_t x_65; -x_64 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__9; -x_65 = lean_string_dec_eq(x_45, x_64); -if (x_65 == 0) +lean_object* x_24; uint8_t x_25; +x_24 = lean_ctor_get(x_22, 0); +x_25 = l_Lean_Expr_isErased(x_24); +lean_dec(x_24); +if (x_25 == 0) { -x_10 = x_2; -x_11 = x_3; -x_12 = lean_box(0); -goto block_43; +lean_object* x_26; lean_object* x_27; lean_object* x_28; +lean_free_object(x_3); +lean_dec(x_13); +lean_dec(x_12); +lean_dec_ref(x_11); +lean_dec(x_8); +lean_dec_ref(x_7); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec_ref(x_2); +x_26 = lean_box(x_14); +x_27 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_27, 0, x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_1); +lean_ctor_set(x_22, 0, x_28); +return x_22; } else { -lean_object* x_66; lean_object* x_67; -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_66 = lean_box(13); -x_67 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_67, 0, x_66); -return x_67; +lean_object* x_29; lean_object* x_30; +lean_free_object(x_22); +x_29 = lean_unsigned_to_nat(1u); +x_30 = lean_nat_add(x_12, x_29); +lean_dec(x_12); +lean_ctor_set(x_3, 1, x_30); +lean_inc_ref(x_2); +{ +lean_object* _tmp_3 = x_2; +x_4 = _tmp_3; +} +goto _start; } } else { -lean_object* x_68; lean_object* x_69; -lean_dec(x_3); +lean_object* x_32; uint8_t x_33; +x_32 = lean_ctor_get(x_22, 0); +lean_inc(x_32); +lean_dec(x_22); +x_33 = l_Lean_Expr_isErased(x_32); +lean_dec(x_32); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +lean_free_object(x_3); +lean_dec(x_13); +lean_dec(x_12); +lean_dec_ref(x_11); +lean_dec(x_8); +lean_dec_ref(x_7); +lean_dec(x_6); +lean_dec_ref(x_5); lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_68 = lean_box(8); -x_69 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_69, 0, x_68); -return x_69; -} +x_34 = lean_box(x_14); +x_35 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_35, 0, x_34); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_1); +x_37 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_37, 0, x_36); +return x_37; } else { -lean_object* x_70; lean_object* x_71; -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_70 = lean_box(6); -x_71 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_71, 0, x_70); -return x_71; +lean_object* x_38; lean_object* x_39; +x_38 = lean_unsigned_to_nat(1u); +x_39 = lean_nat_add(x_12, x_38); +lean_dec(x_12); +lean_ctor_set(x_3, 1, x_39); +lean_inc_ref(x_2); +{ +lean_object* _tmp_3 = x_2; +x_4 = _tmp_3; } +goto _start; } -else -{ -lean_object* x_72; lean_object* x_73; -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_72 = lean_box(9); -x_73 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_73, 0, x_72); -return x_73; } } else { -lean_object* x_74; lean_object* x_75; -lean_dec(x_3); +uint8_t x_41; +lean_free_object(x_3); +lean_dec(x_13); +lean_dec(x_12); +lean_dec_ref(x_11); +lean_dec(x_8); +lean_dec_ref(x_7); +lean_dec(x_6); +lean_dec_ref(x_5); lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_74 = lean_box(0); -x_75 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_75, 0, x_74); -return x_75; -} +x_41 = !lean_is_exclusive(x_22); +if (x_41 == 0) +{ +return x_22; } else { -lean_object* x_76; lean_object* x_77; -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_76 = lean_box(5); -x_77 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_77, 0, x_76); -return x_77; +lean_object* x_42; lean_object* x_43; +x_42 = lean_ctor_get(x_22, 0); +lean_inc(x_42); +lean_dec(x_22); +x_43 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_43, 0, x_42); +return x_43; +} } } else { -lean_object* x_78; lean_object* x_79; -lean_dec(x_3); +uint8_t x_44; +lean_free_object(x_3); +lean_dec(x_13); +lean_dec(x_12); +lean_dec_ref(x_11); +lean_dec(x_8); +lean_dec_ref(x_7); +lean_dec(x_6); +lean_dec_ref(x_5); lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_78 = lean_box(4); -x_79 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_79, 0, x_78); -return x_79; -} +x_44 = !lean_is_exclusive(x_20); +if (x_44 == 0) +{ +return x_20; } else { -lean_object* x_80; lean_object* x_81; -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_80 = lean_box(3); -x_81 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_81, 0, x_80); -return x_81; +lean_object* x_45; lean_object* x_46; +x_45 = lean_ctor_get(x_20, 0); +lean_inc(x_45); +lean_dec(x_20); +x_46 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_46, 0, x_45); +return x_46; +} } } else { -lean_object* x_82; lean_object* x_83; -lean_dec(x_3); +uint8_t x_47; +lean_free_object(x_3); +lean_dec(x_13); +lean_dec(x_12); +lean_dec_ref(x_11); +lean_dec(x_8); +lean_dec_ref(x_7); +lean_dec(x_6); +lean_dec_ref(x_5); lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_82 = lean_box(2); -x_83 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_83, 0, x_82); -return x_83; -} +x_47 = !lean_is_exclusive(x_18); +if (x_47 == 0) +{ +return x_18; } else { -lean_object* x_84; lean_object* x_85; -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_84 = lean_box(1); -x_85 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_85, 0, x_84); -return x_85; +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_18, 0); +lean_inc(x_48); +lean_dec(x_18); +x_49 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_49, 0, x_48); +return x_49; } } -else -{ -x_10 = x_2; -x_11 = x_3; -x_12 = lean_box(0); -goto block_43; } } else { -x_10 = x_2; -x_11 = x_3; -x_12 = lean_box(0); -goto block_43; -} -block_8: +lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; +x_50 = lean_ctor_get(x_3, 0); +x_51 = lean_ctor_get(x_3, 1); +x_52 = lean_ctor_get(x_3, 2); +lean_inc(x_52); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_3); +x_53 = lean_nat_dec_lt(x_51, x_52); +if (x_53 == 0) { -lean_object* x_6; lean_object* x_7; -x_6 = lean_box(8); -x_7 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_7, 0, x_6); -return x_7; +lean_object* x_54; +lean_dec(x_52); +lean_dec(x_51); +lean_dec_ref(x_50); +lean_dec(x_8); +lean_dec_ref(x_7); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec_ref(x_2); +x_54 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_54, 0, x_4); +return x_54; } -block_43: -{ -lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; -x_13 = lean_st_ref_get(x_11); -x_14 = lean_ctor_get(x_13, 0); -lean_inc_ref(x_14); -lean_dec(x_13); -x_15 = 0; -lean_inc_ref(x_14); -x_16 = l_Lean_Environment_find_x3f(x_14, x_1, x_15); -if (lean_obj_tag(x_16) == 1) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec_ref(x_16); -if (lean_obj_tag(x_17) == 5) -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_18 = lean_ctor_get(x_17, 0); -lean_inc_ref(x_18); -lean_dec_ref(x_17); -x_19 = lean_ctor_get(x_18, 4); -lean_inc(x_19); -lean_dec_ref(x_18); -x_20 = lean_unsigned_to_nat(0u); -lean_inc(x_19); -x_21 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg(x_14, x_9, x_20, x_19, x_20, x_10, x_11); -if (lean_obj_tag(x_21) == 0) -{ -uint8_t x_22; -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) +else { -lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_23 = lean_ctor_get(x_21, 0); -x_24 = l_List_lengthTR___redArg(x_19); -lean_dec(x_19); -x_25 = lean_nat_dec_eq(x_23, x_24); -if (x_25 == 0) +lean_object* x_55; lean_object* x_56; lean_object* x_57; +lean_dec_ref(x_4); +x_55 = lean_array_fget_borrowed(x_50, x_51); +x_56 = l_Lean_Expr_fvarId_x21(x_55); +lean_inc_ref(x_5); +x_57 = l_Lean_FVarId_getType___redArg(x_56, x_5, x_7, x_8); +if (lean_obj_tag(x_57) == 0) { -uint8_t x_26; -lean_dec(x_24); -x_26 = lean_nat_dec_eq(x_23, x_20); -lean_dec(x_23); -if (x_26 == 0) +lean_object* x_58; lean_object* x_59; +x_58 = lean_ctor_get(x_57, 0); +lean_inc(x_58); +lean_dec_ref(x_57); +lean_inc(x_8); +lean_inc_ref(x_7); +lean_inc(x_6); +lean_inc_ref(x_5); +x_59 = l_Lean_Compiler_LCNF_toLCNFType(x_58, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_59) == 0) { -lean_object* x_27; -x_27 = lean_box(8); -lean_ctor_set(x_21, 0, x_27); -return x_21; -} -else +lean_object* x_60; lean_object* x_61; +x_60 = lean_ctor_get(x_59, 0); +lean_inc(x_60); +lean_dec_ref(x_59); +lean_inc(x_8); +lean_inc_ref(x_7); +x_61 = l_Lean_Compiler_LCNF_toMonoType(x_60, x_7, x_8); +if (lean_obj_tag(x_61) == 0) { -lean_object* x_28; -x_28 = lean_box(7); -lean_ctor_set(x_21, 0, x_28); -return x_21; -} +lean_object* x_62; lean_object* x_63; uint8_t x_64; +x_62 = lean_ctor_get(x_61, 0); +lean_inc(x_62); +if (lean_is_exclusive(x_61)) { + lean_ctor_release(x_61, 0); + x_63 = x_61; +} else { + lean_dec_ref(x_61); + x_63 = lean_box(0); } -else +x_64 = l_Lean_Expr_isErased(x_62); +lean_dec(x_62); +if (x_64 == 0) { -lean_object* x_29; -lean_dec(x_23); -x_29 = l_Lean_IR_irTypeForEnum(x_24); -lean_dec(x_24); -lean_ctor_set(x_21, 0, x_29); -return x_21; +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +lean_dec(x_52); +lean_dec(x_51); +lean_dec_ref(x_50); +lean_dec(x_8); +lean_dec_ref(x_7); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec_ref(x_2); +x_65 = lean_box(x_53); +x_66 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_66, 0, x_65); +x_67 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_1); +if (lean_is_scalar(x_63)) { + x_68 = lean_alloc_ctor(0, 1, 0); +} else { + x_68 = x_63; } +lean_ctor_set(x_68, 0, x_67); +return x_68; } else { -lean_object* x_30; lean_object* x_31; uint8_t x_32; -x_30 = lean_ctor_get(x_21, 0); -lean_inc(x_30); -lean_dec(x_21); -x_31 = l_List_lengthTR___redArg(x_19); -lean_dec(x_19); -x_32 = lean_nat_dec_eq(x_30, x_31); -if (x_32 == 0) -{ -uint8_t x_33; -lean_dec(x_31); -x_33 = lean_nat_dec_eq(x_30, x_20); -lean_dec(x_30); -if (x_33 == 0) +lean_object* x_69; lean_object* x_70; lean_object* x_71; +lean_dec(x_63); +x_69 = lean_unsigned_to_nat(1u); +x_70 = lean_nat_add(x_51, x_69); +lean_dec(x_51); +x_71 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_71, 0, x_50); +lean_ctor_set(x_71, 1, x_70); +lean_ctor_set(x_71, 2, x_52); +lean_inc_ref(x_2); { -lean_object* x_34; lean_object* x_35; -x_34 = lean_box(8); -x_35 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_35, 0, x_34); -return x_35; +lean_object* _tmp_2 = x_71; +lean_object* _tmp_3 = x_2; +x_3 = _tmp_2; +x_4 = _tmp_3; } -else -{ -lean_object* x_36; lean_object* x_37; -x_36 = lean_box(7); -x_37 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_37, 0, x_36); -return x_37; +goto _start; } } else { -lean_object* x_38; lean_object* x_39; -lean_dec(x_30); -x_38 = l_Lean_IR_irTypeForEnum(x_31); -lean_dec(x_31); -x_39 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_39, 0, x_38); -return x_39; +lean_object* x_73; lean_object* x_74; lean_object* x_75; +lean_dec(x_52); +lean_dec(x_51); +lean_dec_ref(x_50); +lean_dec(x_8); +lean_dec_ref(x_7); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec_ref(x_2); +x_73 = lean_ctor_get(x_61, 0); +lean_inc(x_73); +if (lean_is_exclusive(x_61)) { + lean_ctor_release(x_61, 0); + x_74 = x_61; +} else { + lean_dec_ref(x_61); + x_74 = lean_box(0); } +if (lean_is_scalar(x_74)) { + x_75 = lean_alloc_ctor(1, 1, 0); +} else { + x_75 = x_74; } +lean_ctor_set(x_75, 0, x_73); +return x_75; } -else -{ -uint8_t x_40; -lean_dec(x_19); -x_40 = !lean_is_exclusive(x_21); -if (x_40 == 0) -{ -return x_21; } else { -lean_object* x_41; lean_object* x_42; -x_41 = lean_ctor_get(x_21, 0); -lean_inc(x_41); -lean_dec(x_21); -x_42 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_42, 0, x_41); -return x_42; -} +lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_52); +lean_dec(x_51); +lean_dec_ref(x_50); +lean_dec(x_8); +lean_dec_ref(x_7); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec_ref(x_2); +x_76 = lean_ctor_get(x_59, 0); +lean_inc(x_76); +if (lean_is_exclusive(x_59)) { + lean_ctor_release(x_59, 0); + x_77 = x_59; +} else { + lean_dec_ref(x_59); + x_77 = lean_box(0); } +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 1, 0); +} else { + x_78 = x_77; } -else -{ -lean_dec(x_17); -lean_dec_ref(x_14); -lean_dec(x_11); -lean_dec_ref(x_10); -x_5 = lean_box(0); -goto block_8; +lean_ctor_set(x_78, 0, x_76); +return x_78; } } else { -lean_dec(x_16); -lean_dec_ref(x_14); -lean_dec(x_11); -lean_dec_ref(x_10); -x_5 = lean_box(0); -goto block_8; -} +lean_object* x_79; lean_object* x_80; lean_object* x_81; +lean_dec(x_52); +lean_dec(x_51); +lean_dec_ref(x_50); +lean_dec(x_8); +lean_dec_ref(x_7); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec_ref(x_2); +x_79 = lean_ctor_get(x_57, 0); +lean_inc(x_79); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + x_80 = x_57; +} else { + lean_dec_ref(x_57); + x_80 = lean_box(0); } +if (lean_is_scalar(x_80)) { + x_81 = lean_alloc_ctor(1, 1, 0); +} else { + x_81 = x_80; } +lean_ctor_set(x_81, 0, x_79); +return x_81; } -LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -uint8_t x_10; lean_object* x_11; -x_10 = lean_unbox(x_4); -x_11 = l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2(x_1, x_2, x_3, x_10, x_5, x_6, x_7, x_8); -return x_11; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; -lean_object* x_18 = _args[17]; -lean_object* x_19 = _args[18]; -lean_object* x_20 = _args[19]; -lean_object* x_21 = _args[20]; -lean_object* x_22 = _args[21]; -lean_object* x_23 = _args[22]; -lean_object* x_24 = _args[23]; -lean_object* x_25 = _args[24]; -lean_object* x_26 = _args[25]; -lean_object* x_27 = _args[26]; -lean_object* x_28 = _args[27]; -_start: -{ -lean_object* x_29; -x_29 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_27); -lean_dec(x_22); -lean_dec_ref(x_21); -lean_dec_ref(x_20); -lean_dec_ref(x_19); -lean_dec_ref(x_18); -lean_dec_ref(x_17); -lean_dec_ref(x_16); -lean_dec_ref(x_15); -lean_dec_ref(x_14); -lean_dec_ref(x_13); -lean_dec_ref(x_12); -lean_dec_ref(x_11); -lean_dec_ref(x_10); -lean_dec_ref(x_8); -lean_dec_ref(x_7); -lean_dec_ref(x_6); -lean_dec_ref(x_5); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -return x_29; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___lam__0(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_13; -x_13 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_13; -} -} -LEAN_EXPORT lean_object* l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +lean_object* x_10; lean_object* x_11; lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_31 = lean_unsigned_to_nat(0u); +x_32 = lean_array_get_size(x_3); +x_33 = lean_nat_dec_le(x_2, x_31); +if (x_33 == 0) { -lean_object* x_5; -x_5 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__0(x_1, x_2, x_3); -return x_5; -} +x_10 = x_2; +x_11 = x_32; +goto block_30; } -LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: +else { -uint8_t x_9; lean_object* x_10; -x_9 = lean_unbox(x_3); -x_10 = l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2___redArg(x_1, x_2, x_9, x_4, x_5, x_6, x_7); -return x_10; -} +lean_dec(x_2); +x_10 = x_31; +x_11 = x_32; +goto block_30; } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__1___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: +block_30: { -lean_object* x_10; -x_10 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__1___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_10; -} -} -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = l_Array_toSubarray___redArg(x_3, x_10, x_11); +x_13 = lean_box(0); +x_14 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___lam__0___closed__0; +x_15 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__1___redArg(x_13, x_14, x_12, x_14, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_9; -x_9 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -return x_9; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +uint8_t x_16; +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) { -lean_object* x_5; -x_5 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache(x_1, x_2, x_3); -return x_5; -} -} -static lean_object* _init_l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg___closed__1() { -_start: +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_15, 0); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +lean_dec(x_17); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Name_hash___override___boxed), 1, 0); -return x_1; -} +lean_object* x_19; +x_19 = lean_box(x_1); +lean_ctor_set(x_15, 0, x_19); +return x_15; } -static lean_object* _init_l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg___closed__0() { -_start: +else { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Name_beq___boxed), 2, 0); -return x_1; +lean_object* x_20; +x_20 = lean_ctor_get(x_18, 0); +lean_inc(x_20); +lean_dec_ref(x_18); +lean_ctor_set(x_15, 0, x_20); +return x_15; } } -static lean_object* _init_l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg___closed__1; -x_2 = l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg___closed__0; -x_3 = l_Lean_PersistentHashMap_instInhabited(lean_box(0), lean_box(0), x_2, x_1); -return x_3; -} -} -static lean_object* _init_l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg___closed__3() { -_start: +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_15, 0); +lean_inc(x_21); +lean_dec(x_15); +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +lean_dec(x_21); +if (lean_obj_tag(x_22) == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg___closed__2; -x_2 = lean_box(0); -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; -} +lean_object* x_23; lean_object* x_24; +x_23 = lean_box(x_1); +x_24 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_24, 0, x_23); +return x_24; } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_5 = lean_st_ref_get(x_3); -x_6 = lean_ctor_get(x_5, 0); -lean_inc_ref(x_6); -lean_dec(x_5); -x_7 = lean_ctor_get(x_1, 2); -x_8 = l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg___closed__3; -x_9 = lean_box(0); -x_10 = l___private_Lean_Environment_0__Lean_EnvExtension_getStateUnsafe___redArg(x_8, x_1, x_6, x_7, x_9); -x_11 = lean_ctor_get(x_10, 1); -lean_inc(x_11); -lean_dec(x_10); -x_12 = l_Lean_PersistentHashMap_find_x3f___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__1___redArg(x_11, x_2); -x_13 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_13, 0, x_12); -return x_13; -} +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_22, 0); +lean_inc(x_25); +lean_dec_ref(x_22); +x_26 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_26, 0, x_25); +return x_26; } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_8; -x_8 = l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg(x_3, x_4, x_6); -return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -uint8_t x_4; -x_4 = !lean_is_exclusive(x_3); -if (x_4 == 0) +uint8_t x_27; +x_27 = !lean_is_exclusive(x_15); +if (x_27 == 0) { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_5 = lean_ctor_get(x_3, 0); -x_6 = lean_ctor_get(x_3, 1); -lean_inc(x_1); -x_7 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_7, 0, x_1); -lean_ctor_set(x_7, 1, x_5); -x_8 = l_Lean_PersistentHashMap_insert___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__0___redArg(x_6, x_1, x_2); -lean_ctor_set(x_3, 1, x_8); -lean_ctor_set(x_3, 0, x_7); -return x_3; +return x_15; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_9 = lean_ctor_get(x_3, 0); -x_10 = lean_ctor_get(x_3, 1); -lean_inc(x_10); -lean_inc(x_9); -lean_dec(x_3); -lean_inc(x_1); -x_11 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_11, 0, x_1); -lean_ctor_set(x_11, 1, x_9); -x_12 = l_Lean_PersistentHashMap_insert___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__0___redArg(x_10, x_1, x_2); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_11); -lean_ctor_set(x_13, 1, x_12); -return x_13; +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_15, 0); +lean_inc(x_28); +lean_dec(x_15); +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, x_28); +return x_29; } } } -static lean_object* _init_l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg___closed__0() { +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_1; -x_1 = l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_box(0), lean_box(0)); -return x_1; +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_1); +x_11 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___lam__0(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec_ref(x_4); +return x_11; } } -static lean_object* _init_l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg___closed__1() { +static lean_object* _init_l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__10() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg___closed__0; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_1; +x_1 = l_Lean_Meta_instInhabitedConfigWithKey___private__1; +return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg___closed__2() { +static lean_object* _init_l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg___closed__1; -x_2 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_2, 0, x_1); -lean_ctor_set(x_2, 1, x_1); +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_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_6; uint8_t x_7; -x_6 = lean_st_ref_take(x_4); -x_7 = !lean_is_exclusive(x_6); -if (x_7 == 0) -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_8 = lean_ctor_get(x_6, 0); -x_9 = lean_ctor_get(x_6, 5); -lean_dec(x_9); -x_10 = lean_ctor_get(x_1, 2); -lean_inc(x_10); -x_11 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg___lam__0), 3, 2); -lean_closure_set(x_11, 0, x_2); -lean_closure_set(x_11, 1, x_3); -x_12 = lean_box(0); -x_13 = l_Lean_EnvExtension_modifyState___redArg(x_1, x_8, x_11, x_10, x_12); -lean_dec(x_10); -x_14 = l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg___closed__2; -lean_ctor_set(x_6, 5, x_14); -lean_ctor_set(x_6, 0, x_13); -x_15 = lean_st_ref_set(x_4, x_6); -x_16 = lean_box(0); -x_17 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_17, 0, x_16); -return x_17; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_18 = lean_ctor_get(x_6, 0); -x_19 = lean_ctor_get(x_6, 1); -x_20 = lean_ctor_get(x_6, 2); -x_21 = lean_ctor_get(x_6, 3); -x_22 = lean_ctor_get(x_6, 4); -x_23 = lean_ctor_get(x_6, 6); -x_24 = lean_ctor_get(x_6, 7); -x_25 = lean_ctor_get(x_6, 8); -lean_inc(x_25); -lean_inc(x_24); -lean_inc(x_23); -lean_inc(x_22); -lean_inc(x_21); -lean_inc(x_20); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_6); -x_26 = lean_ctor_get(x_1, 2); -lean_inc(x_26); -x_27 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg___lam__0), 3, 2); -lean_closure_set(x_27, 0, x_2); -lean_closure_set(x_27, 1, x_3); -x_28 = lean_box(0); -x_29 = l_Lean_EnvExtension_modifyState___redArg(x_1, x_18, x_27, x_26, x_28); -lean_dec(x_26); -x_30 = l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg___closed__2; -x_31 = lean_alloc_ctor(0, 9, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_19); -lean_ctor_set(x_31, 2, x_20); -lean_ctor_set(x_31, 3, x_21); -lean_ctor_set(x_31, 4, x_22); -lean_ctor_set(x_31, 5, x_30); -lean_ctor_set(x_31, 6, x_23); -lean_ctor_set(x_31, 7, x_24); -lean_ctor_set(x_31, 8, x_25); -x_32 = lean_st_ref_set(x_4, x_31); -x_33 = lean_box(0); -x_34 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_34, 0, x_33); -return x_34; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: +if (lean_obj_tag(x_4) == 0) { lean_object* x_9; -x_9 = l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg(x_3, x_4, x_5, x_7); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec_ref(x_1); +x_9 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_9, 0, x_5); return x_9; } -} -static lean_object* _init_l_Lean_IR_nameToIRType___closed__0() { -_start: +else { -lean_object* x_1; -x_1 = l_Lean_IR_irTypeExt; -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_IR_nameToIRType(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_29; lean_object* x_30; +x_10 = lean_ctor_get(x_4, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_4, 1); +lean_inc(x_11); +lean_dec_ref(x_4); +x_29 = 0; +lean_inc_ref(x_1); +x_30 = l_Lean_Environment_find_x3f(x_1, x_10, x_29); +if (lean_obj_tag(x_30) == 1) { -lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_5 = l_Lean_IR_nameToIRType___closed__0; -x_6 = l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg(x_5, x_1, x_3); -x_7 = !lean_is_exclusive(x_6); -if (x_7 == 0) +lean_object* x_31; +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +lean_dec_ref(x_30); +if (lean_obj_tag(x_31) == 6) { -lean_object* x_8; -x_8 = lean_ctor_get(x_6, 0); -if (lean_obj_tag(x_8) == 0) +lean_object* x_32; size_t x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; +x_32 = lean_ctor_get(x_31, 0); +lean_inc_ref(x_32); +lean_dec_ref(x_31); +x_33 = 5; +x_34 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__5; +lean_inc_n(x_3, 3); +x_35 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_35, 0, x_3); +lean_ctor_set(x_35, 1, x_3); +lean_ctor_set(x_35, 2, x_3); +lean_ctor_set(x_35, 3, x_34); +lean_ctor_set(x_35, 4, x_34); +lean_ctor_set(x_35, 5, x_34); +lean_ctor_set(x_35, 6, x_34); +lean_ctor_set(x_35, 7, x_34); +lean_ctor_set(x_35, 8, x_34); +x_36 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__6; +x_37 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__7; +x_38 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__8; +lean_inc_n(x_3, 2); +x_39 = lean_alloc_ctor(0, 4, sizeof(size_t)*1); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_37); +lean_ctor_set(x_39, 2, x_3); +lean_ctor_set(x_39, 3, x_3); +lean_ctor_set_usize(x_39, 4, x_33); +x_40 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__9; +lean_inc_ref(x_39); +lean_inc(x_2); +x_41 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_41, 0, x_35); +lean_ctor_set(x_41, 1, x_36); +lean_ctor_set(x_41, 2, x_2); +lean_ctor_set(x_41, 3, x_39); +lean_ctor_set(x_41, 4, x_40); +x_42 = lean_st_mk_ref(x_41); +x_43 = lean_ctor_get(x_32, 0); +lean_inc_ref(x_43); +x_44 = lean_ctor_get(x_32, 3); +lean_inc(x_44); +lean_dec_ref(x_32); +x_45 = !lean_is_exclusive(x_43); +if (x_45 == 0) { -lean_object* x_9; -lean_free_object(x_6); +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; lean_object* x_56; lean_object* x_57; +x_46 = lean_ctor_get(x_43, 2); +x_47 = lean_ctor_get(x_43, 1); +lean_dec(x_47); +x_48 = lean_ctor_get(x_43, 0); +lean_dec(x_48); +x_49 = lean_box(x_29); +x_50 = lean_alloc_closure((void*)(l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___lam__0___boxed), 9, 2); +lean_closure_set(x_50, 0, x_49); +lean_closure_set(x_50, 1, x_44); +x_51 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__10; +x_52 = lean_box(1); +lean_ctor_set(x_43, 2, x_52); +lean_ctor_set(x_43, 1, x_39); +lean_ctor_set(x_43, 0, x_34); +x_53 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__11; +x_54 = lean_box(0); +x_55 = 1; lean_inc(x_3); -lean_inc(x_1); -x_9 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache(x_1, x_2, x_3); -if (lean_obj_tag(x_9) == 0) -{ -lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -lean_dec_ref(x_9); -lean_inc(x_10); -x_11 = l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg(x_5, x_1, x_10, x_3); -lean_dec(x_3); -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) +lean_inc(x_2); +x_56 = lean_alloc_ctor(0, 7, 4); +lean_ctor_set(x_56, 0, x_51); +lean_ctor_set(x_56, 1, x_2); +lean_ctor_set(x_56, 2, x_43); +lean_ctor_set(x_56, 3, x_53); +lean_ctor_set(x_56, 4, x_54); +lean_ctor_set(x_56, 5, x_3); +lean_ctor_set(x_56, 6, x_54); +lean_ctor_set_uint8(x_56, sizeof(void*)*7, x_29); +lean_ctor_set_uint8(x_56, sizeof(void*)*7 + 1, x_29); +lean_ctor_set_uint8(x_56, sizeof(void*)*7 + 2, x_29); +lean_ctor_set_uint8(x_56, sizeof(void*)*7 + 3, x_55); +lean_inc(x_7); +lean_inc_ref(x_6); +lean_inc(x_42); +x_57 = l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2___redArg(x_46, x_50, x_29, x_56, x_42, x_6, x_7); +if (lean_obj_tag(x_57) == 0) { -lean_object* x_13; -x_13 = lean_ctor_get(x_11, 0); -lean_dec(x_13); -lean_ctor_set(x_11, 0, x_10); -return x_11; +lean_object* x_58; lean_object* x_59; uint8_t x_60; +x_58 = lean_ctor_get(x_57, 0); +lean_inc(x_58); +lean_dec_ref(x_57); +x_59 = lean_st_ref_get(x_42); +lean_dec(x_42); +lean_dec(x_59); +x_60 = lean_unbox(x_58); +lean_dec(x_58); +x_12 = x_60; +x_13 = lean_box(0); +goto block_18; } else { -lean_object* x_14; -lean_dec(x_11); -x_14 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_14, 0, x_10); -return x_14; -} +lean_dec(x_42); +if (lean_obj_tag(x_57) == 0) +{ +lean_object* x_61; uint8_t x_62; +x_61 = lean_ctor_get(x_57, 0); +lean_inc(x_61); +lean_dec_ref(x_57); +x_62 = lean_unbox(x_61); +lean_dec(x_61); +x_12 = x_62; +x_13 = lean_box(0); +goto block_18; } else { +uint8_t x_63; +lean_dec(x_11); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec(x_5); lean_dec(x_3); -lean_dec(x_1); -return x_9; -} +lean_dec(x_2); +lean_dec_ref(x_1); +x_63 = !lean_is_exclusive(x_57); +if (x_63 == 0) +{ +return x_57; } else { -lean_object* x_15; -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec(x_1); -x_15 = lean_ctor_get(x_8, 0); -lean_inc(x_15); -lean_dec_ref(x_8); -lean_ctor_set(x_6, 0, x_15); -return x_6; +lean_object* x_64; lean_object* x_65; +x_64 = lean_ctor_get(x_57, 0); +lean_inc(x_64); +lean_dec(x_57); +x_65 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_65, 0, x_64); +return x_65; +} +} } } else { -lean_object* x_16; -x_16 = lean_ctor_get(x_6, 0); -lean_inc(x_16); -lean_dec(x_6); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; uint8_t x_74; lean_object* x_75; lean_object* x_76; +x_66 = lean_ctor_get(x_43, 2); +lean_inc(x_66); +lean_dec(x_43); +x_67 = lean_box(x_29); +x_68 = lean_alloc_closure((void*)(l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___lam__0___boxed), 9, 2); +lean_closure_set(x_68, 0, x_67); +lean_closure_set(x_68, 1, x_44); +x_69 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__10; +x_70 = lean_box(1); +x_71 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_71, 0, x_34); +lean_ctor_set(x_71, 1, x_39); +lean_ctor_set(x_71, 2, x_70); +x_72 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__11; +x_73 = lean_box(0); +x_74 = 1; lean_inc(x_3); -lean_inc(x_1); -x_17 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache(x_1, x_2, x_3); -if (lean_obj_tag(x_17) == 0) +lean_inc(x_2); +x_75 = lean_alloc_ctor(0, 7, 4); +lean_ctor_set(x_75, 0, x_69); +lean_ctor_set(x_75, 1, x_2); +lean_ctor_set(x_75, 2, x_71); +lean_ctor_set(x_75, 3, x_72); +lean_ctor_set(x_75, 4, x_73); +lean_ctor_set(x_75, 5, x_3); +lean_ctor_set(x_75, 6, x_73); +lean_ctor_set_uint8(x_75, sizeof(void*)*7, x_29); +lean_ctor_set_uint8(x_75, sizeof(void*)*7 + 1, x_29); +lean_ctor_set_uint8(x_75, sizeof(void*)*7 + 2, x_29); +lean_ctor_set_uint8(x_75, sizeof(void*)*7 + 3, x_74); +lean_inc(x_7); +lean_inc_ref(x_6); +lean_inc(x_42); +x_76 = l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2___redArg(x_66, x_68, x_29, x_75, x_42, x_6, x_7); +if (lean_obj_tag(x_76) == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -lean_dec_ref(x_17); -lean_inc(x_18); -x_19 = l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg(x_5, x_1, x_18, x_3); -lean_dec(x_3); -if (lean_is_exclusive(x_19)) { - lean_ctor_release(x_19, 0); - x_20 = x_19; -} else { - lean_dec_ref(x_19); - x_20 = lean_box(0); -} -if (lean_is_scalar(x_20)) { - x_21 = lean_alloc_ctor(0, 1, 0); -} else { - x_21 = x_20; -} -lean_ctor_set(x_21, 0, x_18); -return x_21; +lean_object* x_77; lean_object* x_78; uint8_t x_79; +x_77 = lean_ctor_get(x_76, 0); +lean_inc(x_77); +lean_dec_ref(x_76); +x_78 = lean_st_ref_get(x_42); +lean_dec(x_42); +lean_dec(x_78); +x_79 = lean_unbox(x_77); +lean_dec(x_77); +x_12 = x_79; +x_13 = lean_box(0); +goto block_18; } else { -lean_dec(x_3); -lean_dec(x_1); -return x_17; -} -} -else +lean_dec(x_42); +if (lean_obj_tag(x_76) == 0) { -lean_object* x_22; lean_object* x_23; -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec(x_1); -x_22 = lean_ctor_get(x_16, 0); -lean_inc(x_22); -lean_dec_ref(x_16); -x_23 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_23, 0, x_22); -return x_23; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___boxed(lean_object* x_1, lean_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_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -lean_dec(x_2); -return x_8; -} +lean_object* x_80; uint8_t x_81; +x_80 = lean_ctor_get(x_76, 0); +lean_inc(x_80); +lean_dec_ref(x_76); +x_81 = lean_unbox(x_80); +lean_dec(x_80); +x_12 = x_81; +x_13 = lean_box(0); +goto block_18; } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: +else { -lean_object* x_9; -x_9 = l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_object* x_82; lean_object* x_83; lean_object* x_84; +lean_dec(x_11); lean_dec(x_7); lean_dec_ref(x_6); -lean_dec(x_2); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg___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_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg(x_1, x_2, x_3); +lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec_ref(x_1); -return x_5; +x_82 = lean_ctor_get(x_76, 0); +lean_inc(x_82); +if (lean_is_exclusive(x_76)) { + lean_ctor_release(x_76, 0); + x_83 = x_76; +} else { + lean_dec_ref(x_76); + x_83 = lean_box(0); +} +if (lean_is_scalar(x_83)) { + x_84 = lean_alloc_ctor(1, 1, 0); +} else { + x_84 = x_83; +} +lean_ctor_set(x_84, 0, x_82); +return x_84; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +} +} +else { -lean_object* x_6; -x_6 = l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg(x_1, x_2, x_3, x_4); -lean_dec(x_4); -return x_6; +lean_dec(x_31); +lean_inc(x_7); +lean_inc_ref(x_6); +x_19 = x_6; +x_20 = x_7; +x_21 = lean_box(0); +goto block_28; } } -LEAN_EXPORT lean_object* l_Lean_IR_nameToIRType___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +else { -lean_object* x_5; -x_5 = l_Lean_IR_nameToIRType(x_1, x_2, x_3); -return x_5; +lean_dec(x_30); +lean_inc(x_7); +lean_inc_ref(x_6); +x_19 = x_6; +x_20 = x_7; +x_21 = lean_box(0); +goto block_28; } +block_18: +{ +if (x_12 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_unsigned_to_nat(1u); +x_15 = lean_nat_add(x_5, x_14); +lean_dec(x_5); +x_4 = x_11; +x_5 = x_15; +goto _start; } -static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_isAnyProducingType___closed__0() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("lcAny", 5, 5); -return x_1; +x_4 = x_11; +goto _start; } } -LEAN_EXPORT uint8_t l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_isAnyProducingType(lean_object* x_1) { -_start: +block_28: { -switch (lean_obj_tag(x_1)) { -case 4: +lean_object* x_22; lean_object* x_23; +x_22 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__3; +x_23 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__0(x_22, x_19, x_20); +if (lean_obj_tag(x_23) == 0) { -lean_object* x_2; -x_2 = lean_ctor_get(x_1, 0); -if (lean_obj_tag(x_2) == 1) +lean_dec_ref(x_23); +x_4 = x_11; +goto _start; +} +else { -lean_object* x_3; -x_3 = lean_ctor_get(x_2, 0); -if (lean_obj_tag(x_3) == 0) +uint8_t x_25; +lean_dec(x_11); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec_ref(x_1); +x_25 = !lean_is_exclusive(x_23); +if (x_25 == 0) { -lean_object* x_4; lean_object* x_5; uint8_t x_6; -x_4 = lean_ctor_get(x_2, 1); -x_5 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_isAnyProducingType___closed__0; -x_6 = lean_string_dec_eq(x_4, x_5); -return x_6; +return x_23; } else { -uint8_t x_7; -x_7 = 0; -return x_7; +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_23, 0); +lean_inc(x_26); +lean_dec(x_23); +x_27 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_27, 0, x_26); +return x_27; } } -else -{ -uint8_t x_8; -x_8 = 0; -return x_8; } } -case 7: +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21, lean_object* x_22, lean_object* x_23, lean_object* x_24, lean_object* x_25, lean_object* x_26, lean_object* x_27) { +_start: { -lean_object* x_9; -x_9 = lean_ctor_get(x_1, 2); -x_1 = x_9; -goto _start; +lean_object* x_29; +x_29 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg(x_1, x_2, x_9, x_23, x_24, x_26, x_27); +return x_29; } -default: +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -uint8_t x_11; -x_11 = 0; -return x_11; +lean_object* x_13; +x_13 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__1___redArg(x_1, x_2, x_5, x_6, x_8, x_9, x_10, x_11); +return x_13; } } +static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("UInt8", 5, 5); +return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_isAnyProducingType___boxed(lean_object* x_1) { +static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__1() { _start: { -uint8_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_isAnyProducingType(x_1); -lean_dec_ref(x_1); -x_3 = lean_box(x_2); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("UInt16", 6, 6); +return x_1; } } -LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_toIRType_spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__2() { _start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__0___closed__0; -x_6 = lean_panic_fn(x_5, x_1); -x_7 = lean_apply_3(x_6, x_2, x_3, lean_box(0)); -return x_7; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("UInt32", 6, 6); +return x_1; } } -static lean_object* _init_l_Lean_IR_toIRType___closed__0() { +static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__3() { _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_object* x_1; +x_1 = lean_mk_string_unchecked("UInt64", 6, 6); +return x_1; } } -static lean_object* _init_l_Lean_IR_toIRType___closed__1() { +static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__4() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = l_Lean_Expr_sort___override(x_1); -return x_2; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("USize", 5, 5); +return x_1; } } -static lean_object* _init_l_Lean_IR_toIRType___closed__2() { +static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__5() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.IR.toIRType", 16, 16); +x_1 = lean_mk_string_unchecked("Float", 5, 5); return x_1; } } -static lean_object* _init_l_Lean_IR_toIRType___closed__3() { +static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__6() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__2; -x_2 = lean_unsigned_to_nat(41u); -x_3 = lean_unsigned_to_nat(110u); -x_4 = l_Lean_IR_toIRType___closed__2; -x_5 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__0; -x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); -return x_6; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Float32", 7, 7); +return x_1; } } -static lean_object* _init_l_Lean_IR_toIRType___closed__4() { +static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___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_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__2; -x_2 = lean_unsigned_to_nat(9u); -x_3 = lean_unsigned_to_nat(122u); -x_4 = l_Lean_IR_toIRType___closed__2; -x_5 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__0; -x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); -return x_6; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lcErased", 8, 8); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_toIRType(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__8() { _start: { -switch (lean_obj_tag(x_1)) { -case 4: +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Int", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__9() { +_start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = lean_ctor_get(x_1, 0); -lean_inc(x_5); -lean_dec_ref(x_1); -x_6 = l_Lean_IR_toIRType___closed__0; -x_7 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_toIRType_visitApp(x_5, x_6, x_2, x_3); -return x_7; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lcVoid", 6, 6); +return x_1; } -case 5: +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_8; -x_8 = l_Lean_Expr_getAppFn(x_1); -if (lean_obj_tag(x_8) == 4) +lean_object* x_5; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_box(1); +if (lean_obj_tag(x_1) == 1) { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -lean_dec_ref(x_8); -x_10 = l_Lean_IR_toIRType___closed__1; -x_11 = l_Lean_Expr_getAppNumArgs(x_1); -lean_inc(x_11); -x_12 = lean_mk_array(x_11, x_10); -x_13 = lean_unsigned_to_nat(1u); -x_14 = lean_nat_sub(x_11, x_13); -lean_dec(x_11); -x_15 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_12, x_14); -x_16 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_toIRType_visitApp(x_9, x_15, x_2, x_3); -return x_16; +lean_object* x_44; +x_44 = lean_ctor_get(x_1, 0); +if (lean_obj_tag(x_44) == 0) +{ +lean_object* x_45; lean_object* x_46; uint8_t x_47; +x_45 = lean_ctor_get(x_1, 1); +x_46 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__0; +x_47 = lean_string_dec_eq(x_45, x_46); +if (x_47 == 0) +{ +lean_object* x_48; uint8_t x_49; +x_48 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__1; +x_49 = lean_string_dec_eq(x_45, x_48); +if (x_49 == 0) +{ +lean_object* x_50; uint8_t x_51; +x_50 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__2; +x_51 = lean_string_dec_eq(x_45, x_50); +if (x_51 == 0) +{ +lean_object* x_52; uint8_t x_53; +x_52 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__3; +x_53 = lean_string_dec_eq(x_45, x_52); +if (x_53 == 0) +{ +lean_object* x_54; uint8_t x_55; +x_54 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__4; +x_55 = lean_string_dec_eq(x_45, x_54); +if (x_55 == 0) +{ +lean_object* x_56; uint8_t x_57; +x_56 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__5; +x_57 = lean_string_dec_eq(x_45, x_56); +if (x_57 == 0) +{ +lean_object* x_58; uint8_t x_59; +x_58 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__6; +x_59 = lean_string_dec_eq(x_45, x_58); +if (x_59 == 0) +{ +lean_object* x_60; uint8_t x_61; +x_60 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__7; +x_61 = lean_string_dec_eq(x_45, x_60); +if (x_61 == 0) +{ +lean_object* x_62; uint8_t x_63; +x_62 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__8; +x_63 = lean_string_dec_eq(x_45, x_62); +if (x_63 == 0) +{ +lean_object* x_64; uint8_t x_65; +x_64 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__9; +x_65 = lean_string_dec_eq(x_45, x_64); +if (x_65 == 0) +{ +x_10 = x_2; +x_11 = x_3; +x_12 = lean_box(0); +goto block_43; } else { -lean_object* x_17; lean_object* x_18; -lean_dec_ref(x_8); +lean_object* x_66; lean_object* x_67; +lean_dec(x_3); +lean_dec_ref(x_2); lean_dec_ref(x_1); -x_17 = l_Lean_IR_toIRType___closed__3; -x_18 = l_panic___at___00Lean_IR_toIRType_spec__1(x_17, x_2, x_3); -return x_18; +x_66 = lean_box(13); +x_67 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_67, 0, x_66); +return x_67; } } -case 7: +else { -lean_object* x_19; uint8_t x_20; +lean_object* x_68; lean_object* x_69; lean_dec(x_3); lean_dec_ref(x_2); -x_19 = lean_ctor_get(x_1, 2); -lean_inc_ref(x_19); lean_dec_ref(x_1); -x_20 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_isAnyProducingType(x_19); -lean_dec_ref(x_19); -if (x_20 == 0) +x_68 = lean_box(8); +x_69 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_69, 0, x_68); +return x_69; +} +} +else { -lean_object* x_21; lean_object* x_22; -x_21 = lean_box(7); -x_22 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_22, 0, x_21); -return x_22; +lean_object* x_70; lean_object* x_71; +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_70 = lean_box(6); +x_71 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_71, 0, x_70); +return x_71; +} } else { -lean_object* x_23; lean_object* x_24; -x_23 = lean_box(8); -x_24 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_24, 0, x_23); -return x_24; +lean_object* x_72; lean_object* x_73; +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_72 = lean_box(9); +x_73 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_73, 0, x_72); +return x_73; } } -case 10: +else { -lean_object* x_25; -x_25 = lean_ctor_get(x_1, 1); -lean_inc_ref(x_25); +lean_object* x_74; lean_object* x_75; +lean_dec(x_3); +lean_dec_ref(x_2); lean_dec_ref(x_1); -x_1 = x_25; -goto _start; +x_74 = lean_box(0); +x_75 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_75, 0, x_74); +return x_75; } -default: +} +else { -lean_object* x_27; lean_object* x_28; +lean_object* x_76; lean_object* x_77; +lean_dec(x_3); +lean_dec_ref(x_2); lean_dec_ref(x_1); -x_27 = l_Lean_IR_toIRType___closed__4; -x_28 = l_panic___at___00Lean_IR_toIRType_spec__1(x_27, x_2, x_3); -return x_28; +x_76 = lean_box(5); +x_77 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_77, 0, x_76); +return x_77; +} +} +else +{ +lean_object* x_78; lean_object* x_79; +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_78 = lean_box(4); +x_79 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_79, 0, x_78); +return x_79; } } +else +{ +lean_object* x_80; lean_object* x_81; +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_80 = lean_box(3); +x_81 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_81, 0, x_80); +return x_81; } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_toIRType_visitApp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +else { -lean_object* x_6; -lean_inc(x_4); -lean_inc_ref(x_3); -lean_inc(x_1); -x_6 = l_Lean_IR_hasTrivialStructure_x3f(x_1, x_3, x_4); -if (lean_obj_tag(x_6) == 0) +lean_object* x_82; lean_object* x_83; +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_82 = lean_box(2); +x_83 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_83, 0, x_82); +return x_83; +} +} +else { -lean_object* x_7; -x_7 = lean_ctor_get(x_6, 0); -lean_inc(x_7); -lean_dec_ref(x_6); -if (lean_obj_tag(x_7) == 1) +lean_object* x_84; lean_object* x_85; +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_84 = lean_box(1); +x_85 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_85, 0, x_84); +return x_85; +} +} +else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -lean_dec(x_1); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -lean_dec_ref(x_7); -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_8, 1); -lean_inc(x_10); -x_11 = lean_ctor_get(x_8, 2); -lean_inc(x_11); -lean_dec(x_8); +x_10 = x_2; +x_11 = x_3; x_12 = lean_box(0); -lean_inc(x_4); -lean_inc_ref(x_3); -x_13 = l_Lean_Compiler_LCNF_getOtherDeclBaseType(x_9, x_12, x_3, x_4); -if (lean_obj_tag(x_13) == 0) +goto block_43; +} +} +else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_10 = x_2; +x_11 = x_3; +x_12 = lean_box(0); +goto block_43; +} +block_8: +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_box(8); +x_7 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_7, 0, x_6); +return x_7; +} +block_43: +{ +lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; +x_13 = lean_st_ref_get(x_11); x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -lean_dec_ref(x_13); -x_15 = lean_unsigned_to_nat(0u); -x_16 = l_Array_toSubarray___redArg(x_2, x_15, x_10); -x_17 = l_Subarray_toArray___redArg(x_16); -x_18 = l_Lean_Compiler_LCNF_instantiateForall(x_14, x_17, x_3, x_4); -lean_dec_ref(x_17); -if (lean_obj_tag(x_18) == 0) +lean_inc_ref(x_14); +lean_dec(x_13); +x_15 = 0; +lean_inc_ref(x_14); +x_16 = l_Lean_Environment_find_x3f(x_14, x_1, x_15); +if (lean_obj_tag(x_16) == 1) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_19 = lean_ctor_get(x_18, 0); +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec_ref(x_16); +if (lean_obj_tag(x_17) == 5) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_18 = lean_ctor_get(x_17, 0); +lean_inc_ref(x_18); +lean_dec_ref(x_17); +x_19 = lean_ctor_get(x_18, 4); lean_inc(x_19); lean_dec_ref(x_18); -x_20 = l_Lean_instInhabitedExpr; -x_21 = l_Lean_Compiler_LCNF_getParamTypes(x_19); -x_22 = lean_array_get(x_20, x_21, x_11); -lean_dec(x_11); -lean_dec_ref(x_21); -lean_inc(x_4); -lean_inc_ref(x_3); -x_23 = l_Lean_Compiler_LCNF_toMonoType(x_22, x_3, x_4); -if (lean_obj_tag(x_23) == 0) +x_20 = lean_unsigned_to_nat(0u); +lean_inc(x_19); +x_21 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg(x_14, x_9, x_20, x_19, x_20, x_10, x_11); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec_ref(x_23); -x_25 = l_Lean_IR_toIRType(x_24, x_3, x_4); -return x_25; -} -else +uint8_t x_22; +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_23 = lean_ctor_get(x_21, 0); +x_24 = l_List_lengthTR___redArg(x_19); +lean_dec(x_19); +x_25 = lean_nat_dec_eq(x_23, x_24); +if (x_25 == 0) { uint8_t x_26; -lean_dec(x_4); -lean_dec_ref(x_3); -x_26 = !lean_is_exclusive(x_23); +lean_dec(x_24); +x_26 = lean_nat_dec_eq(x_23, x_20); +lean_dec(x_23); if (x_26 == 0) { -return x_23; +lean_object* x_27; +x_27 = lean_box(8); +lean_ctor_set(x_21, 0, x_27); +return x_21; } else { -lean_object* x_27; lean_object* x_28; -x_27 = lean_ctor_get(x_23, 0); -lean_inc(x_27); -lean_dec(x_23); -x_28 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_28, 0, x_27); -return x_28; -} -} +lean_object* x_28; +x_28 = lean_box(7); +lean_ctor_set(x_21, 0, x_28); +return x_21; } -else -{ -uint8_t x_29; -lean_dec(x_11); -lean_dec(x_4); -lean_dec_ref(x_3); -x_29 = !lean_is_exclusive(x_18); -if (x_29 == 0) -{ -return x_18; } else { -lean_object* x_30; lean_object* x_31; -x_30 = lean_ctor_get(x_18, 0); -lean_inc(x_30); -lean_dec(x_18); -x_31 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_31, 0, x_30); -return x_31; -} +lean_object* x_29; +lean_dec(x_23); +x_29 = l_Lean_IR_irTypeForEnum(x_24); +lean_dec(x_24); +lean_ctor_set(x_21, 0, x_29); +return x_21; } } else { -uint8_t x_32; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_4); -lean_dec_ref(x_3); -lean_dec_ref(x_2); -x_32 = !lean_is_exclusive(x_13); +lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_30 = lean_ctor_get(x_21, 0); +lean_inc(x_30); +lean_dec(x_21); +x_31 = l_List_lengthTR___redArg(x_19); +lean_dec(x_19); +x_32 = lean_nat_dec_eq(x_30, x_31); if (x_32 == 0) { -return x_13; +uint8_t x_33; +lean_dec(x_31); +x_33 = lean_nat_dec_eq(x_30, x_20); +lean_dec(x_30); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; +x_34 = lean_box(8); +x_35 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_35, 0, x_34); +return x_35; } else { -lean_object* x_33; lean_object* x_34; -x_33 = lean_ctor_get(x_13, 0); -lean_inc(x_33); -lean_dec(x_13); -x_34 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_34, 0, x_33); -return x_34; -} +lean_object* x_36; lean_object* x_37; +x_36 = lean_box(7); +x_37 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_37, 0, x_36); +return x_37; } } else { -lean_object* x_35; -lean_dec(x_7); -lean_dec_ref(x_2); -x_35 = l_Lean_IR_nameToIRType(x_1, x_3, x_4); -return x_35; +lean_object* x_38; lean_object* x_39; +lean_dec(x_30); +x_38 = l_Lean_IR_irTypeForEnum(x_31); +lean_dec(x_31); +x_39 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_39, 0, x_38); +return x_39; +} } } else { -uint8_t x_36; -lean_dec(x_4); -lean_dec_ref(x_3); -lean_dec_ref(x_2); -lean_dec(x_1); -x_36 = !lean_is_exclusive(x_6); -if (x_36 == 0) +uint8_t x_40; +lean_dec(x_19); +x_40 = !lean_is_exclusive(x_21); +if (x_40 == 0) { -return x_6; +return x_21; } else { -lean_object* x_37; lean_object* x_38; -x_37 = lean_ctor_get(x_6, 0); -lean_inc(x_37); -lean_dec(x_6); -x_38 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_38, 0, x_37); -return x_38; -} +lean_object* x_41; lean_object* x_42; +x_41 = lean_ctor_get(x_21, 0); +lean_inc(x_41); +lean_dec(x_21); +x_42 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_42, 0, x_41); +return x_42; } } } -LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_toIRType_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +else { -lean_object* x_5; -x_5 = l_panic___at___00Lean_IR_toIRType_spec__1(x_1, x_2, x_3); -return x_5; +lean_dec(x_17); +lean_dec_ref(x_14); +lean_dec(x_11); +lean_dec_ref(x_10); +x_5 = lean_box(0); +goto block_8; } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_toIRType_visitApp___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +else { -lean_object* x_6; -x_6 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_toIRType_visitApp(x_1, x_2, x_3, x_4); -return x_6; +lean_dec(x_16); +lean_dec_ref(x_14); +lean_dec(x_11); +lean_dec_ref(x_10); +x_5 = lean_box(0); +goto block_8; } } -LEAN_EXPORT lean_object* l_Lean_IR_toIRType___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_IR_toIRType(x_1, x_2, x_3); -return x_5; } } -LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_ctorIdx(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_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: { -switch (lean_obj_tag(x_1)) { -case 0: -{ -lean_object* x_2; -x_2 = lean_unsigned_to_nat(0u); -return x_2; -} -case 1: -{ -lean_object* x_3; -x_3 = lean_unsigned_to_nat(1u); -return x_3; +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_4); +x_11 = l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2(x_1, x_2, x_3, x_10, x_5, x_6, x_7, x_8); +return x_11; } -case 2: -{ -lean_object* x_4; -x_4 = lean_unsigned_to_nat(2u); -return x_4; } -case 3: -{ -lean_object* x_5; -x_5 = lean_unsigned_to_nat(3u); -return x_5; -} -default: +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; +lean_object* x_21 = _args[20]; +lean_object* x_22 = _args[21]; +lean_object* x_23 = _args[22]; +lean_object* x_24 = _args[23]; +lean_object* x_25 = _args[24]; +lean_object* x_26 = _args[25]; +lean_object* x_27 = _args[26]; +lean_object* x_28 = _args[27]; +_start: { -lean_object* x_6; -x_6 = lean_unsigned_to_nat(4u); -return x_6; -} -} +lean_object* x_29; +x_29 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_27); +lean_dec(x_22); +lean_dec_ref(x_21); +lean_dec_ref(x_20); +lean_dec_ref(x_19); +lean_dec_ref(x_18); +lean_dec_ref(x_17); +lean_dec_ref(x_16); +lean_dec_ref(x_15); +lean_dec_ref(x_14); +lean_dec_ref(x_13); +lean_dec_ref(x_12); +lean_dec_ref(x_11); +lean_dec_ref(x_10); +lean_dec_ref(x_8); +lean_dec_ref(x_7); +lean_dec_ref(x_6); +lean_dec_ref(x_5); +lean_dec_ref(x_4); +lean_dec_ref(x_3); +return x_29; } } -LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_ctorIdx___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_2; -x_2 = l_Lean_IR_CtorFieldInfo_ctorIdx(x_1); -lean_dec(x_1); -return x_2; +lean_object* x_13; +x_13 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; } } -LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_ctorElim___redArg(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -switch (lean_obj_tag(x_1)) { -case 1: -{ -lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_3 = lean_ctor_get(x_1, 0); -lean_inc(x_3); -x_4 = lean_ctor_get(x_1, 1); -lean_inc(x_4); -lean_dec_ref(x_1); -x_5 = lean_apply_2(x_2, x_3, x_4); +lean_object* x_5; +x_5 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__0(x_1, x_2, x_3); return x_5; } -case 2: -{ -lean_object* x_6; lean_object* x_7; -x_6 = lean_ctor_get(x_1, 0); -lean_inc(x_6); -lean_dec_ref(x_1); -x_7 = lean_apply_1(x_2, x_6); -return x_7; -} -case 3: -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_8 = lean_ctor_get(x_1, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_1, 1); -lean_inc(x_9); -x_10 = lean_ctor_get(x_1, 2); -lean_inc(x_10); -lean_dec_ref(x_1); -x_11 = lean_apply_3(x_2, x_8, x_9, x_10); -return x_11; } -default: +LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -lean_dec(x_1); -return x_2; -} -} +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_3); +x_10 = l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2___redArg(x_1, x_2, x_9, x_4, x_5, x_6, x_7); +return x_10; } } -LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_ctorElim(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_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__1___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_6; -x_6 = l_Lean_IR_CtorFieldInfo_ctorElim___redArg(x_3, x_5); -return x_6; +lean_object* x_10; +x_10 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__1___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_10; } } -LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_ctorElim___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_6; -x_6 = l_Lean_IR_CtorFieldInfo_ctorElim(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_2); -return x_6; +lean_object* x_9; +x_9 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +return x_9; } } -LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_erased_elim(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___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_IR_CtorFieldInfo_ctorElim___redArg(x_2, x_4); +x_5 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache(x_1, x_2, x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_erased_elim___redArg(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg___closed__1() { _start: { -lean_object* x_3; -x_3 = l_Lean_IR_CtorFieldInfo_ctorElim___redArg(x_1, x_2); -return x_3; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Name_hash___override___boxed), 1, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_object_elim(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg___closed__0() { _start: { -lean_object* x_5; -x_5 = l_Lean_IR_CtorFieldInfo_ctorElim___redArg(x_2, x_4); -return x_5; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Name_beq___boxed), 2, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_object_elim___redArg(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg___closed__2() { _start: { -lean_object* x_3; -x_3 = l_Lean_IR_CtorFieldInfo_ctorElim___redArg(x_1, x_2); +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg___closed__1; +x_2 = l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg___closed__0; +x_3 = l_Lean_PersistentHashMap_instInhabited(lean_box(0), lean_box(0), x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_usize_elim(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg___closed__3() { _start: { -lean_object* x_5; -x_5 = l_Lean_IR_CtorFieldInfo_ctorElim___redArg(x_2, x_4); -return x_5; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg___closed__2; +x_2 = lean_box(0); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_usize_elim___redArg(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_3; -x_3 = l_Lean_IR_CtorFieldInfo_ctorElim___redArg(x_1, x_2); -return x_3; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_5 = lean_st_ref_get(x_3); +x_6 = lean_ctor_get(x_5, 0); +lean_inc_ref(x_6); +lean_dec(x_5); +x_7 = lean_ctor_get(x_1, 2); +x_8 = l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg___closed__3; +x_9 = lean_box(0); +x_10 = l___private_Lean_Environment_0__Lean_EnvExtension_getStateUnsafe___redArg(x_8, x_1, x_6, x_7, x_9); +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec(x_10); +x_12 = l_Lean_PersistentHashMap_find_x3f___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__1___redArg(x_11, x_2); +x_13 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_13, 0, x_12); +return x_13; } } -LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_scalar_elim(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_5; -x_5 = l_Lean_IR_CtorFieldInfo_ctorElim___redArg(x_2, x_4); -return x_5; +lean_object* x_8; +x_8 = l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg(x_3, x_4, x_6); +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_scalar_elim___redArg(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_3; -x_3 = l_Lean_IR_CtorFieldInfo_ctorElim___redArg(x_1, x_2); +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_3, 1); +lean_inc(x_1); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_1); +lean_ctor_set(x_7, 1, x_5); +x_8 = l_Lean_PersistentHashMap_insert___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__0___redArg(x_6, x_1, x_2); +lean_ctor_set(x_3, 1, x_8); +lean_ctor_set(x_3, 0, x_7); return x_3; } -} -LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_void_elim(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +else { -lean_object* x_5; -x_5 = l_Lean_IR_CtorFieldInfo_ctorElim___redArg(x_2, x_4); -return x_5; -} +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_9 = lean_ctor_get(x_3, 0); +x_10 = lean_ctor_get(x_3, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_3); +lean_inc(x_1); +x_11 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_11, 0, x_1); +lean_ctor_set(x_11, 1, x_9); +x_12 = l_Lean_PersistentHashMap_insert___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__0___redArg(x_10, x_1, x_2); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_11); +lean_ctor_set(x_13, 1, x_12); +return x_13; } -LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_void_elim___redArg(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_Lean_IR_CtorFieldInfo_ctorElim___redArg(x_1, x_2); -return x_3; } } -static lean_object* _init_l_Lean_IR_instInhabitedCtorFieldInfo_default() { +static lean_object* _init_l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg___closed__0() { _start: { lean_object* x_1; -x_1 = lean_box(0); +x_1 = l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_box(0), lean_box(0)); return x_1; } } -static lean_object* _init_l_Lean_IR_instInhabitedCtorFieldInfo() { -_start: -{ -lean_object* x_1; -x_1 = lean_box(0); -return x_1; -} -} -static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__0() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("◾", 3, 1); -return x_1; -} -} -static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__1() { +static lean_object* _init_l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg___closed__1() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_IR_CtorFieldInfo_format___closed__0; -x_2 = lean_alloc_ctor(3, 1, 0); +x_1 = l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg___closed__0; +x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("obj@", 4, 4); -return x_1; -} -} -static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__3() { +static lean_object* _init_l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_IR_CtorFieldInfo_format___closed__2; -x_2 = lean_alloc_ctor(3, 1, 0); +x_1 = l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg___closed__1; +x_2 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); return x_2; } } -static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__4() { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked(":", 1, 1); -return x_1; -} -} -static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__5() { -_start: +lean_object* x_6; uint8_t x_7; +x_6 = lean_st_ref_take(x_4); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_IR_CtorFieldInfo_format___closed__4; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_8 = lean_ctor_get(x_6, 0); +x_9 = lean_ctor_get(x_6, 5); +lean_dec(x_9); +x_10 = lean_ctor_get(x_1, 2); +lean_inc(x_10); +x_11 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg___lam__0), 3, 2); +lean_closure_set(x_11, 0, x_2); +lean_closure_set(x_11, 1, x_3); +x_12 = lean_box(0); +x_13 = l_Lean_EnvExtension_modifyState___redArg(x_1, x_8, x_11, x_10, x_12); +lean_dec(x_10); +x_14 = l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg___closed__2; +lean_ctor_set(x_6, 5, x_14); +lean_ctor_set(x_6, 0, x_13); +x_15 = lean_st_ref_set(x_4, x_6); +x_16 = lean_box(0); +x_17 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_17, 0, x_16); +return x_17; } -static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__6() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("usize@", 6, 6); -return x_1; +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_18 = lean_ctor_get(x_6, 0); +x_19 = lean_ctor_get(x_6, 1); +x_20 = lean_ctor_get(x_6, 2); +x_21 = lean_ctor_get(x_6, 3); +x_22 = lean_ctor_get(x_6, 4); +x_23 = lean_ctor_get(x_6, 6); +x_24 = lean_ctor_get(x_6, 7); +x_25 = lean_ctor_get(x_6, 8); +lean_inc(x_25); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_6); +x_26 = lean_ctor_get(x_1, 2); +lean_inc(x_26); +x_27 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg___lam__0), 3, 2); +lean_closure_set(x_27, 0, x_2); +lean_closure_set(x_27, 1, x_3); +x_28 = lean_box(0); +x_29 = l_Lean_EnvExtension_modifyState___redArg(x_1, x_18, x_27, x_26, x_28); +lean_dec(x_26); +x_30 = l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg___closed__2; +x_31 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_19); +lean_ctor_set(x_31, 2, x_20); +lean_ctor_set(x_31, 3, x_21); +lean_ctor_set(x_31, 4, x_22); +lean_ctor_set(x_31, 5, x_30); +lean_ctor_set(x_31, 6, x_23); +lean_ctor_set(x_31, 7, x_24); +lean_ctor_set(x_31, 8, x_25); +x_32 = lean_st_ref_set(x_4, x_31); +x_33 = lean_box(0); +x_34 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_34, 0, x_33); +return x_34; } } -static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__7() { +} +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_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_1; lean_object* x_2; -x_1 = l_Lean_IR_CtorFieldInfo_format___closed__6; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_9; +x_9 = l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg(x_3, x_4, x_5, x_7); +return x_9; } } -static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__8() { +static lean_object* _init_l_Lean_IR_nameToIRType___closed__0() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("scalar#", 7, 7); +x_1 = l_Lean_IR_irTypeExt; return x_1; } } -static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__9() { +LEAN_EXPORT lean_object* l_Lean_IR_nameToIRType(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_IR_CtorFieldInfo_format___closed__8; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__10() { -_start: +lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_5 = l_Lean_IR_nameToIRType___closed__0; +x_6 = l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg(x_5, x_1, x_3); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("@", 1, 1); -return x_1; -} +lean_object* x_8; +x_8 = lean_ctor_get(x_6, 0); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; +lean_free_object(x_6); +lean_inc(x_3); +lean_inc(x_1); +x_9 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache(x_1, x_2, x_3); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +lean_dec_ref(x_9); +lean_inc(x_10); +x_11 = l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg(x_5, x_1, x_10, x_3); +lean_dec(x_3); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; +x_13 = lean_ctor_get(x_11, 0); +lean_dec(x_13); +lean_ctor_set(x_11, 0, x_10); +return x_11; } -static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__11() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_IR_CtorFieldInfo_format___closed__10; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_14; +lean_dec(x_11); +x_14 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_14, 0, x_10); +return x_14; } } -static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__12() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("void", 4, 4); -return x_1; +lean_dec(x_3); +lean_dec(x_1); +return x_9; } } -static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__13() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_IR_CtorFieldInfo_format___closed__12; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_15; +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec(x_1); +x_15 = lean_ctor_get(x_8, 0); +lean_inc(x_15); +lean_dec_ref(x_8); +lean_ctor_set(x_6, 0, x_15); +return x_6; } } -LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_format(lean_object* x_1) { -_start: -{ -switch (lean_obj_tag(x_1)) { -case 0: +else { -lean_object* x_2; -x_2 = l_Lean_IR_CtorFieldInfo_format___closed__1; -return x_2; -} -case 1: +lean_object* x_16; +x_16 = lean_ctor_get(x_6, 0); +lean_inc(x_16); +lean_dec(x_6); +if (lean_obj_tag(x_16) == 0) { -uint8_t x_3; -x_3 = !lean_is_exclusive(x_1); -if (x_3 == 0) +lean_object* x_17; +lean_inc(x_3); +lean_inc(x_1); +x_17 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache(x_1, x_2, x_3); +if (lean_obj_tag(x_17) == 0) { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_4 = lean_ctor_get(x_1, 0); -x_5 = lean_ctor_get(x_1, 1); -x_6 = l_Lean_IR_CtorFieldInfo_format___closed__3; -x_7 = l_Nat_reprFast(x_4); -x_8 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_8, 0, x_7); -lean_ctor_set_tag(x_1, 5); -lean_ctor_set(x_1, 1, x_8); -lean_ctor_set(x_1, 0, x_6); -x_9 = l_Lean_IR_CtorFieldInfo_format___closed__5; -x_10 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_10, 0, x_1); -lean_ctor_set(x_10, 1, x_9); -x_11 = l_Lean_IR_instToFormatIRType___private__1(x_5); -x_12 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_12, 0, x_10); -lean_ctor_set(x_12, 1, x_11); -return x_12; -} -else -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_13 = lean_ctor_get(x_1, 0); -x_14 = lean_ctor_get(x_1, 1); -lean_inc(x_14); -lean_inc(x_13); -lean_dec(x_1); -x_15 = l_Lean_IR_CtorFieldInfo_format___closed__3; -x_16 = l_Nat_reprFast(x_13); -x_17 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_17, 0, x_16); -x_18 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_18, 0, x_15); -lean_ctor_set(x_18, 1, x_17); -x_19 = l_Lean_IR_CtorFieldInfo_format___closed__5; -x_20 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_20, 0, x_18); -lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_IR_instToFormatIRType___private__1(x_14); -x_22 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -return x_22; +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +lean_dec_ref(x_17); +lean_inc(x_18); +x_19 = l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg(x_5, x_1, x_18, x_3); +lean_dec(x_3); +if (lean_is_exclusive(x_19)) { + lean_ctor_release(x_19, 0); + x_20 = x_19; +} else { + lean_dec_ref(x_19); + x_20 = lean_box(0); } +if (lean_is_scalar(x_20)) { + x_21 = lean_alloc_ctor(0, 1, 0); +} else { + x_21 = x_20; } -case 2: -{ -uint8_t x_23; -x_23 = !lean_is_exclusive(x_1); -if (x_23 == 0) -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_24 = lean_ctor_get(x_1, 0); -x_25 = l_Lean_IR_CtorFieldInfo_format___closed__7; -x_26 = l_Nat_reprFast(x_24); -lean_ctor_set_tag(x_1, 3); -lean_ctor_set(x_1, 0, x_26); -x_27 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_1); -return x_27; +lean_ctor_set(x_21, 0, x_18); +return x_21; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_28 = lean_ctor_get(x_1, 0); -lean_inc(x_28); +lean_dec(x_3); lean_dec(x_1); -x_29 = l_Lean_IR_CtorFieldInfo_format___closed__7; -x_30 = l_Nat_reprFast(x_28); -x_31 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_31, 0, x_30); -x_32 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_32, 0, x_29); -lean_ctor_set(x_32, 1, x_31); -return x_32; -} +return x_17; } -case 3: -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_33 = lean_ctor_get(x_1, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_1, 1); -lean_inc(x_34); -x_35 = lean_ctor_get(x_1, 2); -lean_inc(x_35); -lean_dec_ref(x_1); -x_36 = l_Lean_IR_CtorFieldInfo_format___closed__9; -x_37 = l_Nat_reprFast(x_33); -x_38 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_38, 0, x_37); -x_39 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l_Lean_IR_CtorFieldInfo_format___closed__11; -x_41 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); -x_42 = l_Nat_reprFast(x_34); -x_43 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_43, 0, x_42); -x_44 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_44, 0, x_41); -lean_ctor_set(x_44, 1, x_43); -x_45 = l_Lean_IR_CtorFieldInfo_format___closed__5; -x_46 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -x_47 = l_Lean_IR_instToFormatIRType___private__1(x_35); -x_48 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_48, 0, x_46); -lean_ctor_set(x_48, 1, x_47); -return x_48; } -default: +else { -lean_object* x_49; -x_49 = l_Lean_IR_CtorFieldInfo_format___closed__13; -return x_49; -} -} +lean_object* x_22; lean_object* x_23; +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec(x_1); +x_22 = lean_ctor_get(x_16, 0); +lean_inc(x_22); +lean_dec_ref(x_16); +x_23 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_23, 0, x_22); +return x_23; } } -static lean_object* _init_l_Lean_IR_CtorFieldInfo_instToFormat___closed__0() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_IR_CtorFieldInfo_format), 1, 0); -return x_1; } } -static lean_object* _init_l_Lean_IR_CtorFieldInfo_instToFormat() { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_1; -x_1 = l_Lean_IR_CtorFieldInfo_instToFormat___closed__0; -return x_1; +lean_object* x_8; +x_8 = l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec(x_2); +return x_8; } } -static lean_object* _init_l_Lean_IR_instInhabitedCtorLayout_default___closed__0() { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_1; -x_1 = l_Lean_IR_instInhabitedCtorInfo_default; -return x_1; +lean_object* x_9; +x_9 = l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec(x_2); +return x_9; } } -static lean_object* _init_l_Lean_IR_instInhabitedCtorLayout_default___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_1; -x_1 = l_Array_empty(lean_box(0)); -return x_1; +lean_object* x_5; +x_5 = l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec_ref(x_1); +return x_5; } } -static lean_object* _init_l_Lean_IR_instInhabitedCtorLayout_default___closed__2() { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_IR_instInhabitedCtorLayout_default___closed__1; -x_2 = l_Lean_IR_instInhabitedCtorLayout_default___closed__0; -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; +lean_object* x_6; +x_6 = l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg(x_1, x_2, x_3, x_4); +lean_dec(x_4); +return x_6; } } -static lean_object* _init_l_Lean_IR_instInhabitedCtorLayout_default() { +LEAN_EXPORT lean_object* l_Lean_IR_nameToIRType___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_1; -x_1 = l_Lean_IR_instInhabitedCtorLayout_default___closed__2; -return x_1; +lean_object* x_5; +x_5 = l_Lean_IR_nameToIRType(x_1, x_2, x_3); +return x_5; } } -static lean_object* _init_l_Lean_IR_instInhabitedCtorLayout() { +static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_isAnyProducingType___closed__0() { _start: { lean_object* x_1; -x_1 = l_Lean_IR_instInhabitedCtorLayout_default; +x_1 = lean_mk_string_unchecked("lcAny", 5, 5); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_895574536____hygCtx___hyg_2_() { +LEAN_EXPORT uint8_t l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_isAnyProducingType(lean_object* x_1) { _start: { -lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_IR_instInhabitedCtorLayout_default; -x_3 = l_Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0___redArg(x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_895574536____hygCtx___hyg_2____boxed(lean_object* x_1) { -_start: +switch (lean_obj_tag(x_1)) { +case 4: { lean_object* x_2; -x_2 = l_Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_895574536____hygCtx___hyg_2_(); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__5___redArg(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_10; lean_object* x_11; -x_10 = lean_alloc_closure((void*)(l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2___redArg___lam__0___boxed), 8, 1); -lean_closure_set(x_10, 0, x_2); -x_11 = l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingImp(lean_box(0), x_1, x_10, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_11) == 0) +x_2 = lean_ctor_get(x_1, 0); +if (lean_obj_tag(x_2) == 1) { -uint8_t x_12; -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) +lean_object* x_3; +x_3 = lean_ctor_get(x_2, 0); +if (lean_obj_tag(x_3) == 0) { -return x_11; +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 1); +x_5 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_isAnyProducingType___closed__0; +x_6 = lean_string_dec_eq(x_4, x_5); +return x_6; } else { -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_11, 0); -lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_14, 0, x_13); -return x_14; +uint8_t x_7; +x_7 = 0; +return x_7; } } else { -uint8_t x_15; -x_15 = !lean_is_exclusive(x_11); -if (x_15 == 0) -{ -return x_11; +uint8_t x_8; +x_8 = 0; +return x_8; +} +} +case 7: +{ +lean_object* x_9; +x_9 = lean_ctor_get(x_1, 2); +x_1 = x_9; +goto _start; +} +default: +{ +uint8_t x_11; +x_11 = 0; +return x_11; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_isAnyProducingType___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_isAnyProducingType(x_1); +lean_dec_ref(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__4___redArg(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_alloc_closure((void*)(l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2___redArg___lam__0___boxed), 8, 1); +lean_closure_set(x_10, 0, x_2); +x_11 = l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingImp(lean_box(0), 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) +{ +return x_11; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_14, 0, x_13); +return x_14; +} +} +else +{ +uint8_t x_15; +x_15 = !lean_is_exclusive(x_11); +if (x_15 == 0) +{ +return x_11; +} +else +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_11, 0); +lean_inc(x_16); +lean_dec(x_11); +x_17 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_17, 0, x_16); +return x_17; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_forallTelescopeReducing___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__4___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_11; +} +} +static lean_object* _init_l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__0() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__5; +x_2 = lean_unsigned_to_nat(0u); +x_3 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_2); +lean_ctor_set(x_3, 2, x_2); +lean_ctor_set(x_3, 3, x_1); +lean_ctor_set(x_3, 4, x_1); +lean_ctor_set(x_3, 5, x_1); +lean_ctor_set(x_3, 6, x_1); +lean_ctor_set(x_3, 7, x_1); +lean_ctor_set(x_3, 8, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(32u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__1; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__3() { +_start: +{ +size_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = 5; +x_2 = lean_unsigned_to_nat(0u); +x_3 = l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__1; +x_4 = l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__2; +x_5 = lean_alloc_ctor(0, 4, sizeof(size_t)*1); +lean_ctor_set(x_5, 0, x_4); +lean_ctor_set(x_5, 1, x_3); +lean_ctor_set(x_5, 2, x_2); +lean_ctor_set(x_5, 3, x_2); +lean_ctor_set_usize(x_5, 4, x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(1); +x_2 = l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__3; +x_3 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__5; +x_4 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_1); +return x_4; +} +} +static lean_object* _init_l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("A private declaration `", 23, 23); +return x_1; +} +} +static lean_object* _init_l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__0; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("` (from the current module) exists but would need to be public to access here.", 78, 78); +return x_1; +} +} +static lean_object* _init_l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("A public declaration `", 22, 22); +return x_1; +} +} +static lean_object* _init_l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("` exists but is imported privately; consider adding `public import ", 67, 67); +return x_1; +} +} +static lean_object* _init_l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__6; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("`.", 2, 2); +return x_1; +} +} +static lean_object* _init_l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__8; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("` (from `", 9, 9); +return x_1; +} +} +static lean_object* _init_l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__10; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("`) exists but would need to be public to access here.", 53, 53); +return x_1; +} +} +static lean_object* _init_l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__12; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_5 = lean_st_ref_get(x_3); +x_6 = lean_ctor_get(x_5, 0); +lean_inc_ref(x_6); +lean_dec(x_5); +x_7 = l_Lean_Name_isAnonymous(x_2); +if (x_7 == 0) +{ +uint8_t x_8; +x_8 = lean_ctor_get_uint8(x_6, sizeof(void*)*8); +if (x_8 == 0) +{ +lean_object* x_9; +lean_dec_ref(x_6); +lean_dec(x_2); +x_9 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_9, 0, x_1); +return x_9; +} +else +{ +lean_object* x_10; uint8_t x_11; +lean_inc_ref(x_6); +x_10 = l_Lean_Environment_setExporting(x_6, x_7); +lean_inc(x_2); +lean_inc_ref(x_10); +x_11 = l_Lean_Environment_contains(x_10, x_2, x_8); +if (x_11 == 0) +{ +lean_object* x_12; +lean_dec_ref(x_10); +lean_dec_ref(x_6); +lean_dec(x_2); +x_12 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_12, 0, x_1); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_13 = l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__0; +x_14 = l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__4; +x_15 = lean_box(0); +x_16 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_16, 0, x_10); +lean_ctor_set(x_16, 1, x_13); +lean_ctor_set(x_16, 2, x_14); +lean_ctor_set(x_16, 3, x_15); +lean_inc(x_2); +x_17 = l_Lean_MessageData_ofConstName(x_2, x_7); +x_18 = lean_alloc_ctor(3, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +x_19 = l_Lean_Environment_getModuleIdxFor_x3f(x_6, x_2); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_dec_ref(x_6); +lean_dec(x_2); +x_20 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__1; +x_21 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_18); +x_22 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__3; +x_23 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +x_24 = l_Lean_MessageData_note(x_23); +x_25 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_25, 0, x_1); +lean_ctor_set(x_25, 1, x_24); +x_26 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_26, 0, x_25); +return x_26; +} +else +{ +uint8_t x_27; +x_27 = !lean_is_exclusive(x_19); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_28 = lean_ctor_get(x_19, 0); +x_29 = lean_box(0); +x_30 = l_Lean_Environment_header(x_6); +lean_dec_ref(x_6); +x_31 = l_Lean_EnvironmentHeader_moduleNames(x_30); +x_32 = lean_array_get(x_29, x_31, x_28); +lean_dec(x_28); +lean_dec_ref(x_31); +x_33 = l_Lean_isPrivateName(x_2); +lean_dec(x_2); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_34 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__5; +x_35 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_18); +x_36 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__7; +x_37 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +x_38 = l_Lean_MessageData_ofName(x_32); +x_39 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +x_40 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__9; +x_41 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +x_42 = l_Lean_MessageData_note(x_41); +x_43 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_43, 0, x_1); +lean_ctor_set(x_43, 1, x_42); +lean_ctor_set_tag(x_19, 0); +lean_ctor_set(x_19, 0, x_43); +return x_19; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_44 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__1; +x_45 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_18); +x_46 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__11; +x_47 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +x_48 = l_Lean_MessageData_ofName(x_32); +x_49 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_48); +x_50 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__13; +x_51 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +x_52 = l_Lean_MessageData_note(x_51); +x_53 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_53, 0, x_1); +lean_ctor_set(x_53, 1, x_52); +lean_ctor_set_tag(x_19, 0); +lean_ctor_set(x_19, 0, x_53); +return x_19; +} +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; +x_54 = lean_ctor_get(x_19, 0); +lean_inc(x_54); +lean_dec(x_19); +x_55 = lean_box(0); +x_56 = l_Lean_Environment_header(x_6); +lean_dec_ref(x_6); +x_57 = l_Lean_EnvironmentHeader_moduleNames(x_56); +x_58 = lean_array_get(x_55, x_57, x_54); +lean_dec(x_54); +lean_dec_ref(x_57); +x_59 = l_Lean_isPrivateName(x_2); +lean_dec(x_2); +if (x_59 == 0) +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_60 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__5; +x_61 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_18); +x_62 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__7; +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_64 = l_Lean_MessageData_ofName(x_58); +x_65 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_64); +x_66 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__9; +x_67 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_67, 0, x_65); +lean_ctor_set(x_67, 1, x_66); +x_68 = l_Lean_MessageData_note(x_67); +x_69 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_69, 0, x_1); +lean_ctor_set(x_69, 1, x_68); +x_70 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_70, 0, x_69); +return x_70; +} +else +{ +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_71 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__1; +x_72 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set(x_72, 1, x_18); +x_73 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__11; +x_74 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_74, 0, x_72); +lean_ctor_set(x_74, 1, x_73); +x_75 = l_Lean_MessageData_ofName(x_58); +x_76 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_76, 0, x_74); +lean_ctor_set(x_76, 1, x_75); +x_77 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__13; +x_78 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_77); +x_79 = l_Lean_MessageData_note(x_78); +x_80 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_80, 0, x_1); +lean_ctor_set(x_80, 1, x_79); +x_81 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_81, 0, x_80); +return x_81; +} +} +} +} +} +} +else +{ +lean_object* x_82; +lean_dec_ref(x_6); +lean_dec(x_2); +x_82 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_82, 0, x_1); +return x_82; +} +} +} +LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_6; +x_6 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg(x_1, x_2, x_4); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_5 = lean_st_ref_get(x_3); +x_6 = lean_ctor_get(x_5, 0); +lean_inc_ref(x_6); +lean_dec(x_5); +x_7 = lean_ctor_get(x_2, 2); +x_8 = l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__0; +x_9 = l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__4; +lean_inc(x_7); +x_10 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_10, 0, x_6); +lean_ctor_set(x_10, 1, x_8); +lean_ctor_set(x_10, 2, x_9); +lean_ctor_set(x_10, 3, x_7); +x_11 = lean_alloc_ctor(3, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_1); +x_12 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_12, 0, x_11); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_5 = lean_ctor_get(x_2, 5); +x_6 = l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11(x_1, x_2, x_3); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_6, 0); +lean_inc(x_5); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_5); +lean_ctor_set(x_9, 1, x_8); +lean_ctor_set_tag(x_6, 1); +lean_ctor_set(x_6, 0, x_9); +return x_6; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_ctor_get(x_6, 0); +lean_inc(x_10); +lean_dec(x_6); +lean_inc(x_5); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_5); +lean_ctor_set(x_11, 1, x_10); +x_12 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_12, 0, x_11); +return x_12; +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_6; +x_6 = l_Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11___redArg(x_2, x_3, x_4); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__20___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint8_t x_6; +x_6 = !lean_is_exclusive(x_3); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_ctor_get(x_3, 5); +x_8 = l_Lean_replaceRef(x_1, x_7); +lean_dec(x_7); +lean_ctor_set(x_3, 5, x_8); +x_9 = l_Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11___redArg(x_2, x_3, x_4); +lean_dec_ref(x_3); +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_10 = lean_ctor_get(x_3, 0); +x_11 = lean_ctor_get(x_3, 1); +x_12 = lean_ctor_get(x_3, 2); +x_13 = lean_ctor_get(x_3, 3); +x_14 = lean_ctor_get(x_3, 4); +x_15 = lean_ctor_get(x_3, 5); +x_16 = lean_ctor_get(x_3, 6); +x_17 = lean_ctor_get(x_3, 7); +x_18 = lean_ctor_get(x_3, 8); +x_19 = lean_ctor_get(x_3, 9); +x_20 = lean_ctor_get(x_3, 10); +x_21 = lean_ctor_get(x_3, 11); +x_22 = lean_ctor_get_uint8(x_3, sizeof(void*)*14); +x_23 = lean_ctor_get(x_3, 12); +x_24 = lean_ctor_get_uint8(x_3, sizeof(void*)*14 + 1); +x_25 = lean_ctor_get(x_3, 13); +lean_inc(x_25); +lean_inc(x_23); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_dec(x_3); +x_26 = l_Lean_replaceRef(x_1, x_15); +lean_dec(x_15); +x_27 = lean_alloc_ctor(0, 14, 2); +lean_ctor_set(x_27, 0, x_10); +lean_ctor_set(x_27, 1, x_11); +lean_ctor_set(x_27, 2, x_12); +lean_ctor_set(x_27, 3, x_13); +lean_ctor_set(x_27, 4, x_14); +lean_ctor_set(x_27, 5, x_26); +lean_ctor_set(x_27, 6, x_16); +lean_ctor_set(x_27, 7, x_17); +lean_ctor_set(x_27, 8, x_18); +lean_ctor_set(x_27, 9, x_19); +lean_ctor_set(x_27, 10, x_20); +lean_ctor_set(x_27, 11, x_21); +lean_ctor_set(x_27, 12, x_23); +lean_ctor_set(x_27, 13, x_25); +lean_ctor_set_uint8(x_27, sizeof(void*)*14, x_22); +lean_ctor_set_uint8(x_27, sizeof(void*)*14 + 1, x_24); +x_28 = l_Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11___redArg(x_2, x_27, x_4); +lean_dec_ref(x_27); +return x_28; +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__20(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__20___redArg(x_2, x_3, x_4, x_5); +return x_7; +} +} +static lean_object* _init_l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Unknown constant `", 18, 18); +return x_1; +} +} +static lean_object* _init_l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg___closed__0; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("`", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_unknownIdentifierMessageTag; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_6; uint8_t x_7; +x_6 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg(x_1, x_2, x_4); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_ctor_get(x_6, 0); +x_9 = l_Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19___closed__0; +x_10 = lean_alloc_ctor(8, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_8); +lean_ctor_set(x_6, 0, x_10); +return x_6; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = lean_ctor_get(x_6, 0); +lean_inc(x_11); +lean_dec(x_6); +x_12 = l_Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19___closed__0; +x_13 = lean_alloc_ctor(8, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_11); +x_14 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_14, 0, x_13); +return x_14; +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = l_Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19(x_2, x_3, x_4, x_5); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +lean_dec_ref(x_7); +x_9 = l_Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__20___redArg(x_1, x_8, x_4, x_5); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_6; uint8_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_6 = l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg___closed__1; +x_7 = 0; +lean_inc(x_2); +x_8 = l_Lean_MessageData_ofConstName(x_2, x_7); +x_9 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_9, 0, x_6); +lean_ctor_set(x_9, 1, x_8); +x_10 = l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg___closed__3; +x_11 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_11, 0, x_9); +lean_ctor_set(x_11, 1, x_10); +x_12 = l_Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18___redArg(x_1, x_11, x_2, x_3, x_4); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_2, 5); +lean_inc(x_5); +x_6 = l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg(x_5, x_1, x_2, x_3); +lean_dec(x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_6; +x_6 = l_Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15___redArg(x_2, x_3, x_4); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg(x_2, x_3, x_4, x_5); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18___redArg(x_2, x_3, x_4, x_5, x_6); +return x_8; +} +} +static lean_object* _init_l_Lean_IR_getCtorLayout___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_IR_ctorLayoutExt; +return x_1; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.IR.getCtorLayout.fillCache", 31, 31); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__2; +x_2 = lean_unsigned_to_nat(64u); +x_3 = lean_unsigned_to_nat(209u); +x_4 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__0; +x_5 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__0; +x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__0___closed__0; +x_6 = lean_panic_fn(x_5, x_1); +x_7 = lean_apply_3(x_6, x_2, x_3, lean_box(0)); +return x_7; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___closed__0() { +_start: +{ +size_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = 5; +x_2 = lean_unsigned_to_nat(0u); +x_3 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__7; +x_4 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__8; +x_5 = lean_alloc_ctor(0, 4, sizeof(size_t)*1); +lean_ctor_set(x_5, 0, x_4); +lean_ctor_set(x_5, 1, x_3); +lean_ctor_set(x_5, 2, x_2); +lean_ctor_set(x_5, 3, x_2); +lean_ctor_set_usize(x_5, 4, x_1); +return x_5; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__9; +x_2 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___closed__0; +x_3 = lean_box(1); +x_4 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__6; +x_5 = l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__0; +x_6 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_4); +lean_ctor_set(x_6, 2, x_3); +lean_ctor_set(x_6, 3, x_2); +lean_ctor_set(x_6, 4, x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__0(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_usize_dec_lt(x_3, x_2); +if (x_6 == 0) +{ +lean_object* x_7; +x_7 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_7, 0, x_4); +lean_ctor_set(x_7, 1, x_5); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_8 = lean_array_uget(x_4, x_3); +x_9 = lean_unsigned_to_nat(0u); +x_10 = lean_array_uset(x_4, x_3, x_9); +switch (lean_obj_tag(x_8)) { +case 1: +{ +x_11 = x_8; +x_12 = x_5; +goto block_17; +} +case 2: +{ +x_11 = x_8; +x_12 = x_5; +goto block_17; +} +case 3: +{ +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = lean_ctor_get(x_8, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_8, 2); +lean_inc(x_19); +x_20 = lean_nat_dec_eq(x_18, x_1); +if (x_20 == 0) +{ +lean_dec(x_19); +lean_dec(x_18); +x_11 = x_8; +x_12 = x_5; +goto block_17; +} +else +{ +uint8_t x_21; +x_21 = !lean_is_exclusive(x_8); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_22 = lean_ctor_get(x_8, 2); +lean_dec(x_22); +x_23 = lean_ctor_get(x_8, 1); +lean_dec(x_23); +x_24 = lean_ctor_get(x_8, 0); +lean_dec(x_24); +x_25 = lean_nat_add(x_5, x_18); +lean_ctor_set(x_8, 1, x_5); +x_11 = x_8; +x_12 = x_25; +goto block_17; +} +else +{ +lean_object* x_26; lean_object* x_27; +lean_dec(x_8); +x_26 = lean_nat_add(x_5, x_18); +x_27 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_27, 0, x_18); +lean_ctor_set(x_27, 1, x_5); +lean_ctor_set(x_27, 2, x_19); +x_11 = x_27; +x_12 = x_26; +goto block_17; +} +} +} +default: +{ +x_11 = x_8; +x_12 = x_5; +goto block_17; +} +} +block_17: +{ +size_t x_13; size_t x_14; lean_object* x_15; +x_13 = 1; +x_14 = lean_usize_add(x_3, x_13); +x_15 = lean_array_uset(x_10, x_3, x_11); +x_3 = x_14; +x_4 = x_15; +x_5 = x_12; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___lam__0(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_array_size(x_1); +x_5 = 0; +x_6 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__0(x_2, x_4, x_5, x_1, x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___lam__0(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +static lean_object* _init_l_Lean_IR_toIRType___closed__0() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_getConstInfoInduct___at___00Lean_IR_unboxedIRType_spec__9___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("` is not an inductive type", 26, 26); +return x_1; +} +} +static lean_object* _init_l_Lean_getConstInfoInduct___at___00Lean_IR_unboxedIRType_spec__9___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_getConstInfoInduct___at___00Lean_IR_unboxedIRType_spec__9___closed__0; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_getConstInfoInduct___at___00Lean_IR_unboxedIRType_spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_st_ref_get(x_3); +x_6 = lean_ctor_get(x_5, 0); +lean_inc_ref(x_6); +lean_dec(x_5); +lean_inc(x_1); +x_7 = l_Lean_isInductiveCore_x3f(x_6, x_1); +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_8 = l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg___closed__3; +x_9 = 0; +x_10 = l_Lean_MessageData_ofConstName(x_1, x_9); +x_11 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_11, 0, x_8); +lean_ctor_set(x_11, 1, x_10); +x_12 = l_Lean_getConstInfoInduct___at___00Lean_IR_unboxedIRType_spec__9___closed__1; +x_13 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_13, 0, x_11); +lean_ctor_set(x_13, 1, x_12); +x_14 = l_Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11___redArg(x_13, x_2, x_3); +return x_14; +} +else +{ +uint8_t x_15; +lean_dec(x_1); +x_15 = !lean_is_exclusive(x_7); +if (x_15 == 0) +{ +lean_ctor_set_tag(x_7, 0); +return x_7; +} +else +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_7, 0); +lean_inc(x_16); +lean_dec(x_7); +x_17 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_17, 0, x_16); +return x_17; +} +} +} +} +static lean_object* _init_l_Lean_IR_unboxedIRType___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("too few parameters for type ", 28, 28); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_unboxedIRType___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_IR_unboxedIRType___closed__0; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_IR_unboxedIRType___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(": expected ", 11, 11); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_unboxedIRType___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_IR_unboxedIRType___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_IR_unboxedIRType___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" but got only ", 14, 14); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_unboxedIRType___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_IR_unboxedIRType___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_unboxedIRType(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_24; +lean_inc(x_1); +x_24 = l_Lean_getConstInfoInduct___at___00Lean_IR_unboxedIRType_spec__9(x_1, x_3, x_4); +if (lean_obj_tag(x_24) == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_41; uint8_t x_42; +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +if (lean_is_exclusive(x_24)) { + lean_ctor_release(x_24, 0); + x_26 = x_24; +} else { + lean_dec_ref(x_24); + x_26 = lean_box(0); +} +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +x_28 = lean_ctor_get(x_25, 4); +lean_inc(x_28); +lean_dec(x_25); +x_41 = lean_array_get_size(x_2); +x_42 = lean_nat_dec_lt(x_41, x_27); +if (x_42 == 0) +{ +x_29 = x_3; +x_30 = x_4; +x_31 = lean_box(0); +goto block_40; +} +else +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_43 = l_Lean_IR_unboxedIRType___closed__1; +lean_inc(x_1); +x_44 = l_Lean_MessageData_ofName(x_1); +x_45 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +x_46 = l_Lean_IR_unboxedIRType___closed__3; +x_47 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +lean_inc(x_27); +x_48 = l_Nat_reprFast(x_27); +x_49 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_49, 0, x_48); +x_50 = l_Lean_MessageData_ofFormat(x_49); +x_51 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_51, 0, x_47); +lean_ctor_set(x_51, 1, x_50); +x_52 = l_Lean_IR_unboxedIRType___closed__5; +x_53 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +x_54 = l_Nat_reprFast(x_41); +x_55 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_55, 0, x_54); +x_56 = l_Lean_MessageData_ofFormat(x_55); +x_57 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_57, 0, x_53); +lean_ctor_set(x_57, 1, x_56); +x_58 = l_Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11___redArg(x_57, x_3, x_4); +if (lean_obj_tag(x_58) == 0) +{ +lean_dec_ref(x_58); +x_29 = x_3; +x_30 = x_4; +x_31 = lean_box(0); +goto block_40; +} +else +{ +uint8_t x_59; +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec(x_1); +x_59 = !lean_is_exclusive(x_58); +if (x_59 == 0) +{ +return x_58; +} +else +{ +lean_object* x_60; lean_object* x_61; +x_60 = lean_ctor_get(x_58, 0); +lean_inc(x_60); +lean_dec(x_58); +x_61 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_61, 0, x_60); +return x_61; +} +} +} +block_40: +{ +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_32; lean_object* x_33; +lean_dec(x_30); +lean_dec_ref(x_29); +lean_dec(x_27); +lean_dec(x_1); +x_32 = lean_box(6); +if (lean_is_scalar(x_26)) { + x_33 = lean_alloc_ctor(0, 1, 0); +} else { + x_33 = x_26; +} +lean_ctor_set(x_33, 0, x_32); +return x_33; +} +else +{ +lean_object* x_34; lean_object* x_35; +lean_dec(x_26); +x_34 = lean_unsigned_to_nat(0u); +x_35 = l_Array_extract___redArg(x_2, x_34, x_27); +if (lean_obj_tag(x_28) == 1) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_28, 1); +if (lean_obj_tag(x_36) == 0) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_28, 0); +lean_inc(x_37); +lean_dec_ref(x_28); +x_38 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_38, 0, x_1); +x_39 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor(x_38, x_37, x_35, x_29, x_30); +return x_39; +} +else +{ +x_6 = x_28; +x_7 = x_34; +x_8 = x_35; +x_9 = x_29; +x_10 = x_30; +x_11 = lean_box(0); +goto block_23; +} +} +else +{ +x_6 = x_28; +x_7 = x_34; +x_8 = x_35; +x_9 = x_29; +x_10 = x_30; +x_11 = lean_box(0); +goto block_23; +} +} +} +} +else +{ +uint8_t x_62; +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec(x_1); +x_62 = !lean_is_exclusive(x_24); +if (x_62 == 0) +{ +return x_24; +} +else +{ +lean_object* x_63; lean_object* x_64; +x_63 = lean_ctor_get(x_24, 0); +lean_inc(x_63); +lean_dec(x_24); +x_64 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_64, 0, x_63); +return x_64; +} +} +block_23: +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_mk_empty_array_with_capacity(x_7); +lean_dec(x_7); +x_13 = l_List_forIn_x27_loop___at___00Lean_IR_unboxedIRType_spec__10___redArg(x_8, x_6, x_12, x_9, x_10); +if (lean_obj_tag(x_13) == 0) +{ +uint8_t x_14; +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_alloc_ctor(11, 2, 0); +lean_ctor_set(x_16, 0, x_1); +lean_ctor_set(x_16, 1, x_15); +lean_ctor_set(x_13, 0, x_16); +return x_13; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_13, 0); +lean_inc(x_17); +lean_dec(x_13); +x_18 = lean_alloc_ctor(11, 2, 0); +lean_ctor_set(x_18, 0, x_1); +lean_ctor_set(x_18, 1, x_17); +x_19 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_19, 0, x_18); +return x_19; +} +} +else +{ +uint8_t x_20; +lean_dec(x_1); +x_20 = !lean_is_exclusive(x_13); +if (x_20 == 0) +{ +return x_13; +} +else +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_13, 0); +lean_inc(x_21); +lean_dec(x_13); +x_22 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_22, 0, x_21); +return x_22; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_toIRType_visitApp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_6; +lean_inc(x_4); +lean_inc_ref(x_3); +lean_inc(x_1); +x_6 = l_Lean_IR_hasTrivialStructure_x3f(x_1, x_3, x_4); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_7; +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +lean_dec_ref(x_6); +if (lean_obj_tag(x_7) == 1) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +lean_dec(x_1); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +lean_dec_ref(x_7); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +x_11 = lean_ctor_get(x_8, 2); +lean_inc(x_11); +lean_dec(x_8); +x_12 = lean_box(0); +lean_inc(x_4); +lean_inc_ref(x_3); +x_13 = l_Lean_Compiler_LCNF_getOtherDeclBaseType(x_9, x_12, x_3, x_4); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +lean_dec_ref(x_13); +x_15 = lean_unsigned_to_nat(0u); +x_16 = l_Array_toSubarray___redArg(x_2, x_15, x_10); +x_17 = l_Subarray_toArray___redArg(x_16); +x_18 = l_Lean_Compiler_LCNF_instantiateForall(x_14, x_17, x_3, x_4); +lean_dec_ref(x_17); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +lean_dec_ref(x_18); +x_20 = l_Lean_instInhabitedExpr; +x_21 = l_Lean_Compiler_LCNF_getParamTypes(x_19); +x_22 = lean_array_get(x_20, x_21, x_11); +lean_dec(x_11); +lean_dec_ref(x_21); +lean_inc(x_4); +lean_inc_ref(x_3); +x_23 = l_Lean_Compiler_LCNF_toMonoType(x_22, x_3, x_4); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec_ref(x_23); +x_25 = l_Lean_IR_toIRType(x_24, x_3, x_4); +return x_25; +} +else +{ +uint8_t x_26; +lean_dec(x_4); +lean_dec_ref(x_3); +x_26 = !lean_is_exclusive(x_23); +if (x_26 == 0) +{ +return x_23; +} +else +{ +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_23, 0); +lean_inc(x_27); +lean_dec(x_23); +x_28 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_28, 0, x_27); +return x_28; +} +} +} +else +{ +uint8_t x_29; +lean_dec(x_11); +lean_dec(x_4); +lean_dec_ref(x_3); +x_29 = !lean_is_exclusive(x_18); +if (x_29 == 0) +{ +return x_18; +} +else +{ +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_18, 0); +lean_inc(x_30); +lean_dec(x_18); +x_31 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_31, 0, x_30); +return x_31; +} +} +} +else +{ +uint8_t x_32; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +x_32 = !lean_is_exclusive(x_13); +if (x_32 == 0) +{ +return x_13; +} +else +{ +lean_object* x_33; lean_object* x_34; +x_33 = lean_ctor_get(x_13, 0); +lean_inc(x_33); +lean_dec(x_13); +x_34 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_34, 0, x_33); +return x_34; +} +} +} +else +{ +lean_object* x_35; lean_object* x_36; uint8_t x_37; +lean_dec(x_7); +x_35 = lean_st_ref_get(x_4); +x_36 = lean_ctor_get(x_35, 0); +lean_inc_ref(x_36); +lean_dec(x_35); +lean_inc(x_1); +x_37 = l_Lean_IR_UnboxResult_hasUnboxAttr(x_36, x_1); +if (x_37 == 0) +{ +lean_object* x_38; +lean_dec_ref(x_2); +x_38 = l_Lean_IR_nameToIRType(x_1, x_3, x_4); +return x_38; +} +else +{ +lean_object* x_39; +x_39 = l_Lean_IR_unboxedIRType(x_1, x_2, x_3, x_4); +lean_dec_ref(x_2); +return x_39; +} +} +} +else +{ +uint8_t x_40; +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +lean_dec(x_1); +x_40 = !lean_is_exclusive(x_6); +if (x_40 == 0) +{ +return x_6; +} +else +{ +lean_object* x_41; lean_object* x_42; +x_41 = lean_ctor_get(x_6, 0); +lean_inc(x_41); +lean_dec(x_6); +x_42 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_42, 0, x_41); +return x_42; +} +} +} +} +static lean_object* _init_l_Lean_IR_toIRType___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = l_Lean_Expr_sort___override(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_IR_toIRType___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.IR.toIRType", 16, 16); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_toIRType___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_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__2; +x_2 = lean_unsigned_to_nat(41u); +x_3 = lean_unsigned_to_nat(142u); +x_4 = l_Lean_IR_toIRType___closed__2; +x_5 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__0; +x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_toIRType_spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__0___closed__0; +x_6 = lean_panic_fn(x_5, x_1); +x_7 = lean_apply_3(x_6, x_2, x_3, lean_box(0)); +return x_7; +} +} +static lean_object* _init_l_Lean_IR_toIRType___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__2; +x_2 = lean_unsigned_to_nat(9u); +x_3 = lean_unsigned_to_nat(154u); +x_4 = l_Lean_IR_toIRType___closed__2; +x_5 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__0; +x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_toIRType(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 4: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +lean_dec_ref(x_1); +x_6 = l_Lean_IR_toIRType___closed__0; +x_7 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_toIRType_visitApp(x_5, x_6, x_2, x_3); +return x_7; +} +case 5: +{ +lean_object* x_8; +x_8 = l_Lean_Expr_getAppFn(x_1); +if (lean_obj_tag(x_8) == 4) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +lean_dec_ref(x_8); +x_10 = l_Lean_IR_toIRType___closed__1; +x_11 = l_Lean_Expr_getAppNumArgs(x_1); +lean_inc(x_11); +x_12 = lean_mk_array(x_11, x_10); +x_13 = lean_unsigned_to_nat(1u); +x_14 = lean_nat_sub(x_11, x_13); +lean_dec(x_11); +x_15 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_12, x_14); +x_16 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_toIRType_visitApp(x_9, x_15, x_2, x_3); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; +lean_dec_ref(x_8); +lean_dec_ref(x_1); +x_17 = l_Lean_IR_toIRType___closed__3; +x_18 = l_panic___at___00Lean_IR_toIRType_spec__6(x_17, x_2, x_3); +return x_18; +} +} +case 7: +{ +lean_object* x_19; uint8_t x_20; +lean_dec(x_3); +lean_dec_ref(x_2); +x_19 = lean_ctor_get(x_1, 2); +lean_inc_ref(x_19); +lean_dec_ref(x_1); +x_20 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_isAnyProducingType(x_19); +lean_dec_ref(x_19); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_box(7); +x_22 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_22, 0, x_21); +return x_22; +} +else +{ +lean_object* x_23; lean_object* x_24; +x_23 = lean_box(8); +x_24 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_24, 0, x_23); +return x_24; +} +} +case 10: +{ +lean_object* x_25; +x_25 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_25); +lean_dec_ref(x_1); +x_1 = x_25; +goto _start; +} +default: +{ +lean_object* x_27; lean_object* x_28; +lean_dec_ref(x_1); +x_27 = l_Lean_IR_toIRType___closed__4; +x_28 = l_panic___at___00Lean_IR_toIRType_spec__6(x_27, x_2, x_3); +return x_28; +} +} +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(lean_object* x_1, uint8_t x_2, uint8_t x_3, uint8_t x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_13 = lean_array_push(x_1, x_7); +x_14 = lean_box(x_5); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_6); +x_16 = lean_box(x_4); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_15); +x_18 = lean_box(x_3); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +x_20 = lean_box(x_2); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_13); +lean_ctor_set(x_22, 1, x_21); +x_23 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_23, 0, x_22); +x_24 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_24, 0, x_23); +return x_24; +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg(lean_object* x_1, lean_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_9; +x_9 = !lean_is_exclusive(x_2); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_10 = lean_ctor_get(x_2, 0); +x_11 = lean_ctor_get(x_2, 1); +x_12 = lean_ctor_get(x_2, 2); +x_13 = lean_nat_dec_lt(x_11, x_12); +if (x_13 == 0) +{ +lean_object* x_14; +lean_free_object(x_2); +lean_dec(x_12); +lean_dec(x_11); +lean_dec_ref(x_10); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec(x_1); +x_14 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_14, 0, x_3); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_array_fget_borrowed(x_10, x_11); +x_16 = l_Lean_Expr_fvarId_x21(x_15); +lean_inc_ref(x_4); +x_17 = l_Lean_FVarId_getType___redArg(x_16, x_4, x_6, x_7); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +lean_dec_ref(x_17); +lean_inc(x_7); +lean_inc_ref(x_6); +lean_inc(x_5); +lean_inc_ref(x_4); +x_19 = l_Lean_Compiler_LCNF_toLCNFType(x_18, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +lean_dec_ref(x_19); +lean_inc(x_7); +lean_inc_ref(x_6); +x_21 = l_Lean_Compiler_LCNF_toMonoType(x_20, x_6, x_7); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +lean_dec_ref(x_21); +lean_inc(x_7); +lean_inc_ref(x_6); +x_23 = l_Lean_IR_toIRType(x_22, x_6, x_7); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_24 = lean_ctor_get(x_3, 1); +lean_inc(x_24); +x_25 = lean_ctor_get(x_24, 1); +lean_inc(x_25); +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +x_28 = lean_ctor_get(x_23, 0); +lean_inc(x_28); +lean_dec_ref(x_23); +x_29 = lean_ctor_get(x_3, 0); +lean_inc(x_29); +lean_dec_ref(x_3); +x_30 = lean_ctor_get(x_24, 0); +lean_inc(x_30); +lean_dec(x_24); +x_31 = lean_ctor_get(x_25, 0); +lean_inc(x_31); +lean_dec(x_25); +x_32 = lean_ctor_get(x_26, 0); +lean_inc(x_32); +lean_dec(x_26); +x_33 = !lean_is_exclusive(x_27); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_34 = lean_ctor_get(x_27, 0); +x_35 = lean_ctor_get(x_27, 1); +x_36 = lean_unsigned_to_nat(1u); +x_37 = lean_nat_add(x_11, x_36); +lean_dec(x_11); +lean_ctor_set(x_2, 1, x_37); +switch (lean_obj_tag(x_28)) { +case 0: +{ +lean_object* x_53; lean_object* x_54; uint8_t x_55; uint8_t x_56; uint8_t x_57; lean_object* x_58; +lean_free_object(x_27); +lean_dec(x_34); +x_53 = lean_unsigned_to_nat(8u); +lean_inc(x_1); +x_54 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set(x_54, 1, x_1); +lean_ctor_set(x_54, 2, x_28); +x_55 = lean_unbox(x_30); +lean_dec(x_30); +x_56 = lean_unbox(x_31); +lean_dec(x_31); +x_57 = lean_unbox(x_32); +lean_dec(x_32); +x_58 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_29, x_55, x_56, x_57, x_13, x_35, x_54, x_4, x_5, x_6, x_7); +x_38 = x_58; +goto block_52; +} +case 1: +{ +lean_object* x_59; uint8_t x_60; uint8_t x_61; uint8_t x_62; lean_object* x_63; +lean_free_object(x_27); +lean_dec(x_30); +lean_inc(x_1); +x_59 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_59, 0, x_36); +lean_ctor_set(x_59, 1, x_1); +lean_ctor_set(x_59, 2, x_28); +x_60 = lean_unbox(x_31); +lean_dec(x_31); +x_61 = lean_unbox(x_32); +lean_dec(x_32); +x_62 = lean_unbox(x_34); +lean_dec(x_34); +x_63 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_29, x_13, x_60, x_61, x_62, x_35, x_59, x_4, x_5, x_6, x_7); +x_38 = x_63; +goto block_52; +} +case 2: +{ +lean_object* x_64; lean_object* x_65; uint8_t x_66; uint8_t x_67; uint8_t x_68; lean_object* x_69; +lean_free_object(x_27); +lean_dec(x_31); +x_64 = lean_unsigned_to_nat(2u); +lean_inc(x_1); +x_65 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_65, 0, x_64); +lean_ctor_set(x_65, 1, x_1); +lean_ctor_set(x_65, 2, x_28); +x_66 = lean_unbox(x_30); +lean_dec(x_30); +x_67 = lean_unbox(x_32); +lean_dec(x_32); +x_68 = lean_unbox(x_34); +lean_dec(x_34); +x_69 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_29, x_66, x_13, x_67, x_68, x_35, x_65, x_4, x_5, x_6, x_7); +x_38 = x_69; +goto block_52; +} +case 3: +{ +lean_object* x_70; lean_object* x_71; uint8_t x_72; uint8_t x_73; uint8_t x_74; lean_object* x_75; +lean_free_object(x_27); +lean_dec(x_32); +x_70 = lean_unsigned_to_nat(4u); +lean_inc(x_1); +x_71 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_1); +lean_ctor_set(x_71, 2, x_28); +x_72 = lean_unbox(x_30); +lean_dec(x_30); +x_73 = lean_unbox(x_31); +lean_dec(x_31); +x_74 = lean_unbox(x_34); +lean_dec(x_34); +x_75 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_29, x_72, x_73, x_13, x_74, x_35, x_71, x_4, x_5, x_6, x_7); +x_38 = x_75; +goto block_52; +} +case 4: +{ +lean_object* x_76; lean_object* x_77; uint8_t x_78; uint8_t x_79; uint8_t x_80; lean_object* x_81; +lean_free_object(x_27); +lean_dec(x_34); +x_76 = lean_unsigned_to_nat(8u); +lean_inc(x_1); +x_77 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_77, 0, x_76); +lean_ctor_set(x_77, 1, x_1); +lean_ctor_set(x_77, 2, x_28); +x_78 = lean_unbox(x_30); +lean_dec(x_30); +x_79 = lean_unbox(x_31); +lean_dec(x_31); +x_80 = lean_unbox(x_32); +lean_dec(x_32); +x_81 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_29, x_78, x_79, x_80, x_13, x_35, x_77, x_4, x_5, x_6, x_7); +x_38 = x_81; +goto block_52; +} +case 5: +{ +lean_object* x_82; uint8_t x_83; uint8_t x_84; uint8_t x_85; uint8_t x_86; lean_object* x_87; +lean_free_object(x_27); +lean_inc(x_1); +x_82 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_82, 0, x_1); +x_83 = lean_unbox(x_30); +lean_dec(x_30); +x_84 = lean_unbox(x_31); +lean_dec(x_31); +x_85 = lean_unbox(x_32); +lean_dec(x_32); +x_86 = lean_unbox(x_34); +lean_dec(x_34); +x_87 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_29, x_83, x_84, x_85, x_86, x_35, x_82, x_4, x_5, x_6, x_7); +x_38 = x_87; +goto block_52; } -else +case 6: { -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_11, 0); -lean_inc(x_16); -lean_dec(x_11); -x_17 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_17, 0, x_16); -return x_17; -} -} -} +lean_object* x_88; uint8_t x_89; uint8_t x_90; uint8_t x_91; uint8_t x_92; lean_object* x_93; +lean_free_object(x_27); +x_88 = lean_box(0); +x_89 = lean_unbox(x_30); +lean_dec(x_30); +x_90 = lean_unbox(x_31); +lean_dec(x_31); +x_91 = lean_unbox(x_32); +lean_dec(x_32); +x_92 = lean_unbox(x_34); +lean_dec(x_34); +x_93 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_29, x_89, x_90, x_91, x_92, x_35, x_88, x_4, x_5, x_6, x_7); +x_38 = x_93; +goto block_52; } -LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: +case 9: { -lean_object* x_11; -x_11 = l_Lean_Meta_forallTelescopeReducing___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__5___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_11; -} +lean_object* x_94; lean_object* x_95; uint8_t x_96; uint8_t x_97; uint8_t x_98; lean_object* x_99; +lean_free_object(x_27); +lean_dec(x_32); +x_94 = lean_unsigned_to_nat(4u); +lean_inc(x_1); +x_95 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_95, 0, x_94); +lean_ctor_set(x_95, 1, x_1); +lean_ctor_set(x_95, 2, x_28); +x_96 = lean_unbox(x_30); +lean_dec(x_30); +x_97 = lean_unbox(x_31); +lean_dec(x_31); +x_98 = lean_unbox(x_34); +lean_dec(x_34); +x_99 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_29, x_96, x_97, x_13, x_98, x_35, x_95, x_4, x_5, x_6, x_7); +x_38 = x_99; +goto block_52; } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(lean_object* x_1, uint8_t x_2, uint8_t x_3, uint8_t x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: +case 10: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_13 = lean_array_push(x_1, x_7); -x_14 = lean_box(x_5); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_15, 1, x_6); -x_16 = lean_box(x_4); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_15); -x_18 = lean_box(x_3); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_17); -x_20 = lean_box(x_2); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_19); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_13); -lean_ctor_set(x_22, 1, x_21); -x_23 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_23, 0, x_22); -x_24 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_24, 0, x_23); -return x_24; -} +lean_object* x_100; uint8_t x_101; uint8_t x_102; uint8_t x_103; uint8_t x_104; lean_object* x_105; +x_100 = lean_nat_add(x_35, x_36); +lean_ctor_set_tag(x_27, 1); +lean_ctor_set(x_27, 1, x_28); +lean_ctor_set(x_27, 0, x_35); +x_101 = lean_unbox(x_30); +lean_dec(x_30); +x_102 = lean_unbox(x_31); +lean_dec(x_31); +x_103 = lean_unbox(x_32); +lean_dec(x_32); +x_104 = lean_unbox(x_34); +lean_dec(x_34); +x_105 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_29, x_101, x_102, x_103, x_104, x_100, x_27, x_4, x_5, x_6, x_7); +x_38 = x_105; +goto block_52; } -static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__0() { -_start: +case 11: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.IR.getCtorLayout.fillCache", 31, 31); -return x_1; -} +lean_object* x_106; uint8_t x_107; uint8_t x_108; uint8_t x_109; uint8_t x_110; lean_object* x_111; +x_106 = lean_nat_add(x_35, x_36); +lean_ctor_set_tag(x_27, 1); +lean_ctor_set(x_27, 1, x_28); +lean_ctor_set(x_27, 0, x_35); +x_107 = lean_unbox(x_30); +lean_dec(x_30); +x_108 = lean_unbox(x_31); +lean_dec(x_31); +x_109 = lean_unbox(x_32); +lean_dec(x_32); +x_110 = lean_unbox(x_34); +lean_dec(x_34); +x_111 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_29, x_107, x_108, x_109, x_110, x_106, x_27, x_4, x_5, x_6, x_7); +x_38 = x_111; +goto block_52; } -static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___closed__0() { -_start: +case 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_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__2; -x_2 = lean_unsigned_to_nat(34u); -x_3 = lean_unsigned_to_nat(208u); -x_4 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__0; -x_5 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__0; -x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); -return x_6; -} +lean_object* x_112; uint8_t x_113; uint8_t x_114; uint8_t x_115; uint8_t x_116; lean_object* x_117; +lean_free_object(x_27); +x_112 = lean_box(4); +x_113 = lean_unbox(x_30); +lean_dec(x_30); +x_114 = lean_unbox(x_31); +lean_dec(x_31); +x_115 = lean_unbox(x_32); +lean_dec(x_32); +x_116 = lean_unbox(x_34); +lean_dec(x_34); +x_117 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_29, x_113, x_114, x_115, x_116, x_35, x_112, x_4, x_5, x_6, x_7); +x_38 = x_117; +goto block_52; } -static lean_object* _init_l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___closed__0() { -_start: +default: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_instInhabitedMetaM___lam__0___boxed), 5, 0); -return x_1; +lean_object* x_118; uint8_t x_119; uint8_t x_120; uint8_t x_121; uint8_t x_122; lean_object* x_123; +x_118 = lean_nat_add(x_35, x_36); +lean_ctor_set_tag(x_27, 1); +lean_ctor_set(x_27, 1, x_28); +lean_ctor_set(x_27, 0, x_35); +x_119 = lean_unbox(x_30); +lean_dec(x_30); +x_120 = lean_unbox(x_31); +lean_dec(x_31); +x_121 = lean_unbox(x_32); +lean_dec(x_32); +x_122 = lean_unbox(x_34); +lean_dec(x_34); +x_123 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_29, x_119, x_120, x_121, x_122, x_118, x_27, x_4, x_5, x_6, x_7); +x_38 = x_123; +goto block_52; } } -LEAN_EXPORT lean_object* l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +block_52: { -lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_7 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___closed__0; -x_8 = lean_panic_fn(x_7, x_1); -x_9 = lean_apply_5(x_8, x_2, x_3, x_4, x_5, lean_box(0)); -return x_9; -} -} -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg(lean_object* x_1, lean_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_38) == 0) { -uint8_t x_9; -x_9 = !lean_is_exclusive(x_2); -if (x_9 == 0) +uint8_t x_39; +x_39 = !lean_is_exclusive(x_38); +if (x_39 == 0) { -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); -x_13 = lean_nat_dec_lt(x_11, x_12); -if (x_13 == 0) +lean_object* x_40; +x_40 = lean_ctor_get(x_38, 0); +if (lean_obj_tag(x_40) == 0) { -lean_object* x_14; -lean_free_object(x_2); -lean_dec(x_12); -lean_dec(x_11); -lean_dec_ref(x_10); +lean_object* x_41; +lean_dec_ref(x_2); lean_dec(x_7); lean_dec_ref(x_6); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_1); -x_14 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_14, 0, x_3); -return x_14; +x_41 = lean_ctor_get(x_40, 0); +lean_inc(x_41); +lean_dec_ref(x_40); +lean_ctor_set(x_38, 0, x_41); +return x_38; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_array_fget_borrowed(x_10, x_11); -x_16 = l_Lean_Expr_fvarId_x21(x_15); -lean_inc_ref(x_4); -x_17 = l_Lean_FVarId_getType___redArg(x_16, x_4, x_6, x_7); -if (lean_obj_tag(x_17) == 0) -{ -lean_object* x_18; lean_object* x_19; -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -lean_dec_ref(x_17); -lean_inc(x_7); -lean_inc_ref(x_6); -lean_inc(x_5); -lean_inc_ref(x_4); -x_19 = l_Lean_Compiler_LCNF_toLCNFType(x_18, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_19) == 0) +lean_object* x_42; +lean_free_object(x_38); +x_42 = lean_ctor_get(x_40, 0); +lean_inc(x_42); +lean_dec_ref(x_40); +x_3 = x_42; +goto _start; +} +} +else { -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -lean_dec_ref(x_19); -lean_inc(x_7); -lean_inc_ref(x_6); -x_21 = l_Lean_Compiler_LCNF_toMonoType(x_20, x_6, x_7); -if (lean_obj_tag(x_21) == 0) +lean_object* x_44; +x_44 = lean_ctor_get(x_38, 0); +lean_inc(x_44); +lean_dec(x_38); +if (lean_obj_tag(x_44) == 0) { -lean_object* x_22; lean_object* x_23; -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -lean_dec_ref(x_21); -lean_inc(x_7); -lean_inc_ref(x_6); -x_23 = l_Lean_IR_toIRType(x_22, x_6, x_7); -if (lean_obj_tag(x_23) == 0) +lean_object* x_45; lean_object* x_46; +lean_dec_ref(x_2); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec(x_1); +x_45 = lean_ctor_get(x_44, 0); +lean_inc(x_45); +lean_dec_ref(x_44); +x_46 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_46, 0, x_45); +return x_46; +} +else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; -x_24 = lean_ctor_get(x_3, 1); -lean_inc(x_24); -x_25 = lean_ctor_get(x_24, 1); -lean_inc(x_25); -x_26 = lean_ctor_get(x_25, 1); -lean_inc(x_26); -x_27 = lean_ctor_get(x_26, 1); -lean_inc(x_27); -x_28 = lean_ctor_get(x_23, 0); -lean_inc(x_28); -lean_dec_ref(x_23); -x_29 = lean_ctor_get(x_3, 0); -lean_inc(x_29); -lean_dec_ref(x_3); -x_30 = lean_ctor_get(x_24, 0); -lean_inc(x_30); -lean_dec(x_24); -x_31 = lean_ctor_get(x_25, 0); -lean_inc(x_31); -lean_dec(x_25); -x_32 = lean_ctor_get(x_26, 0); -lean_inc(x_32); -lean_dec(x_26); -x_33 = !lean_is_exclusive(x_27); -if (x_33 == 0) +lean_object* x_47; +x_47 = lean_ctor_get(x_44, 0); +lean_inc(x_47); +lean_dec_ref(x_44); +x_3 = x_47; +goto _start; +} +} +} +else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_34 = lean_ctor_get(x_27, 0); -x_35 = lean_ctor_get(x_27, 1); -x_36 = lean_unsigned_to_nat(1u); -x_37 = lean_nat_add(x_11, x_36); +uint8_t x_49; +lean_dec_ref(x_2); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec(x_1); +x_49 = !lean_is_exclusive(x_38); +if (x_49 == 0) +{ +return x_38; +} +else +{ +lean_object* x_50; lean_object* x_51; +x_50 = lean_ctor_get(x_38, 0); +lean_inc(x_50); +lean_dec(x_38); +x_51 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_51, 0, x_50); +return x_51; +} +} +} +} +else +{ +lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; +x_124 = lean_ctor_get(x_27, 0); +x_125 = lean_ctor_get(x_27, 1); +lean_inc(x_125); +lean_inc(x_124); +lean_dec(x_27); +x_126 = lean_unsigned_to_nat(1u); +x_127 = lean_nat_add(x_11, x_126); lean_dec(x_11); -lean_ctor_set(x_2, 1, x_37); +lean_ctor_set(x_2, 1, x_127); switch (lean_obj_tag(x_28)) { case 0: { -lean_object* x_53; lean_object* x_54; uint8_t x_55; uint8_t x_56; uint8_t x_57; lean_object* x_58; -lean_free_object(x_27); -lean_dec(x_34); -x_53 = lean_unsigned_to_nat(8u); +lean_object* x_139; lean_object* x_140; uint8_t x_141; uint8_t x_142; uint8_t x_143; lean_object* x_144; +lean_dec(x_124); +x_139 = lean_unsigned_to_nat(8u); lean_inc(x_1); -x_54 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_1); -lean_ctor_set(x_54, 2, x_28); -x_55 = lean_unbox(x_30); +x_140 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_140, 0, x_139); +lean_ctor_set(x_140, 1, x_1); +lean_ctor_set(x_140, 2, x_28); +x_141 = lean_unbox(x_30); lean_dec(x_30); -x_56 = lean_unbox(x_31); +x_142 = lean_unbox(x_31); lean_dec(x_31); -x_57 = lean_unbox(x_32); +x_143 = lean_unbox(x_32); lean_dec(x_32); -x_58 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_29, x_55, x_56, x_57, x_13, x_35, x_54, x_4, x_5, x_6, x_7); -x_38 = x_58; -goto block_52; +x_144 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_29, x_141, x_142, x_143, x_13, x_125, x_140, x_4, x_5, x_6, x_7); +x_128 = x_144; +goto block_138; } case 1: { -lean_object* x_59; uint8_t x_60; uint8_t x_61; uint8_t x_62; lean_object* x_63; -lean_free_object(x_27); +lean_object* x_145; uint8_t x_146; uint8_t x_147; uint8_t x_148; lean_object* x_149; lean_dec(x_30); lean_inc(x_1); -x_59 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_59, 0, x_36); -lean_ctor_set(x_59, 1, x_1); -lean_ctor_set(x_59, 2, x_28); -x_60 = lean_unbox(x_31); +x_145 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_145, 0, x_126); +lean_ctor_set(x_145, 1, x_1); +lean_ctor_set(x_145, 2, x_28); +x_146 = lean_unbox(x_31); lean_dec(x_31); -x_61 = lean_unbox(x_32); +x_147 = lean_unbox(x_32); lean_dec(x_32); -x_62 = lean_unbox(x_34); -lean_dec(x_34); -x_63 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_29, x_13, x_60, x_61, x_62, x_35, x_59, x_4, x_5, x_6, x_7); -x_38 = x_63; -goto block_52; +x_148 = lean_unbox(x_124); +lean_dec(x_124); +x_149 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_29, x_13, x_146, x_147, x_148, x_125, x_145, x_4, x_5, x_6, x_7); +x_128 = x_149; +goto block_138; } case 2: { -lean_object* x_64; lean_object* x_65; uint8_t x_66; uint8_t x_67; uint8_t x_68; lean_object* x_69; -lean_free_object(x_27); +lean_object* x_150; lean_object* x_151; uint8_t x_152; uint8_t x_153; uint8_t x_154; lean_object* x_155; lean_dec(x_31); -x_64 = lean_unsigned_to_nat(2u); +x_150 = lean_unsigned_to_nat(2u); lean_inc(x_1); -x_65 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_65, 0, x_64); -lean_ctor_set(x_65, 1, x_1); -lean_ctor_set(x_65, 2, x_28); -x_66 = lean_unbox(x_30); +x_151 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_151, 0, x_150); +lean_ctor_set(x_151, 1, x_1); +lean_ctor_set(x_151, 2, x_28); +x_152 = lean_unbox(x_30); lean_dec(x_30); -x_67 = lean_unbox(x_32); +x_153 = lean_unbox(x_32); lean_dec(x_32); -x_68 = lean_unbox(x_34); -lean_dec(x_34); -x_69 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_29, x_66, x_13, x_67, x_68, x_35, x_65, x_4, x_5, x_6, x_7); -x_38 = x_69; -goto block_52; +x_154 = lean_unbox(x_124); +lean_dec(x_124); +x_155 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_29, x_152, x_13, x_153, x_154, x_125, x_151, x_4, x_5, x_6, x_7); +x_128 = x_155; +goto block_138; } case 3: { -lean_object* x_70; lean_object* x_71; uint8_t x_72; uint8_t x_73; uint8_t x_74; lean_object* x_75; -lean_free_object(x_27); +lean_object* x_156; lean_object* x_157; uint8_t x_158; uint8_t x_159; uint8_t x_160; lean_object* x_161; lean_dec(x_32); -x_70 = lean_unsigned_to_nat(4u); +x_156 = lean_unsigned_to_nat(4u); lean_inc(x_1); -x_71 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_1); -lean_ctor_set(x_71, 2, x_28); -x_72 = lean_unbox(x_30); +x_157 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_157, 0, x_156); +lean_ctor_set(x_157, 1, x_1); +lean_ctor_set(x_157, 2, x_28); +x_158 = lean_unbox(x_30); lean_dec(x_30); -x_73 = lean_unbox(x_31); +x_159 = lean_unbox(x_31); lean_dec(x_31); -x_74 = lean_unbox(x_34); -lean_dec(x_34); -x_75 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_29, x_72, x_73, x_13, x_74, x_35, x_71, x_4, x_5, x_6, x_7); -x_38 = x_75; -goto block_52; +x_160 = lean_unbox(x_124); +lean_dec(x_124); +x_161 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_29, x_158, x_159, x_13, x_160, x_125, x_157, x_4, x_5, x_6, x_7); +x_128 = x_161; +goto block_138; } case 4: { -lean_object* x_76; lean_object* x_77; uint8_t x_78; uint8_t x_79; uint8_t x_80; lean_object* x_81; -lean_free_object(x_27); -lean_dec(x_34); -x_76 = lean_unsigned_to_nat(8u); +lean_object* x_162; lean_object* x_163; uint8_t x_164; uint8_t x_165; uint8_t x_166; lean_object* x_167; +lean_dec(x_124); +x_162 = lean_unsigned_to_nat(8u); lean_inc(x_1); -x_77 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_77, 0, x_76); -lean_ctor_set(x_77, 1, x_1); -lean_ctor_set(x_77, 2, x_28); -x_78 = lean_unbox(x_30); +x_163 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_163, 0, x_162); +lean_ctor_set(x_163, 1, x_1); +lean_ctor_set(x_163, 2, x_28); +x_164 = lean_unbox(x_30); lean_dec(x_30); -x_79 = lean_unbox(x_31); +x_165 = lean_unbox(x_31); lean_dec(x_31); -x_80 = lean_unbox(x_32); +x_166 = lean_unbox(x_32); lean_dec(x_32); -x_81 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_29, x_78, x_79, x_80, x_13, x_35, x_77, x_4, x_5, x_6, x_7); -x_38 = x_81; -goto block_52; +x_167 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_29, x_164, x_165, x_166, x_13, x_125, x_163, x_4, x_5, x_6, x_7); +x_128 = x_167; +goto block_138; } case 5: { -lean_object* x_82; uint8_t x_83; uint8_t x_84; uint8_t x_85; uint8_t x_86; lean_object* x_87; -lean_free_object(x_27); +lean_object* x_168; uint8_t x_169; uint8_t x_170; uint8_t x_171; uint8_t x_172; lean_object* x_173; lean_inc(x_1); -x_82 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_82, 0, x_1); -x_83 = lean_unbox(x_30); +x_168 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_168, 0, x_1); +x_169 = lean_unbox(x_30); lean_dec(x_30); -x_84 = lean_unbox(x_31); +x_170 = lean_unbox(x_31); lean_dec(x_31); -x_85 = lean_unbox(x_32); +x_171 = lean_unbox(x_32); lean_dec(x_32); -x_86 = lean_unbox(x_34); -lean_dec(x_34); -x_87 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_29, x_83, x_84, x_85, x_86, x_35, x_82, x_4, x_5, x_6, x_7); -x_38 = x_87; -goto block_52; +x_172 = lean_unbox(x_124); +lean_dec(x_124); +x_173 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_29, x_169, x_170, x_171, x_172, x_125, x_168, x_4, x_5, x_6, x_7); +x_128 = x_173; +goto block_138; } case 6: { -lean_object* x_88; uint8_t x_89; uint8_t x_90; uint8_t x_91; uint8_t x_92; lean_object* x_93; -lean_free_object(x_27); -x_88 = lean_box(0); -x_89 = lean_unbox(x_30); +lean_object* x_174; uint8_t x_175; uint8_t x_176; uint8_t x_177; uint8_t x_178; lean_object* x_179; +x_174 = lean_box(0); +x_175 = lean_unbox(x_30); lean_dec(x_30); -x_90 = lean_unbox(x_31); +x_176 = lean_unbox(x_31); lean_dec(x_31); -x_91 = lean_unbox(x_32); +x_177 = lean_unbox(x_32); lean_dec(x_32); -x_92 = lean_unbox(x_34); -lean_dec(x_34); -x_93 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_29, x_89, x_90, x_91, x_92, x_35, x_88, x_4, x_5, x_6, x_7); -x_38 = x_93; -goto block_52; +x_178 = lean_unbox(x_124); +lean_dec(x_124); +x_179 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_29, x_175, x_176, x_177, x_178, x_125, x_174, x_4, x_5, x_6, x_7); +x_128 = x_179; +goto block_138; } case 9: { -lean_object* x_94; lean_object* x_95; uint8_t x_96; uint8_t x_97; uint8_t x_98; lean_object* x_99; -lean_free_object(x_27); +lean_object* x_180; lean_object* x_181; uint8_t x_182; uint8_t x_183; uint8_t x_184; lean_object* x_185; lean_dec(x_32); -x_94 = lean_unsigned_to_nat(4u); +x_180 = lean_unsigned_to_nat(4u); lean_inc(x_1); -x_95 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_95, 0, x_94); -lean_ctor_set(x_95, 1, x_1); -lean_ctor_set(x_95, 2, x_28); -x_96 = lean_unbox(x_30); +x_181 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_181, 0, x_180); +lean_ctor_set(x_181, 1, x_1); +lean_ctor_set(x_181, 2, x_28); +x_182 = lean_unbox(x_30); lean_dec(x_30); -x_97 = lean_unbox(x_31); +x_183 = lean_unbox(x_31); lean_dec(x_31); -x_98 = lean_unbox(x_34); -lean_dec(x_34); -x_99 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_29, x_96, x_97, x_13, x_98, x_35, x_95, x_4, x_5, x_6, x_7); -x_38 = x_99; -goto block_52; +x_184 = lean_unbox(x_124); +lean_dec(x_124); +x_185 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_29, x_182, x_183, x_13, x_184, x_125, x_181, x_4, x_5, x_6, x_7); +x_128 = x_185; +goto block_138; } case 10: { -lean_object* x_100; lean_object* x_101; -lean_free_object(x_27); -lean_dec_ref(x_28); -x_100 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___closed__0; -lean_inc(x_7); -lean_inc_ref(x_6); -lean_inc(x_5); -lean_inc_ref(x_4); -x_101 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2(x_100, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_101) == 0) -{ -lean_object* x_102; uint8_t x_103; uint8_t x_104; uint8_t x_105; uint8_t x_106; lean_object* x_107; -x_102 = lean_ctor_get(x_101, 0); -lean_inc(x_102); -lean_dec_ref(x_101); -x_103 = lean_unbox(x_30); +lean_object* x_186; lean_object* x_187; uint8_t x_188; uint8_t x_189; uint8_t x_190; uint8_t x_191; lean_object* x_192; +x_186 = lean_nat_add(x_125, x_126); +x_187 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_187, 0, x_125); +lean_ctor_set(x_187, 1, x_28); +x_188 = lean_unbox(x_30); lean_dec(x_30); -x_104 = lean_unbox(x_31); +x_189 = lean_unbox(x_31); lean_dec(x_31); -x_105 = lean_unbox(x_32); +x_190 = lean_unbox(x_32); lean_dec(x_32); -x_106 = lean_unbox(x_34); -lean_dec(x_34); -x_107 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_29, x_103, x_104, x_105, x_106, x_35, x_102, x_4, x_5, x_6, x_7); -x_38 = x_107; -goto block_52; +x_191 = lean_unbox(x_124); +lean_dec(x_124); +x_192 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_29, x_188, x_189, x_190, x_191, x_186, x_187, x_4, x_5, x_6, x_7); +x_128 = x_192; +goto block_138; +} +case 11: +{ +lean_object* x_193; lean_object* x_194; uint8_t x_195; uint8_t x_196; uint8_t x_197; uint8_t x_198; lean_object* x_199; +x_193 = lean_nat_add(x_125, x_126); +x_194 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_194, 0, x_125); +lean_ctor_set(x_194, 1, x_28); +x_195 = lean_unbox(x_30); +lean_dec(x_30); +x_196 = lean_unbox(x_31); +lean_dec(x_31); +x_197 = lean_unbox(x_32); +lean_dec(x_32); +x_198 = lean_unbox(x_124); +lean_dec(x_124); +x_199 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_29, x_195, x_196, x_197, x_198, x_193, x_194, x_4, x_5, x_6, x_7); +x_128 = x_199; +goto block_138; +} +case 13: +{ +lean_object* x_200; uint8_t x_201; uint8_t x_202; uint8_t x_203; uint8_t x_204; lean_object* x_205; +x_200 = lean_box(4); +x_201 = lean_unbox(x_30); +lean_dec(x_30); +x_202 = lean_unbox(x_31); +lean_dec(x_31); +x_203 = lean_unbox(x_32); +lean_dec(x_32); +x_204 = lean_unbox(x_124); +lean_dec(x_124); +x_205 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_29, x_201, x_202, x_203, x_204, x_125, x_200, x_4, x_5, x_6, x_7); +x_128 = x_205; +goto block_138; +} +default: +{ +lean_object* x_206; lean_object* x_207; uint8_t x_208; uint8_t x_209; uint8_t x_210; uint8_t x_211; lean_object* x_212; +x_206 = lean_nat_add(x_125, x_126); +x_207 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_207, 0, x_125); +lean_ctor_set(x_207, 1, x_28); +x_208 = lean_unbox(x_30); +lean_dec(x_30); +x_209 = lean_unbox(x_31); +lean_dec(x_31); +x_210 = lean_unbox(x_32); +lean_dec(x_32); +x_211 = lean_unbox(x_124); +lean_dec(x_124); +x_212 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_29, x_208, x_209, x_210, x_211, x_206, x_207, x_4, x_5, x_6, x_7); +x_128 = x_212; +goto block_138; +} +} +block_138: +{ +if (lean_obj_tag(x_128) == 0) +{ +lean_object* x_129; lean_object* x_130; +x_129 = lean_ctor_get(x_128, 0); +lean_inc(x_129); +if (lean_is_exclusive(x_128)) { + lean_ctor_release(x_128, 0); + x_130 = x_128; +} else { + lean_dec_ref(x_128); + x_130 = lean_box(0); +} +if (lean_obj_tag(x_129) == 0) +{ +lean_object* x_131; lean_object* x_132; +lean_dec_ref(x_2); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec(x_1); +x_131 = lean_ctor_get(x_129, 0); +lean_inc(x_131); +lean_dec_ref(x_129); +if (lean_is_scalar(x_130)) { + x_132 = lean_alloc_ctor(0, 1, 0); +} else { + x_132 = x_130; +} +lean_ctor_set(x_132, 0, x_131); +return x_132; +} +else +{ +lean_object* x_133; +lean_dec(x_130); +x_133 = lean_ctor_get(x_129, 0); +lean_inc(x_133); +lean_dec_ref(x_129); +x_3 = x_133; +goto _start; +} +} +else +{ +lean_object* x_135; lean_object* x_136; lean_object* x_137; +lean_dec_ref(x_2); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec(x_1); +x_135 = lean_ctor_get(x_128, 0); +lean_inc(x_135); +if (lean_is_exclusive(x_128)) { + lean_ctor_release(x_128, 0); + x_136 = x_128; +} else { + lean_dec_ref(x_128); + x_136 = lean_box(0); +} +if (lean_is_scalar(x_136)) { + x_137 = lean_alloc_ctor(1, 1, 0); +} else { + x_137 = x_136; +} +lean_ctor_set(x_137, 0, x_135); +return x_137; +} +} +} +} +else +{ +uint8_t x_213; +lean_free_object(x_2); +lean_dec(x_12); +lean_dec(x_11); +lean_dec_ref(x_10); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec_ref(x_3); +lean_dec(x_1); +x_213 = !lean_is_exclusive(x_23); +if (x_213 == 0) +{ +return x_23; +} +else +{ +lean_object* x_214; lean_object* x_215; +x_214 = lean_ctor_get(x_23, 0); +lean_inc(x_214); +lean_dec(x_23); +x_215 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_215, 0, x_214); +return x_215; +} +} +} +else +{ +uint8_t x_216; +lean_free_object(x_2); +lean_dec(x_12); +lean_dec(x_11); +lean_dec_ref(x_10); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec_ref(x_3); +lean_dec(x_1); +x_216 = !lean_is_exclusive(x_21); +if (x_216 == 0) +{ +return x_21; +} +else +{ +lean_object* x_217; lean_object* x_218; +x_217 = lean_ctor_get(x_21, 0); +lean_inc(x_217); +lean_dec(x_21); +x_218 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_218, 0, x_217); +return x_218; +} +} +} +else +{ +uint8_t x_219; +lean_free_object(x_2); +lean_dec(x_12); +lean_dec(x_11); +lean_dec_ref(x_10); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec_ref(x_3); +lean_dec(x_1); +x_219 = !lean_is_exclusive(x_19); +if (x_219 == 0) +{ +return x_19; } else { -uint8_t x_108; -lean_dec_ref(x_2); -lean_dec(x_35); -lean_dec(x_34); -lean_dec(x_32); -lean_dec(x_31); -lean_dec(x_30); -lean_dec(x_29); +lean_object* x_220; lean_object* x_221; +x_220 = lean_ctor_get(x_19, 0); +lean_inc(x_220); +lean_dec(x_19); +x_221 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_221, 0, x_220); +return x_221; +} +} +} +else +{ +uint8_t x_222; +lean_free_object(x_2); +lean_dec(x_12); +lean_dec(x_11); +lean_dec_ref(x_10); lean_dec(x_7); lean_dec_ref(x_6); lean_dec(x_5); lean_dec_ref(x_4); +lean_dec_ref(x_3); lean_dec(x_1); -x_108 = !lean_is_exclusive(x_101); -if (x_108 == 0) +x_222 = !lean_is_exclusive(x_17); +if (x_222 == 0) { -return x_101; +return x_17; } else { -lean_object* x_109; lean_object* x_110; -x_109 = lean_ctor_get(x_101, 0); -lean_inc(x_109); -lean_dec(x_101); -x_110 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_110, 0, x_109); -return x_110; +lean_object* x_223; lean_object* x_224; +x_223 = lean_ctor_get(x_17, 0); +lean_inc(x_223); +lean_dec(x_17); +x_224 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_224, 0, x_223); +return x_224; } } } -case 11: -{ -lean_object* x_111; lean_object* x_112; -lean_free_object(x_27); -lean_dec_ref(x_28); -x_111 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___closed__0; -lean_inc(x_7); -lean_inc_ref(x_6); -lean_inc(x_5); -lean_inc_ref(x_4); -x_112 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2(x_111, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_112) == 0) -{ -lean_object* x_113; uint8_t x_114; uint8_t x_115; uint8_t x_116; uint8_t x_117; lean_object* x_118; -x_113 = lean_ctor_get(x_112, 0); -lean_inc(x_113); -lean_dec_ref(x_112); -x_114 = lean_unbox(x_30); -lean_dec(x_30); -x_115 = lean_unbox(x_31); -lean_dec(x_31); -x_116 = lean_unbox(x_32); -lean_dec(x_32); -x_117 = lean_unbox(x_34); -lean_dec(x_34); -x_118 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_29, x_114, x_115, x_116, x_117, x_35, x_113, x_4, x_5, x_6, x_7); -x_38 = x_118; -goto block_52; } else { -uint8_t x_119; -lean_dec_ref(x_2); -lean_dec(x_35); -lean_dec(x_34); -lean_dec(x_32); -lean_dec(x_31); -lean_dec(x_30); -lean_dec(x_29); +lean_object* x_225; lean_object* x_226; lean_object* x_227; uint8_t x_228; +x_225 = lean_ctor_get(x_2, 0); +x_226 = lean_ctor_get(x_2, 1); +x_227 = lean_ctor_get(x_2, 2); +lean_inc(x_227); +lean_inc(x_226); +lean_inc(x_225); +lean_dec(x_2); +x_228 = lean_nat_dec_lt(x_226, x_227); +if (x_228 == 0) +{ +lean_object* x_229; +lean_dec(x_227); +lean_dec(x_226); +lean_dec_ref(x_225); lean_dec(x_7); lean_dec_ref(x_6); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_1); -x_119 = !lean_is_exclusive(x_112); -if (x_119 == 0) -{ -return x_112; +x_229 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_229, 0, x_3); +return x_229; } else { -lean_object* x_120; lean_object* x_121; -x_120 = lean_ctor_get(x_112, 0); -lean_inc(x_120); -lean_dec(x_112); -x_121 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_121, 0, x_120); -return x_121; +lean_object* x_230; lean_object* x_231; lean_object* x_232; +x_230 = lean_array_fget_borrowed(x_225, x_226); +x_231 = l_Lean_Expr_fvarId_x21(x_230); +lean_inc_ref(x_4); +x_232 = l_Lean_FVarId_getType___redArg(x_231, x_4, x_6, x_7); +if (lean_obj_tag(x_232) == 0) +{ +lean_object* x_233; lean_object* x_234; +x_233 = lean_ctor_get(x_232, 0); +lean_inc(x_233); +lean_dec_ref(x_232); +lean_inc(x_7); +lean_inc_ref(x_6); +lean_inc(x_5); +lean_inc_ref(x_4); +x_234 = l_Lean_Compiler_LCNF_toLCNFType(x_233, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_234) == 0) +{ +lean_object* x_235; lean_object* x_236; +x_235 = lean_ctor_get(x_234, 0); +lean_inc(x_235); +lean_dec_ref(x_234); +lean_inc(x_7); +lean_inc_ref(x_6); +x_236 = l_Lean_Compiler_LCNF_toMonoType(x_235, x_6, x_7); +if (lean_obj_tag(x_236) == 0) +{ +lean_object* x_237; lean_object* x_238; +x_237 = lean_ctor_get(x_236, 0); +lean_inc(x_237); +lean_dec_ref(x_236); +lean_inc(x_7); +lean_inc_ref(x_6); +x_238 = l_Lean_IR_toIRType(x_237, x_6, x_7); +if (lean_obj_tag(x_238) == 0) +{ +lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; +x_239 = lean_ctor_get(x_3, 1); +lean_inc(x_239); +x_240 = lean_ctor_get(x_239, 1); +lean_inc(x_240); +x_241 = lean_ctor_get(x_240, 1); +lean_inc(x_241); +x_242 = lean_ctor_get(x_241, 1); +lean_inc(x_242); +x_243 = lean_ctor_get(x_238, 0); +lean_inc(x_243); +lean_dec_ref(x_238); +x_244 = lean_ctor_get(x_3, 0); +lean_inc(x_244); +lean_dec_ref(x_3); +x_245 = lean_ctor_get(x_239, 0); +lean_inc(x_245); +lean_dec(x_239); +x_246 = lean_ctor_get(x_240, 0); +lean_inc(x_246); +lean_dec(x_240); +x_247 = lean_ctor_get(x_241, 0); +lean_inc(x_247); +lean_dec(x_241); +x_248 = lean_ctor_get(x_242, 0); +lean_inc(x_248); +x_249 = lean_ctor_get(x_242, 1); +lean_inc(x_249); +if (lean_is_exclusive(x_242)) { + lean_ctor_release(x_242, 0); + lean_ctor_release(x_242, 1); + x_250 = x_242; +} else { + lean_dec_ref(x_242); + x_250 = lean_box(0); +} +x_251 = lean_unsigned_to_nat(1u); +x_252 = lean_nat_add(x_226, x_251); +lean_dec(x_226); +x_253 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_253, 0, x_225); +lean_ctor_set(x_253, 1, x_252); +lean_ctor_set(x_253, 2, x_227); +switch (lean_obj_tag(x_243)) { +case 0: +{ +lean_object* x_265; lean_object* x_266; uint8_t x_267; uint8_t x_268; uint8_t x_269; lean_object* x_270; +lean_dec(x_250); +lean_dec(x_248); +x_265 = lean_unsigned_to_nat(8u); +lean_inc(x_1); +x_266 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_266, 0, x_265); +lean_ctor_set(x_266, 1, x_1); +lean_ctor_set(x_266, 2, x_243); +x_267 = lean_unbox(x_245); +lean_dec(x_245); +x_268 = lean_unbox(x_246); +lean_dec(x_246); +x_269 = lean_unbox(x_247); +lean_dec(x_247); +x_270 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_244, x_267, x_268, x_269, x_228, x_249, x_266, x_4, x_5, x_6, x_7); +x_254 = x_270; +goto block_264; +} +case 1: +{ +lean_object* x_271; uint8_t x_272; uint8_t x_273; uint8_t x_274; lean_object* x_275; +lean_dec(x_250); +lean_dec(x_245); +lean_inc(x_1); +x_271 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_271, 0, x_251); +lean_ctor_set(x_271, 1, x_1); +lean_ctor_set(x_271, 2, x_243); +x_272 = lean_unbox(x_246); +lean_dec(x_246); +x_273 = lean_unbox(x_247); +lean_dec(x_247); +x_274 = lean_unbox(x_248); +lean_dec(x_248); +x_275 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_244, x_228, x_272, x_273, x_274, x_249, x_271, x_4, x_5, x_6, x_7); +x_254 = x_275; +goto block_264; +} +case 2: +{ +lean_object* x_276; lean_object* x_277; uint8_t x_278; uint8_t x_279; uint8_t x_280; lean_object* x_281; +lean_dec(x_250); +lean_dec(x_246); +x_276 = lean_unsigned_to_nat(2u); +lean_inc(x_1); +x_277 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_277, 0, x_276); +lean_ctor_set(x_277, 1, x_1); +lean_ctor_set(x_277, 2, x_243); +x_278 = lean_unbox(x_245); +lean_dec(x_245); +x_279 = lean_unbox(x_247); +lean_dec(x_247); +x_280 = lean_unbox(x_248); +lean_dec(x_248); +x_281 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_244, x_278, x_228, x_279, x_280, x_249, x_277, x_4, x_5, x_6, x_7); +x_254 = x_281; +goto block_264; +} +case 3: +{ +lean_object* x_282; lean_object* x_283; uint8_t x_284; uint8_t x_285; uint8_t x_286; lean_object* x_287; +lean_dec(x_250); +lean_dec(x_247); +x_282 = lean_unsigned_to_nat(4u); +lean_inc(x_1); +x_283 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_283, 0, x_282); +lean_ctor_set(x_283, 1, x_1); +lean_ctor_set(x_283, 2, x_243); +x_284 = lean_unbox(x_245); +lean_dec(x_245); +x_285 = lean_unbox(x_246); +lean_dec(x_246); +x_286 = lean_unbox(x_248); +lean_dec(x_248); +x_287 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_244, x_284, x_285, x_228, x_286, x_249, x_283, x_4, x_5, x_6, x_7); +x_254 = x_287; +goto block_264; +} +case 4: +{ +lean_object* x_288; lean_object* x_289; uint8_t x_290; uint8_t x_291; uint8_t x_292; lean_object* x_293; +lean_dec(x_250); +lean_dec(x_248); +x_288 = lean_unsigned_to_nat(8u); +lean_inc(x_1); +x_289 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_289, 0, x_288); +lean_ctor_set(x_289, 1, x_1); +lean_ctor_set(x_289, 2, x_243); +x_290 = lean_unbox(x_245); +lean_dec(x_245); +x_291 = lean_unbox(x_246); +lean_dec(x_246); +x_292 = lean_unbox(x_247); +lean_dec(x_247); +x_293 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_244, x_290, x_291, x_292, x_228, x_249, x_289, x_4, x_5, x_6, x_7); +x_254 = x_293; +goto block_264; +} +case 5: +{ +lean_object* x_294; uint8_t x_295; uint8_t x_296; uint8_t x_297; uint8_t x_298; lean_object* x_299; +lean_dec(x_250); +lean_inc(x_1); +x_294 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_294, 0, x_1); +x_295 = lean_unbox(x_245); +lean_dec(x_245); +x_296 = lean_unbox(x_246); +lean_dec(x_246); +x_297 = lean_unbox(x_247); +lean_dec(x_247); +x_298 = lean_unbox(x_248); +lean_dec(x_248); +x_299 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_244, x_295, x_296, x_297, x_298, x_249, x_294, x_4, x_5, x_6, x_7); +x_254 = x_299; +goto block_264; +} +case 6: +{ +lean_object* x_300; uint8_t x_301; uint8_t x_302; uint8_t x_303; uint8_t x_304; lean_object* x_305; +lean_dec(x_250); +x_300 = lean_box(0); +x_301 = lean_unbox(x_245); +lean_dec(x_245); +x_302 = lean_unbox(x_246); +lean_dec(x_246); +x_303 = lean_unbox(x_247); +lean_dec(x_247); +x_304 = lean_unbox(x_248); +lean_dec(x_248); +x_305 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_244, x_301, x_302, x_303, x_304, x_249, x_300, x_4, x_5, x_6, x_7); +x_254 = x_305; +goto block_264; +} +case 9: +{ +lean_object* x_306; lean_object* x_307; uint8_t x_308; uint8_t x_309; uint8_t x_310; lean_object* x_311; +lean_dec(x_250); +lean_dec(x_247); +x_306 = lean_unsigned_to_nat(4u); +lean_inc(x_1); +x_307 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_307, 0, x_306); +lean_ctor_set(x_307, 1, x_1); +lean_ctor_set(x_307, 2, x_243); +x_308 = lean_unbox(x_245); +lean_dec(x_245); +x_309 = lean_unbox(x_246); +lean_dec(x_246); +x_310 = lean_unbox(x_248); +lean_dec(x_248); +x_311 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_244, x_308, x_309, x_228, x_310, x_249, x_307, x_4, x_5, x_6, x_7); +x_254 = x_311; +goto block_264; +} +case 10: +{ +lean_object* x_312; lean_object* x_313; uint8_t x_314; uint8_t x_315; uint8_t x_316; uint8_t x_317; lean_object* x_318; +x_312 = lean_nat_add(x_249, x_251); +if (lean_is_scalar(x_250)) { + x_313 = lean_alloc_ctor(1, 2, 0); +} else { + x_313 = x_250; + lean_ctor_set_tag(x_313, 1); +} +lean_ctor_set(x_313, 0, x_249); +lean_ctor_set(x_313, 1, x_243); +x_314 = lean_unbox(x_245); +lean_dec(x_245); +x_315 = lean_unbox(x_246); +lean_dec(x_246); +x_316 = lean_unbox(x_247); +lean_dec(x_247); +x_317 = lean_unbox(x_248); +lean_dec(x_248); +x_318 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_244, x_314, x_315, x_316, x_317, x_312, x_313, x_4, x_5, x_6, x_7); +x_254 = x_318; +goto block_264; } +case 11: +{ +lean_object* x_319; lean_object* x_320; uint8_t x_321; uint8_t x_322; uint8_t x_323; uint8_t x_324; lean_object* x_325; +x_319 = lean_nat_add(x_249, x_251); +if (lean_is_scalar(x_250)) { + x_320 = lean_alloc_ctor(1, 2, 0); +} else { + x_320 = x_250; + lean_ctor_set_tag(x_320, 1); } +lean_ctor_set(x_320, 0, x_249); +lean_ctor_set(x_320, 1, x_243); +x_321 = lean_unbox(x_245); +lean_dec(x_245); +x_322 = lean_unbox(x_246); +lean_dec(x_246); +x_323 = lean_unbox(x_247); +lean_dec(x_247); +x_324 = lean_unbox(x_248); +lean_dec(x_248); +x_325 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_244, x_321, x_322, x_323, x_324, x_319, x_320, x_4, x_5, x_6, x_7); +x_254 = x_325; +goto block_264; } case 13: { -lean_object* x_122; uint8_t x_123; uint8_t x_124; uint8_t x_125; uint8_t x_126; lean_object* x_127; -lean_free_object(x_27); -x_122 = lean_box(4); -x_123 = lean_unbox(x_30); -lean_dec(x_30); -x_124 = lean_unbox(x_31); -lean_dec(x_31); -x_125 = lean_unbox(x_32); -lean_dec(x_32); -x_126 = lean_unbox(x_34); -lean_dec(x_34); -x_127 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_29, x_123, x_124, x_125, x_126, x_35, x_122, x_4, x_5, x_6, x_7); -x_38 = x_127; -goto block_52; +lean_object* x_326; uint8_t x_327; uint8_t x_328; uint8_t x_329; uint8_t x_330; lean_object* x_331; +lean_dec(x_250); +x_326 = lean_box(4); +x_327 = lean_unbox(x_245); +lean_dec(x_245); +x_328 = lean_unbox(x_246); +lean_dec(x_246); +x_329 = lean_unbox(x_247); +lean_dec(x_247); +x_330 = lean_unbox(x_248); +lean_dec(x_248); +x_331 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_244, x_327, x_328, x_329, x_330, x_249, x_326, x_4, x_5, x_6, x_7); +x_254 = x_331; +goto block_264; } default: { -lean_object* x_128; uint8_t x_129; uint8_t x_130; uint8_t x_131; uint8_t x_132; lean_object* x_133; -x_128 = lean_nat_add(x_35, x_36); -lean_ctor_set_tag(x_27, 1); -lean_ctor_set(x_27, 1, x_28); -lean_ctor_set(x_27, 0, x_35); -x_129 = lean_unbox(x_30); -lean_dec(x_30); -x_130 = lean_unbox(x_31); -lean_dec(x_31); -x_131 = lean_unbox(x_32); -lean_dec(x_32); -x_132 = lean_unbox(x_34); -lean_dec(x_34); -x_133 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_29, x_129, x_130, x_131, x_132, x_128, x_27, x_4, x_5, x_6, x_7); -x_38 = x_133; -goto block_52; +lean_object* x_332; lean_object* x_333; uint8_t x_334; uint8_t x_335; uint8_t x_336; uint8_t x_337; lean_object* x_338; +x_332 = lean_nat_add(x_249, x_251); +if (lean_is_scalar(x_250)) { + x_333 = lean_alloc_ctor(1, 2, 0); +} else { + x_333 = x_250; + lean_ctor_set_tag(x_333, 1); } +lean_ctor_set(x_333, 0, x_249); +lean_ctor_set(x_333, 1, x_243); +x_334 = lean_unbox(x_245); +lean_dec(x_245); +x_335 = lean_unbox(x_246); +lean_dec(x_246); +x_336 = lean_unbox(x_247); +lean_dec(x_247); +x_337 = lean_unbox(x_248); +lean_dec(x_248); +x_338 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_244, x_334, x_335, x_336, x_337, x_332, x_333, x_4, x_5, x_6, x_7); +x_254 = x_338; +goto block_264; } -block_52: -{ -if (lean_obj_tag(x_38) == 0) +} +block_264: { -uint8_t x_39; -x_39 = !lean_is_exclusive(x_38); -if (x_39 == 0) +if (lean_obj_tag(x_254) == 0) { -lean_object* x_40; -x_40 = lean_ctor_get(x_38, 0); -if (lean_obj_tag(x_40) == 0) +lean_object* x_255; lean_object* x_256; +x_255 = lean_ctor_get(x_254, 0); +lean_inc(x_255); +if (lean_is_exclusive(x_254)) { + lean_ctor_release(x_254, 0); + x_256 = x_254; +} else { + lean_dec_ref(x_254); + x_256 = lean_box(0); +} +if (lean_obj_tag(x_255) == 0) { -lean_object* x_41; -lean_dec_ref(x_2); +lean_object* x_257; lean_object* x_258; +lean_dec_ref(x_253); lean_dec(x_7); lean_dec_ref(x_6); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_1); -x_41 = lean_ctor_get(x_40, 0); -lean_inc(x_41); -lean_dec_ref(x_40); -lean_ctor_set(x_38, 0, x_41); -return x_38; +x_257 = lean_ctor_get(x_255, 0); +lean_inc(x_257); +lean_dec_ref(x_255); +if (lean_is_scalar(x_256)) { + x_258 = lean_alloc_ctor(0, 1, 0); +} else { + x_258 = x_256; +} +lean_ctor_set(x_258, 0, x_257); +return x_258; } else { -lean_object* x_42; -lean_free_object(x_38); -x_42 = lean_ctor_get(x_40, 0); -lean_inc(x_42); -lean_dec_ref(x_40); -x_3 = x_42; +lean_object* x_259; +lean_dec(x_256); +x_259 = lean_ctor_get(x_255, 0); +lean_inc(x_259); +lean_dec_ref(x_255); +x_2 = x_253; +x_3 = x_259; goto _start; } } else { -lean_object* x_44; -x_44 = lean_ctor_get(x_38, 0); -lean_inc(x_44); -lean_dec(x_38); -if (lean_obj_tag(x_44) == 0) -{ -lean_object* x_45; lean_object* x_46; -lean_dec_ref(x_2); +lean_object* x_261; lean_object* x_262; lean_object* x_263; +lean_dec_ref(x_253); lean_dec(x_7); lean_dec_ref(x_6); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_1); -x_45 = lean_ctor_get(x_44, 0); -lean_inc(x_45); -lean_dec_ref(x_44); -x_46 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_46, 0, x_45); -return x_46; +x_261 = lean_ctor_get(x_254, 0); +lean_inc(x_261); +if (lean_is_exclusive(x_254)) { + lean_ctor_release(x_254, 0); + x_262 = x_254; +} else { + lean_dec_ref(x_254); + x_262 = lean_box(0); +} +if (lean_is_scalar(x_262)) { + x_263 = lean_alloc_ctor(1, 1, 0); +} else { + x_263 = x_262; +} +lean_ctor_set(x_263, 0, x_261); +return x_263; +} +} } else { -lean_object* x_47; -x_47 = lean_ctor_get(x_44, 0); -lean_inc(x_47); -lean_dec_ref(x_44); -x_3 = x_47; -goto _start; +lean_object* x_339; lean_object* x_340; lean_object* x_341; +lean_dec(x_227); +lean_dec(x_226); +lean_dec_ref(x_225); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec_ref(x_3); +lean_dec(x_1); +x_339 = lean_ctor_get(x_238, 0); +lean_inc(x_339); +if (lean_is_exclusive(x_238)) { + lean_ctor_release(x_238, 0); + x_340 = x_238; +} else { + lean_dec_ref(x_238); + x_340 = lean_box(0); +} +if (lean_is_scalar(x_340)) { + x_341 = lean_alloc_ctor(1, 1, 0); +} else { + x_341 = x_340; } +lean_ctor_set(x_341, 0, x_339); +return x_341; } } else { -uint8_t x_49; -lean_dec_ref(x_2); +lean_object* x_342; lean_object* x_343; lean_object* x_344; +lean_dec(x_227); +lean_dec(x_226); +lean_dec_ref(x_225); lean_dec(x_7); lean_dec_ref(x_6); lean_dec(x_5); lean_dec_ref(x_4); +lean_dec_ref(x_3); lean_dec(x_1); -x_49 = !lean_is_exclusive(x_38); -if (x_49 == 0) -{ -return x_38; +x_342 = lean_ctor_get(x_236, 0); +lean_inc(x_342); +if (lean_is_exclusive(x_236)) { + lean_ctor_release(x_236, 0); + x_343 = x_236; +} else { + lean_dec_ref(x_236); + x_343 = lean_box(0); +} +if (lean_is_scalar(x_343)) { + x_344 = lean_alloc_ctor(1, 1, 0); +} else { + x_344 = x_343; +} +lean_ctor_set(x_344, 0, x_342); +return x_344; +} } else { -lean_object* x_50; lean_object* x_51; -x_50 = lean_ctor_get(x_38, 0); -lean_inc(x_50); -lean_dec(x_38); -x_51 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_51, 0, x_50); -return x_51; +lean_object* x_345; lean_object* x_346; lean_object* x_347; +lean_dec(x_227); +lean_dec(x_226); +lean_dec_ref(x_225); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec_ref(x_3); +lean_dec(x_1); +x_345 = lean_ctor_get(x_234, 0); +lean_inc(x_345); +if (lean_is_exclusive(x_234)) { + lean_ctor_release(x_234, 0); + x_346 = x_234; +} else { + lean_dec_ref(x_234); + x_346 = lean_box(0); } +if (lean_is_scalar(x_346)) { + x_347 = lean_alloc_ctor(1, 1, 0); +} else { + x_347 = x_346; } +lean_ctor_set(x_347, 0, x_345); +return x_347; } } else { -lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; -x_134 = lean_ctor_get(x_27, 0); -x_135 = lean_ctor_get(x_27, 1); -lean_inc(x_135); -lean_inc(x_134); -lean_dec(x_27); -x_136 = lean_unsigned_to_nat(1u); -x_137 = lean_nat_add(x_11, x_136); -lean_dec(x_11); -lean_ctor_set(x_2, 1, x_137); -switch (lean_obj_tag(x_28)) { -case 0: +lean_object* x_348; lean_object* x_349; lean_object* x_350; +lean_dec(x_227); +lean_dec(x_226); +lean_dec_ref(x_225); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec_ref(x_3); +lean_dec(x_1); +x_348 = lean_ctor_get(x_232, 0); +lean_inc(x_348); +if (lean_is_exclusive(x_232)) { + lean_ctor_release(x_232, 0); + x_349 = x_232; +} else { + lean_dec_ref(x_232); + x_349 = lean_box(0); +} +if (lean_is_scalar(x_349)) { + x_350 = lean_alloc_ctor(1, 1, 0); +} else { + x_350 = x_349; +} +lean_ctor_set(x_350, 0, x_348); +return x_350; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_usize_dec_lt(x_2, x_1); +if (x_5 == 0) { -lean_object* x_149; lean_object* x_150; uint8_t x_151; uint8_t x_152; uint8_t x_153; lean_object* x_154; -lean_dec(x_134); -x_149 = lean_unsigned_to_nat(8u); -lean_inc(x_1); -x_150 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_150, 0, x_149); -lean_ctor_set(x_150, 1, x_1); -lean_ctor_set(x_150, 2, x_28); -x_151 = lean_unbox(x_30); -lean_dec(x_30); -x_152 = lean_unbox(x_31); -lean_dec(x_31); -x_153 = lean_unbox(x_32); -lean_dec(x_32); -x_154 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_29, x_151, x_152, x_153, x_13, x_135, x_150, x_4, x_5, x_6, x_7); -x_138 = x_154; -goto block_148; +lean_object* x_6; +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_3); +lean_ctor_set(x_6, 1, x_4); +return x_6; } +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_7 = lean_array_uget(x_3, x_2); +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_array_uset(x_3, x_2, x_8); +switch (lean_obj_tag(x_7)) { case 1: { -lean_object* x_155; uint8_t x_156; uint8_t x_157; uint8_t x_158; lean_object* x_159; -lean_dec(x_30); -lean_inc(x_1); -x_155 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_155, 0, x_136); -lean_ctor_set(x_155, 1, x_1); -lean_ctor_set(x_155, 2, x_28); -x_156 = lean_unbox(x_31); -lean_dec(x_31); -x_157 = lean_unbox(x_32); -lean_dec(x_32); -x_158 = lean_unbox(x_134); -lean_dec(x_134); -x_159 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_29, x_13, x_156, x_157, x_158, x_135, x_155, x_4, x_5, x_6, x_7); -x_138 = x_159; -goto block_148; +x_10 = x_7; +x_11 = x_4; +goto block_16; } case 2: { -lean_object* x_160; lean_object* x_161; uint8_t x_162; uint8_t x_163; uint8_t x_164; lean_object* x_165; -lean_dec(x_31); -x_160 = lean_unsigned_to_nat(2u); -lean_inc(x_1); -x_161 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_161, 0, x_160); -lean_ctor_set(x_161, 1, x_1); -lean_ctor_set(x_161, 2, x_28); -x_162 = lean_unbox(x_30); -lean_dec(x_30); -x_163 = lean_unbox(x_32); -lean_dec(x_32); -x_164 = lean_unbox(x_134); -lean_dec(x_134); -x_165 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_29, x_162, x_13, x_163, x_164, x_135, x_161, x_4, x_5, x_6, x_7); -x_138 = x_165; -goto block_148; -} -case 3: +uint8_t x_17; +x_17 = !lean_is_exclusive(x_7); +if (x_17 == 0) { -lean_object* x_166; lean_object* x_167; uint8_t x_168; uint8_t x_169; uint8_t x_170; lean_object* x_171; -lean_dec(x_32); -x_166 = lean_unsigned_to_nat(4u); -lean_inc(x_1); -x_167 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_167, 0, x_166); -lean_ctor_set(x_167, 1, x_1); -lean_ctor_set(x_167, 2, x_28); -x_168 = lean_unbox(x_30); -lean_dec(x_30); -x_169 = lean_unbox(x_31); -lean_dec(x_31); -x_170 = lean_unbox(x_134); -lean_dec(x_134); -x_171 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_29, x_168, x_169, x_13, x_170, x_135, x_167, x_4, x_5, x_6, x_7); -x_138 = x_171; -goto block_148; +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_7, 0); +lean_dec(x_18); +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_add(x_4, x_19); +lean_ctor_set(x_7, 0, x_4); +x_10 = x_7; +x_11 = x_20; +goto block_16; } -case 4: +else { -lean_object* x_172; lean_object* x_173; uint8_t x_174; uint8_t x_175; uint8_t x_176; lean_object* x_177; -lean_dec(x_134); -x_172 = lean_unsigned_to_nat(8u); -lean_inc(x_1); -x_173 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_173, 0, x_172); -lean_ctor_set(x_173, 1, x_1); -lean_ctor_set(x_173, 2, x_28); -x_174 = lean_unbox(x_30); -lean_dec(x_30); -x_175 = lean_unbox(x_31); -lean_dec(x_31); -x_176 = lean_unbox(x_32); -lean_dec(x_32); -x_177 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_29, x_174, x_175, x_176, x_13, x_135, x_173, x_4, x_5, x_6, x_7); -x_138 = x_177; -goto block_148; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +lean_dec(x_7); +x_21 = lean_unsigned_to_nat(1u); +x_22 = lean_nat_add(x_4, x_21); +x_23 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_23, 0, x_4); +x_10 = x_23; +x_11 = x_22; +goto block_16; } -case 5: -{ -lean_object* x_178; uint8_t x_179; uint8_t x_180; uint8_t x_181; uint8_t x_182; lean_object* x_183; -lean_inc(x_1); -x_178 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_178, 0, x_1); -x_179 = lean_unbox(x_30); -lean_dec(x_30); -x_180 = lean_unbox(x_31); -lean_dec(x_31); -x_181 = lean_unbox(x_32); -lean_dec(x_32); -x_182 = lean_unbox(x_134); -lean_dec(x_134); -x_183 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_29, x_179, x_180, x_181, x_182, x_135, x_178, x_4, x_5, x_6, x_7); -x_138 = x_183; -goto block_148; } -case 6: +case 3: { -lean_object* x_184; uint8_t x_185; uint8_t x_186; uint8_t x_187; uint8_t x_188; lean_object* x_189; -x_184 = lean_box(0); -x_185 = lean_unbox(x_30); -lean_dec(x_30); -x_186 = lean_unbox(x_31); -lean_dec(x_31); -x_187 = lean_unbox(x_32); -lean_dec(x_32); -x_188 = lean_unbox(x_134); -lean_dec(x_134); -x_189 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_29, x_185, x_186, x_187, x_188, x_135, x_184, x_4, x_5, x_6, x_7); -x_138 = x_189; -goto block_148; +x_10 = x_7; +x_11 = x_4; +goto block_16; } -case 9: +default: { -lean_object* x_190; lean_object* x_191; uint8_t x_192; uint8_t x_193; uint8_t x_194; lean_object* x_195; -lean_dec(x_32); -x_190 = lean_unsigned_to_nat(4u); -lean_inc(x_1); -x_191 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_191, 0, x_190); -lean_ctor_set(x_191, 1, x_1); -lean_ctor_set(x_191, 2, x_28); -x_192 = lean_unbox(x_30); -lean_dec(x_30); -x_193 = lean_unbox(x_31); -lean_dec(x_31); -x_194 = lean_unbox(x_134); -lean_dec(x_134); -x_195 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_29, x_192, x_193, x_13, x_194, x_135, x_191, x_4, x_5, x_6, x_7); -x_138 = x_195; -goto block_148; +x_10 = x_7; +x_11 = x_4; +goto block_16; } -case 10: -{ -lean_object* x_196; lean_object* x_197; -lean_dec_ref(x_28); -x_196 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___closed__0; -lean_inc(x_7); -lean_inc_ref(x_6); -lean_inc(x_5); -lean_inc_ref(x_4); -x_197 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2(x_196, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_197) == 0) -{ -lean_object* x_198; uint8_t x_199; uint8_t x_200; uint8_t x_201; uint8_t x_202; lean_object* x_203; -x_198 = lean_ctor_get(x_197, 0); -lean_inc(x_198); -lean_dec_ref(x_197); -x_199 = lean_unbox(x_30); -lean_dec(x_30); -x_200 = lean_unbox(x_31); -lean_dec(x_31); -x_201 = lean_unbox(x_32); -lean_dec(x_32); -x_202 = lean_unbox(x_134); -lean_dec(x_134); -x_203 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_29, x_199, x_200, x_201, x_202, x_135, x_198, x_4, x_5, x_6, x_7); -x_138 = x_203; -goto block_148; } -else +block_16: { -lean_object* x_204; lean_object* x_205; lean_object* x_206; -lean_dec_ref(x_2); -lean_dec(x_135); -lean_dec(x_134); -lean_dec(x_32); -lean_dec(x_31); -lean_dec(x_30); -lean_dec(x_29); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_1); -x_204 = lean_ctor_get(x_197, 0); -lean_inc(x_204); -if (lean_is_exclusive(x_197)) { - lean_ctor_release(x_197, 0); - x_205 = x_197; -} else { - lean_dec_ref(x_197); - x_205 = lean_box(0); +size_t x_12; size_t x_13; lean_object* x_14; +x_12 = 1; +x_13 = lean_usize_add(x_2, x_12); +x_14 = lean_array_uset(x_9, x_2, x_10); +x_2 = x_13; +x_3 = x_14; +x_4 = x_11; +goto _start; } -if (lean_is_scalar(x_205)) { - x_206 = lean_alloc_ctor(1, 1, 0); -} else { - x_206 = x_205; } -lean_ctor_set(x_206, 0, x_204); -return x_206; } } -case 11: +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___lam__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -lean_object* x_207; lean_object* x_208; -lean_dec_ref(x_28); -x_207 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___closed__0; -lean_inc(x_7); -lean_inc_ref(x_6); -lean_inc(x_5); -lean_inc_ref(x_4); -x_208 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2(x_207, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_208) == 0) -{ -lean_object* x_209; uint8_t x_210; uint8_t x_211; uint8_t x_212; uint8_t x_213; lean_object* x_214; -x_209 = lean_ctor_get(x_208, 0); -lean_inc(x_209); -lean_dec_ref(x_208); -x_210 = lean_unbox(x_30); +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_14 = lean_mk_empty_array_with_capacity(x_1); +x_15 = lean_unsigned_to_nat(0u); +x_16 = lean_nat_add(x_2, x_1); +x_17 = l_Array_toSubarray___redArg(x_7, x_2, x_16); +x_18 = lean_box(x_3); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_15); +x_20 = lean_box(x_3); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +x_22 = lean_box(x_3); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_21); +x_24 = lean_box(x_3); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_23); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_14); +lean_ctor_set(x_26, 1, x_25); +x_27 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg(x_15, x_17, x_26, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; size_t x_40; size_t x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_81; +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +if (lean_is_exclusive(x_27)) { + lean_ctor_release(x_27, 0); + x_29 = x_27; +} else { + lean_dec_ref(x_27); + x_29 = lean_box(0); +} +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +x_33 = lean_ctor_get(x_32, 1); +lean_inc(x_33); +x_34 = lean_ctor_get(x_28, 0); +lean_inc(x_34); +lean_dec(x_28); +x_35 = lean_ctor_get(x_30, 0); +lean_inc(x_35); lean_dec(x_30); -x_211 = lean_unbox(x_31); +x_36 = lean_ctor_get(x_31, 0); +lean_inc(x_36); lean_dec(x_31); -x_212 = lean_unbox(x_32); +x_37 = lean_ctor_get(x_32, 0); +lean_inc(x_37); lean_dec(x_32); -x_213 = lean_unbox(x_134); -lean_dec(x_134); -x_214 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_29, x_210, x_211, x_212, x_213, x_135, x_209, x_4, x_5, x_6, x_7); -x_138 = x_214; -goto block_148; +x_38 = lean_ctor_get(x_33, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_33, 1); +lean_inc(x_39); +lean_dec(x_33); +x_40 = lean_array_size(x_34); +x_41 = 0; +lean_inc(x_39); +x_42 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3(x_40, x_41, x_34, x_39); +x_43 = lean_ctor_get(x_42, 0); +lean_inc(x_43); +x_44 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +if (lean_is_exclusive(x_42)) { + lean_ctor_release(x_42, 0); + lean_ctor_release(x_42, 1); + x_45 = x_42; +} else { + lean_dec_ref(x_42); + x_45 = lean_box(0); +} +x_46 = lean_nat_sub(x_44, x_39); +lean_dec(x_44); +x_81 = lean_unbox(x_38); +lean_dec(x_38); +if (x_81 == 0) +{ +x_72 = x_43; +x_73 = x_15; +x_74 = lean_box(0); +goto block_80; } else { -lean_object* x_215; lean_object* x_216; lean_object* x_217; -lean_dec_ref(x_2); -lean_dec(x_135); -lean_dec(x_134); -lean_dec(x_32); -lean_dec(x_31); -lean_dec(x_30); -lean_dec(x_29); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_1); -x_215 = lean_ctor_get(x_208, 0); -lean_inc(x_215); -if (lean_is_exclusive(x_208)) { - lean_ctor_release(x_208, 0); - x_216 = x_208; +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_82 = lean_unsigned_to_nat(8u); +lean_inc_ref(x_6); +x_83 = lean_apply_3(x_6, x_43, x_82, x_15); +x_84 = lean_ctor_get(x_83, 0); +lean_inc(x_84); +x_85 = lean_ctor_get(x_83, 1); +lean_inc(x_85); +lean_dec_ref(x_83); +x_72 = x_84; +x_73 = x_85; +x_74 = lean_box(0); +goto block_80; +} +block_53: +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_50, 0, x_4); +lean_ctor_set(x_50, 1, x_5); +lean_ctor_set(x_50, 2, x_39); +lean_ctor_set(x_50, 3, x_46); +lean_ctor_set(x_50, 4, x_48); +if (lean_is_scalar(x_45)) { + x_51 = lean_alloc_ctor(0, 2, 0); } else { - lean_dec_ref(x_208); - x_216 = lean_box(0); + x_51 = x_45; } -if (lean_is_scalar(x_216)) { - x_217 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_47); +if (lean_is_scalar(x_29)) { + x_52 = lean_alloc_ctor(0, 1, 0); } else { - x_217 = x_216; -} -lean_ctor_set(x_217, 0, x_215); -return x_217; + x_52 = x_29; } +lean_ctor_set(x_52, 0, x_51); +return x_52; } -case 13: +block_62: { -lean_object* x_218; uint8_t x_219; uint8_t x_220; uint8_t x_221; uint8_t x_222; lean_object* x_223; -x_218 = lean_box(4); -x_219 = lean_unbox(x_30); -lean_dec(x_30); -x_220 = lean_unbox(x_31); -lean_dec(x_31); -x_221 = lean_unbox(x_32); -lean_dec(x_32); -x_222 = lean_unbox(x_134); -lean_dec(x_134); -x_223 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_29, x_219, x_220, x_221, x_222, x_135, x_218, x_4, x_5, x_6, x_7); -x_138 = x_223; -goto block_148; +uint8_t x_57; +x_57 = lean_unbox(x_35); +lean_dec(x_35); +if (x_57 == 0) +{ +lean_dec_ref(x_6); +x_47 = x_54; +x_48 = x_55; +x_49 = lean_box(0); +goto block_53; } -default: +else { -lean_object* x_224; lean_object* x_225; uint8_t x_226; uint8_t x_227; uint8_t x_228; uint8_t x_229; lean_object* x_230; -x_224 = lean_nat_add(x_135, x_136); -x_225 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_225, 0, x_135); -lean_ctor_set(x_225, 1, x_28); -x_226 = lean_unbox(x_30); -lean_dec(x_30); -x_227 = lean_unbox(x_31); -lean_dec(x_31); -x_228 = lean_unbox(x_32); -lean_dec(x_32); -x_229 = lean_unbox(x_134); -lean_dec(x_134); -x_230 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_29, x_226, x_227, x_228, x_229, x_224, x_225, x_4, x_5, x_6, x_7); -x_138 = x_230; -goto block_148; +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_58 = lean_unsigned_to_nat(1u); +x_59 = lean_apply_3(x_6, x_54, x_58, x_55); +x_60 = lean_ctor_get(x_59, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_59, 1); +lean_inc(x_61); +lean_dec_ref(x_59); +x_47 = x_60; +x_48 = x_61; +x_49 = lean_box(0); +goto block_53; } } -block_148: +block_71: { -if (lean_obj_tag(x_138) == 0) +uint8_t x_66; +x_66 = lean_unbox(x_36); +lean_dec(x_36); +if (x_66 == 0) { -lean_object* x_139; lean_object* x_140; -x_139 = lean_ctor_get(x_138, 0); -lean_inc(x_139); -if (lean_is_exclusive(x_138)) { - lean_ctor_release(x_138, 0); - x_140 = x_138; -} else { - lean_dec_ref(x_138); - x_140 = lean_box(0); +x_54 = x_63; +x_55 = x_64; +x_56 = lean_box(0); +goto block_62; } -if (lean_obj_tag(x_139) == 0) +else { -lean_object* x_141; lean_object* x_142; -lean_dec_ref(x_2); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_1); -x_141 = lean_ctor_get(x_139, 0); -lean_inc(x_141); -lean_dec_ref(x_139); -if (lean_is_scalar(x_140)) { - x_142 = lean_alloc_ctor(0, 1, 0); -} else { - x_142 = x_140; +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_67 = lean_unsigned_to_nat(2u); +lean_inc_ref(x_6); +x_68 = lean_apply_3(x_6, x_63, x_67, x_64); +x_69 = lean_ctor_get(x_68, 0); +lean_inc(x_69); +x_70 = lean_ctor_get(x_68, 1); +lean_inc(x_70); +lean_dec_ref(x_68); +x_54 = x_69; +x_55 = x_70; +x_56 = lean_box(0); +goto block_62; } -lean_ctor_set(x_142, 0, x_141); -return x_142; } -else +block_80: { -lean_object* x_143; -lean_dec(x_140); -x_143 = lean_ctor_get(x_139, 0); -lean_inc(x_143); -lean_dec_ref(x_139); -x_3 = x_143; -goto _start; -} +uint8_t x_75; +x_75 = lean_unbox(x_37); +lean_dec(x_37); +if (x_75 == 0) +{ +x_63 = x_72; +x_64 = x_73; +x_65 = lean_box(0); +goto block_71; } else { -lean_object* x_145; lean_object* x_146; lean_object* x_147; -lean_dec_ref(x_2); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_1); -x_145 = lean_ctor_get(x_138, 0); -lean_inc(x_145); -if (lean_is_exclusive(x_138)) { - lean_ctor_release(x_138, 0); - x_146 = x_138; -} else { - lean_dec_ref(x_138); - x_146 = lean_box(0); -} -if (lean_is_scalar(x_146)) { - x_147 = lean_alloc_ctor(1, 1, 0); -} else { - x_147 = x_146; -} -lean_ctor_set(x_147, 0, x_145); -return x_147; -} +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_76 = lean_unsigned_to_nat(4u); +lean_inc_ref(x_6); +x_77 = lean_apply_3(x_6, x_72, x_76, x_73); +x_78 = lean_ctor_get(x_77, 0); +lean_inc(x_78); +x_79 = lean_ctor_get(x_77, 1); +lean_inc(x_79); +lean_dec_ref(x_77); +x_63 = x_78; +x_64 = x_79; +x_65 = lean_box(0); +goto block_71; } } } else { -uint8_t x_231; -lean_free_object(x_2); -lean_dec(x_12); -lean_dec(x_11); -lean_dec_ref(x_10); -lean_dec(x_7); +uint8_t x_86; lean_dec_ref(x_6); lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -lean_dec(x_1); -x_231 = !lean_is_exclusive(x_23); -if (x_231 == 0) +lean_dec(x_4); +x_86 = !lean_is_exclusive(x_27); +if (x_86 == 0) { -return x_23; +return x_27; } else { -lean_object* x_232; lean_object* x_233; -x_232 = lean_ctor_get(x_23, 0); -lean_inc(x_232); -lean_dec(x_23); -x_233 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_233, 0, x_232); -return x_233; +lean_object* x_87; lean_object* x_88; +x_87 = lean_ctor_get(x_27, 0); +lean_inc(x_87); +lean_dec(x_27); +x_88 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_88, 0, x_87); +return x_88; } } } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: { -uint8_t x_234; -lean_free_object(x_2); -lean_dec(x_12); -lean_dec(x_11); -lean_dec_ref(x_10); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec_ref(x_3); +uint8_t x_14; lean_object* x_15; +x_14 = lean_unbox(x_3); +x_15 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___lam__1(x_1, x_2, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec_ref(x_8); lean_dec(x_1); -x_234 = !lean_is_exclusive(x_21); -if (x_234 == 0) -{ -return x_21; +return x_15; } -else +} +static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___closed__2() { +_start: { -lean_object* x_235; lean_object* x_236; -x_235 = lean_ctor_get(x_21, 0); -lean_inc(x_235); -lean_dec(x_21); -x_236 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_236, 0, x_235); -return x_236; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(1); +x_2 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___closed__0; +x_3 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__5; +x_4 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_1); +return x_4; } } +static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___closed__3() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; lean_object* x_9; +x_1 = 1; +x_2 = lean_unsigned_to_nat(0u); +x_3 = lean_box(0); +x_4 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__11; +x_5 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___closed__2; +x_6 = lean_box(1); +x_7 = 0; +x_8 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__10; +x_9 = lean_alloc_ctor(0, 7, 4); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_6); +lean_ctor_set(x_9, 2, x_5); +lean_ctor_set(x_9, 3, x_4); +lean_ctor_set(x_9, 4, x_3); +lean_ctor_set(x_9, 5, x_2); +lean_ctor_set(x_9, 6, x_3); +lean_ctor_set_uint8(x_9, sizeof(void*)*7, x_7); +lean_ctor_set_uint8(x_9, sizeof(void*)*7 + 1, x_7); +lean_ctor_set_uint8(x_9, sizeof(void*)*7 + 2, x_7); +lean_ctor_set_uint8(x_9, sizeof(void*)*7 + 3, x_1); +return x_9; } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -uint8_t x_237; -lean_free_object(x_2); -lean_dec(x_12); +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; +x_11 = lean_st_ref_get(x_3); +x_12 = lean_ctor_get(x_11, 0); +lean_inc_ref(x_12); lean_dec(x_11); -lean_dec_ref(x_10); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -lean_dec(x_1); -x_237 = !lean_is_exclusive(x_19); -if (x_237 == 0) +x_13 = 0; +lean_inc(x_1); +x_14 = l_Lean_Environment_find_x3f(x_12, x_1, x_13); +if (lean_obj_tag(x_14) == 1) { -return x_19; -} -else +lean_object* x_15; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +lean_dec_ref(x_14); +if (lean_obj_tag(x_15) == 6) { -lean_object* x_238; lean_object* x_239; -x_238 = lean_ctor_get(x_19, 0); -lean_inc(x_238); -lean_dec(x_19); -x_239 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_239, 0, x_238); -return x_239; -} -} -} -else +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_16 = lean_ctor_get(x_15, 0); +lean_inc_ref(x_16); +lean_dec_ref(x_15); +x_17 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___closed__1; +x_18 = lean_st_mk_ref(x_17); +x_19 = lean_ctor_get(x_16, 0); +lean_inc_ref(x_19); +x_20 = lean_ctor_get(x_16, 2); +lean_inc(x_20); +x_21 = lean_ctor_get(x_16, 3); +lean_inc(x_21); +x_22 = lean_ctor_get(x_16, 4); +lean_inc(x_22); +lean_dec_ref(x_16); +x_23 = lean_ctor_get(x_19, 2); +lean_inc_ref(x_23); +lean_dec_ref(x_19); +x_24 = lean_alloc_closure((void*)(l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___lam__0___boxed), 3, 0); +x_25 = lean_box(x_13); +x_26 = lean_alloc_closure((void*)(l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___lam__1___boxed), 13, 6); +lean_closure_set(x_26, 0, x_22); +lean_closure_set(x_26, 1, x_21); +lean_closure_set(x_26, 2, x_25); +lean_closure_set(x_26, 3, x_1); +lean_closure_set(x_26, 4, x_20); +lean_closure_set(x_26, 5, x_24); +x_27 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___closed__3; +lean_inc(x_18); +x_28 = l_Lean_Meta_forallTelescopeReducing___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__4___redArg(x_23, x_26, x_13, x_13, x_27, x_18, x_2, x_3); +if (lean_obj_tag(x_28) == 0) { -uint8_t x_240; -lean_free_object(x_2); -lean_dec(x_12); -lean_dec(x_11); -lean_dec_ref(x_10); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -lean_dec(x_1); -x_240 = !lean_is_exclusive(x_17); -if (x_240 == 0) +uint8_t x_29; +x_29 = !lean_is_exclusive(x_28); +if (x_29 == 0) { -return x_17; +lean_object* x_30; +x_30 = lean_st_ref_get(x_18); +lean_dec(x_18); +lean_dec(x_30); +return x_28; } else { -lean_object* x_241; lean_object* x_242; -x_241 = lean_ctor_get(x_17, 0); -lean_inc(x_241); -lean_dec(x_17); -x_242 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_242, 0, x_241); -return x_242; +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_28, 0); +lean_inc(x_31); +lean_dec(x_28); +x_32 = lean_st_ref_get(x_18); +lean_dec(x_18); +lean_dec(x_32); +x_33 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_33, 0, x_31); +return x_33; } } +else +{ +lean_dec(x_18); +return x_28; } } else { -lean_object* x_243; lean_object* x_244; lean_object* x_245; uint8_t x_246; -x_243 = lean_ctor_get(x_2, 0); -x_244 = lean_ctor_get(x_2, 1); -x_245 = lean_ctor_get(x_2, 2); -lean_inc(x_245); -lean_inc(x_244); -lean_inc(x_243); -lean_dec(x_2); -x_246 = lean_nat_dec_lt(x_244, x_245); -if (x_246 == 0) -{ -lean_object* x_247; -lean_dec(x_245); -lean_dec(x_244); -lean_dec_ref(x_243); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec(x_5); -lean_dec_ref(x_4); +lean_dec(x_15); lean_dec(x_1); -x_247 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_247, 0, x_3); -return x_247; +x_5 = x_2; +x_6 = x_3; +x_7 = lean_box(0); +goto block_10; +} } else { -lean_object* x_248; lean_object* x_249; lean_object* x_250; -x_248 = lean_array_fget_borrowed(x_243, x_244); -x_249 = l_Lean_Expr_fvarId_x21(x_248); -lean_inc_ref(x_4); -x_250 = l_Lean_FVarId_getType___redArg(x_249, x_4, x_6, x_7); -if (lean_obj_tag(x_250) == 0) -{ -lean_object* x_251; lean_object* x_252; -x_251 = lean_ctor_get(x_250, 0); -lean_inc(x_251); -lean_dec_ref(x_250); -lean_inc(x_7); -lean_inc_ref(x_6); -lean_inc(x_5); -lean_inc_ref(x_4); -x_252 = l_Lean_Compiler_LCNF_toLCNFType(x_251, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_252) == 0) -{ -lean_object* x_253; lean_object* x_254; -x_253 = lean_ctor_get(x_252, 0); -lean_inc(x_253); -lean_dec_ref(x_252); -lean_inc(x_7); -lean_inc_ref(x_6); -x_254 = l_Lean_Compiler_LCNF_toMonoType(x_253, x_6, x_7); -if (lean_obj_tag(x_254) == 0) -{ -lean_object* x_255; lean_object* x_256; -x_255 = lean_ctor_get(x_254, 0); -lean_inc(x_255); -lean_dec_ref(x_254); -lean_inc(x_7); -lean_inc_ref(x_6); -x_256 = l_Lean_IR_toIRType(x_255, x_6, x_7); -if (lean_obj_tag(x_256) == 0) -{ -lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; -x_257 = lean_ctor_get(x_3, 1); -lean_inc(x_257); -x_258 = lean_ctor_get(x_257, 1); -lean_inc(x_258); -x_259 = lean_ctor_get(x_258, 1); -lean_inc(x_259); -x_260 = lean_ctor_get(x_259, 1); -lean_inc(x_260); -x_261 = lean_ctor_get(x_256, 0); -lean_inc(x_261); -lean_dec_ref(x_256); -x_262 = lean_ctor_get(x_3, 0); -lean_inc(x_262); -lean_dec_ref(x_3); -x_263 = lean_ctor_get(x_257, 0); -lean_inc(x_263); -lean_dec(x_257); -x_264 = lean_ctor_get(x_258, 0); -lean_inc(x_264); -lean_dec(x_258); -x_265 = lean_ctor_get(x_259, 0); -lean_inc(x_265); -lean_dec(x_259); -x_266 = lean_ctor_get(x_260, 0); -lean_inc(x_266); -x_267 = lean_ctor_get(x_260, 1); -lean_inc(x_267); -if (lean_is_exclusive(x_260)) { - lean_ctor_release(x_260, 0); - lean_ctor_release(x_260, 1); - x_268 = x_260; -} else { - lean_dec_ref(x_260); - x_268 = lean_box(0); -} -x_269 = lean_unsigned_to_nat(1u); -x_270 = lean_nat_add(x_244, x_269); -lean_dec(x_244); -x_271 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_271, 0, x_243); -lean_ctor_set(x_271, 1, x_270); -lean_ctor_set(x_271, 2, x_245); -switch (lean_obj_tag(x_261)) { -case 0: -{ -lean_object* x_283; lean_object* x_284; uint8_t x_285; uint8_t x_286; uint8_t x_287; lean_object* x_288; -lean_dec(x_268); -lean_dec(x_266); -x_283 = lean_unsigned_to_nat(8u); -lean_inc(x_1); -x_284 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_284, 0, x_283); -lean_ctor_set(x_284, 1, x_1); -lean_ctor_set(x_284, 2, x_261); -x_285 = lean_unbox(x_263); -lean_dec(x_263); -x_286 = lean_unbox(x_264); -lean_dec(x_264); -x_287 = lean_unbox(x_265); -lean_dec(x_265); -x_288 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_262, x_285, x_286, x_287, x_246, x_267, x_284, x_4, x_5, x_6, x_7); -x_272 = x_288; -goto block_282; +lean_dec(x_14); +lean_dec(x_1); +x_5 = x_2; +x_6 = x_3; +x_7 = lean_box(0); +goto block_10; } -case 1: +block_10: { -lean_object* x_289; uint8_t x_290; uint8_t x_291; uint8_t x_292; lean_object* x_293; -lean_dec(x_268); -lean_dec(x_263); -lean_inc(x_1); -x_289 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_289, 0, x_269); -lean_ctor_set(x_289, 1, x_1); -lean_ctor_set(x_289, 2, x_261); -x_290 = lean_unbox(x_264); -lean_dec(x_264); -x_291 = lean_unbox(x_265); -lean_dec(x_265); -x_292 = lean_unbox(x_266); -lean_dec(x_266); -x_293 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_262, x_246, x_290, x_291, x_292, x_267, x_289, x_4, x_5, x_6, x_7); -x_272 = x_293; -goto block_282; +lean_object* x_8; lean_object* x_9; +x_8 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__1; +x_9 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__1(x_8, x_5, x_6); +return x_9; } -case 2: -{ -lean_object* x_294; lean_object* x_295; uint8_t x_296; uint8_t x_297; uint8_t x_298; lean_object* x_299; -lean_dec(x_268); -lean_dec(x_264); -x_294 = lean_unsigned_to_nat(2u); -lean_inc(x_1); -x_295 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_295, 0, x_294); -lean_ctor_set(x_295, 1, x_1); -lean_ctor_set(x_295, 2, x_261); -x_296 = lean_unbox(x_263); -lean_dec(x_263); -x_297 = lean_unbox(x_265); -lean_dec(x_265); -x_298 = lean_unbox(x_266); -lean_dec(x_266); -x_299 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_262, x_296, x_246, x_297, x_298, x_267, x_295, x_4, x_5, x_6, x_7); -x_272 = x_299; -goto block_282; } -case 3: -{ -lean_object* x_300; lean_object* x_301; uint8_t x_302; uint8_t x_303; uint8_t x_304; lean_object* x_305; -lean_dec(x_268); -lean_dec(x_265); -x_300 = lean_unsigned_to_nat(4u); -lean_inc(x_1); -x_301 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_301, 0, x_300); -lean_ctor_set(x_301, 1, x_1); -lean_ctor_set(x_301, 2, x_261); -x_302 = lean_unbox(x_263); -lean_dec(x_263); -x_303 = lean_unbox(x_264); -lean_dec(x_264); -x_304 = lean_unbox(x_266); -lean_dec(x_266); -x_305 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_262, x_302, x_303, x_246, x_304, x_267, x_301, x_4, x_5, x_6, x_7); -x_272 = x_305; -goto block_282; } -case 4: +LEAN_EXPORT lean_object* l_Lean_IR_getCtorLayout(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_306; lean_object* x_307; uint8_t x_308; uint8_t x_309; uint8_t x_310; lean_object* x_311; -lean_dec(x_268); -lean_dec(x_266); -x_306 = lean_unsigned_to_nat(8u); -lean_inc(x_1); -x_307 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_307, 0, x_306); -lean_ctor_set(x_307, 1, x_1); -lean_ctor_set(x_307, 2, x_261); -x_308 = lean_unbox(x_263); -lean_dec(x_263); -x_309 = lean_unbox(x_264); -lean_dec(x_264); -x_310 = lean_unbox(x_265); -lean_dec(x_265); -x_311 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_262, x_308, x_309, x_310, x_246, x_267, x_307, x_4, x_5, x_6, x_7); -x_272 = x_311; -goto block_282; -} -case 5: +lean_object* x_5; lean_object* x_6; +x_5 = l_Lean_IR_getCtorLayout___closed__0; +x_6 = l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg(x_5, x_1, x_3); +if (lean_obj_tag(x_6) == 0) { -lean_object* x_312; uint8_t x_313; uint8_t x_314; uint8_t x_315; uint8_t x_316; lean_object* x_317; -lean_dec(x_268); -lean_inc(x_1); -x_312 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_312, 0, x_1); -x_313 = lean_unbox(x_263); -lean_dec(x_263); -x_314 = lean_unbox(x_264); -lean_dec(x_264); -x_315 = lean_unbox(x_265); -lean_dec(x_265); -x_316 = lean_unbox(x_266); -lean_dec(x_266); -x_317 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_262, x_313, x_314, x_315, x_316, x_267, x_312, x_4, x_5, x_6, x_7); -x_272 = x_317; -goto block_282; -} -case 6: +uint8_t x_7; +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) { -lean_object* x_318; uint8_t x_319; uint8_t x_320; uint8_t x_321; uint8_t x_322; lean_object* x_323; -lean_dec(x_268); -x_318 = lean_box(0); -x_319 = lean_unbox(x_263); -lean_dec(x_263); -x_320 = lean_unbox(x_264); -lean_dec(x_264); -x_321 = lean_unbox(x_265); -lean_dec(x_265); -x_322 = lean_unbox(x_266); -lean_dec(x_266); -x_323 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_262, x_319, x_320, x_321, x_322, x_267, x_318, x_4, x_5, x_6, x_7); -x_272 = x_323; -goto block_282; -} -case 9: +lean_object* x_8; +x_8 = lean_ctor_get(x_6, 0); +if (lean_obj_tag(x_8) == 0) { -lean_object* x_324; lean_object* x_325; uint8_t x_326; uint8_t x_327; uint8_t x_328; lean_object* x_329; -lean_dec(x_268); -lean_dec(x_265); -x_324 = lean_unsigned_to_nat(4u); +lean_object* x_9; +lean_free_object(x_6); +lean_inc(x_3); lean_inc(x_1); -x_325 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_325, 0, x_324); -lean_ctor_set(x_325, 1, x_1); -lean_ctor_set(x_325, 2, x_261); -x_326 = lean_unbox(x_263); -lean_dec(x_263); -x_327 = lean_unbox(x_264); -lean_dec(x_264); -x_328 = lean_unbox(x_266); -lean_dec(x_266); -x_329 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_262, x_326, x_327, x_246, x_328, x_267, x_325, x_4, x_5, x_6, x_7); -x_272 = x_329; -goto block_282; -} -case 10: -{ -lean_object* x_330; lean_object* x_331; -lean_dec(x_268); -lean_dec_ref(x_261); -x_330 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___closed__0; -lean_inc(x_7); -lean_inc_ref(x_6); -lean_inc(x_5); -lean_inc_ref(x_4); -x_331 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2(x_330, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_331) == 0) -{ -lean_object* x_332; uint8_t x_333; uint8_t x_334; uint8_t x_335; uint8_t x_336; lean_object* x_337; -x_332 = lean_ctor_get(x_331, 0); -lean_inc(x_332); -lean_dec_ref(x_331); -x_333 = lean_unbox(x_263); -lean_dec(x_263); -x_334 = lean_unbox(x_264); -lean_dec(x_264); -x_335 = lean_unbox(x_265); -lean_dec(x_265); -x_336 = lean_unbox(x_266); -lean_dec(x_266); -x_337 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_262, x_333, x_334, x_335, x_336, x_267, x_332, x_4, x_5, x_6, x_7); -x_272 = x_337; -goto block_282; -} -else -{ -lean_object* x_338; lean_object* x_339; lean_object* x_340; -lean_dec_ref(x_271); -lean_dec(x_267); -lean_dec(x_266); -lean_dec(x_265); -lean_dec(x_264); -lean_dec(x_263); -lean_dec(x_262); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_1); -x_338 = lean_ctor_get(x_331, 0); -lean_inc(x_338); -if (lean_is_exclusive(x_331)) { - lean_ctor_release(x_331, 0); - x_339 = x_331; -} else { - lean_dec_ref(x_331); - x_339 = lean_box(0); -} -if (lean_is_scalar(x_339)) { - x_340 = lean_alloc_ctor(1, 1, 0); -} else { - x_340 = x_339; -} -lean_ctor_set(x_340, 0, x_338); -return x_340; -} -} -case 11: +x_9 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache(x_1, x_2, x_3); +if (lean_obj_tag(x_9) == 0) { -lean_object* x_341; lean_object* x_342; -lean_dec(x_268); -lean_dec_ref(x_261); -x_341 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___closed__0; -lean_inc(x_7); -lean_inc_ref(x_6); -lean_inc(x_5); -lean_inc_ref(x_4); -x_342 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2(x_341, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_342) == 0) -{ -lean_object* x_343; uint8_t x_344; uint8_t x_345; uint8_t x_346; uint8_t x_347; lean_object* x_348; -x_343 = lean_ctor_get(x_342, 0); -lean_inc(x_343); -lean_dec_ref(x_342); -x_344 = lean_unbox(x_263); -lean_dec(x_263); -x_345 = lean_unbox(x_264); -lean_dec(x_264); -x_346 = lean_unbox(x_265); -lean_dec(x_265); -x_347 = lean_unbox(x_266); -lean_dec(x_266); -x_348 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_262, x_344, x_345, x_346, x_347, x_267, x_343, x_4, x_5, x_6, x_7); -x_272 = x_348; -goto block_282; -} -else -{ -lean_object* x_349; lean_object* x_350; lean_object* x_351; -lean_dec_ref(x_271); -lean_dec(x_267); -lean_dec(x_266); -lean_dec(x_265); -lean_dec(x_264); -lean_dec(x_263); -lean_dec(x_262); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_1); -x_349 = lean_ctor_get(x_342, 0); -lean_inc(x_349); -if (lean_is_exclusive(x_342)) { - lean_ctor_release(x_342, 0); - x_350 = x_342; -} else { - lean_dec_ref(x_342); - x_350 = lean_box(0); -} -if (lean_is_scalar(x_350)) { - x_351 = lean_alloc_ctor(1, 1, 0); -} else { - x_351 = x_350; -} -lean_ctor_set(x_351, 0, x_349); -return x_351; -} -} -case 13: +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +lean_dec_ref(x_9); +lean_inc(x_10); +x_11 = l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg(x_5, x_1, x_10, x_3); +lean_dec(x_3); +if (lean_obj_tag(x_11) == 0) { -lean_object* x_352; uint8_t x_353; uint8_t x_354; uint8_t x_355; uint8_t x_356; lean_object* x_357; -lean_dec(x_268); -x_352 = lean_box(4); -x_353 = lean_unbox(x_263); -lean_dec(x_263); -x_354 = lean_unbox(x_264); -lean_dec(x_264); -x_355 = lean_unbox(x_265); -lean_dec(x_265); -x_356 = lean_unbox(x_266); -lean_dec(x_266); -x_357 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_262, x_353, x_354, x_355, x_356, x_267, x_352, x_4, x_5, x_6, x_7); -x_272 = x_357; -goto block_282; -} -default: +uint8_t x_12; +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) { -lean_object* x_358; lean_object* x_359; uint8_t x_360; uint8_t x_361; uint8_t x_362; uint8_t x_363; lean_object* x_364; -x_358 = lean_nat_add(x_267, x_269); -if (lean_is_scalar(x_268)) { - x_359 = lean_alloc_ctor(1, 2, 0); -} else { - x_359 = x_268; - lean_ctor_set_tag(x_359, 1); -} -lean_ctor_set(x_359, 0, x_267); -lean_ctor_set(x_359, 1, x_261); -x_360 = lean_unbox(x_263); -lean_dec(x_263); -x_361 = lean_unbox(x_264); -lean_dec(x_264); -x_362 = lean_unbox(x_265); -lean_dec(x_265); -x_363 = lean_unbox(x_266); -lean_dec(x_266); -x_364 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_262, x_360, x_361, x_362, x_363, x_358, x_359, x_4, x_5, x_6, x_7); -x_272 = x_364; -goto block_282; -} -} -block_282: -{ -if (lean_obj_tag(x_272) == 0) -{ -lean_object* x_273; lean_object* x_274; -x_273 = lean_ctor_get(x_272, 0); -lean_inc(x_273); -if (lean_is_exclusive(x_272)) { - lean_ctor_release(x_272, 0); - x_274 = x_272; -} else { - lean_dec_ref(x_272); - x_274 = lean_box(0); +lean_object* x_13; +x_13 = lean_ctor_get(x_11, 0); +lean_dec(x_13); +lean_ctor_set(x_11, 0, x_10); +return x_11; } -if (lean_obj_tag(x_273) == 0) +else { -lean_object* x_275; lean_object* x_276; -lean_dec_ref(x_271); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_1); -x_275 = lean_ctor_get(x_273, 0); -lean_inc(x_275); -lean_dec_ref(x_273); -if (lean_is_scalar(x_274)) { - x_276 = lean_alloc_ctor(0, 1, 0); -} else { - x_276 = x_274; +lean_object* x_14; +lean_dec(x_11); +x_14 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_14, 0, x_10); +return x_14; } -lean_ctor_set(x_276, 0, x_275); -return x_276; } else { -lean_object* x_277; -lean_dec(x_274); -x_277 = lean_ctor_get(x_273, 0); -lean_inc(x_277); -lean_dec_ref(x_273); -x_2 = x_271; -x_3 = x_277; -goto _start; -} +uint8_t x_15; +lean_dec(x_10); +x_15 = !lean_is_exclusive(x_11); +if (x_15 == 0) +{ +return x_11; } else { -lean_object* x_279; lean_object* x_280; lean_object* x_281; -lean_dec_ref(x_271); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_1); -x_279 = lean_ctor_get(x_272, 0); -lean_inc(x_279); -if (lean_is_exclusive(x_272)) { - lean_ctor_release(x_272, 0); - x_280 = x_272; -} else { - lean_dec_ref(x_272); - x_280 = lean_box(0); -} -if (lean_is_scalar(x_280)) { - x_281 = lean_alloc_ctor(1, 1, 0); -} else { - x_281 = x_280; -} -lean_ctor_set(x_281, 0, x_279); -return x_281; +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_11, 0); +lean_inc(x_16); +lean_dec(x_11); +x_17 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_17, 0, x_16); +return x_17; } } } else { -lean_object* x_365; lean_object* x_366; lean_object* x_367; -lean_dec(x_245); -lean_dec(x_244); -lean_dec_ref(x_243); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec_ref(x_3); +lean_dec(x_3); lean_dec(x_1); -x_365 = lean_ctor_get(x_256, 0); -lean_inc(x_365); -if (lean_is_exclusive(x_256)) { - lean_ctor_release(x_256, 0); - x_366 = x_256; -} else { - lean_dec_ref(x_256); - x_366 = lean_box(0); -} -if (lean_is_scalar(x_366)) { - x_367 = lean_alloc_ctor(1, 1, 0); -} else { - x_367 = x_366; -} -lean_ctor_set(x_367, 0, x_365); -return x_367; +return x_9; } } else { -lean_object* x_368; lean_object* x_369; lean_object* x_370; -lean_dec(x_245); -lean_dec(x_244); -lean_dec_ref(x_243); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec_ref(x_3); +lean_object* x_18; +lean_dec(x_3); +lean_dec_ref(x_2); lean_dec(x_1); -x_368 = lean_ctor_get(x_254, 0); -lean_inc(x_368); -if (lean_is_exclusive(x_254)) { - lean_ctor_release(x_254, 0); - x_369 = x_254; -} else { - lean_dec_ref(x_254); - x_369 = lean_box(0); +x_18 = lean_ctor_get(x_8, 0); +lean_inc(x_18); +lean_dec_ref(x_8); +lean_ctor_set(x_6, 0, x_18); +return x_6; } -if (lean_is_scalar(x_369)) { - x_370 = lean_alloc_ctor(1, 1, 0); +} +else +{ +lean_object* x_19; +x_19 = lean_ctor_get(x_6, 0); +lean_inc(x_19); +lean_dec(x_6); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; +lean_inc(x_3); +lean_inc(x_1); +x_20 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache(x_1, x_2, x_3); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +lean_dec_ref(x_20); +lean_inc(x_21); +x_22 = l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg(x_5, x_1, x_21, x_3); +lean_dec(x_3); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; lean_object* x_24; +if (lean_is_exclusive(x_22)) { + lean_ctor_release(x_22, 0); + x_23 = x_22; } else { - x_370 = x_369; + lean_dec_ref(x_22); + x_23 = lean_box(0); } -lean_ctor_set(x_370, 0, x_368); -return x_370; +if (lean_is_scalar(x_23)) { + x_24 = lean_alloc_ctor(0, 1, 0); +} else { + x_24 = x_23; } +lean_ctor_set(x_24, 0, x_21); +return x_24; } else { -lean_object* x_371; lean_object* x_372; lean_object* x_373; -lean_dec(x_245); -lean_dec(x_244); -lean_dec_ref(x_243); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -lean_dec(x_1); -x_371 = lean_ctor_get(x_252, 0); -lean_inc(x_371); -if (lean_is_exclusive(x_252)) { - lean_ctor_release(x_252, 0); - x_372 = x_252; +lean_object* x_25; lean_object* x_26; lean_object* x_27; +lean_dec(x_21); +x_25 = lean_ctor_get(x_22, 0); +lean_inc(x_25); +if (lean_is_exclusive(x_22)) { + lean_ctor_release(x_22, 0); + x_26 = x_22; } else { - lean_dec_ref(x_252); - x_372 = lean_box(0); + lean_dec_ref(x_22); + x_26 = lean_box(0); } -if (lean_is_scalar(x_372)) { - x_373 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_26)) { + x_27 = lean_alloc_ctor(1, 1, 0); } else { - x_373 = x_372; + x_27 = x_26; } -lean_ctor_set(x_373, 0, x_371); -return x_373; +lean_ctor_set(x_27, 0, x_25); +return x_27; } } else { -lean_object* x_374; lean_object* x_375; lean_object* x_376; -lean_dec(x_245); -lean_dec(x_244); -lean_dec_ref(x_243); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec_ref(x_3); +lean_dec(x_3); lean_dec(x_1); -x_374 = lean_ctor_get(x_250, 0); -lean_inc(x_374); -if (lean_is_exclusive(x_250)) { - lean_ctor_release(x_250, 0); - x_375 = x_250; -} else { - lean_dec_ref(x_250); - x_375 = lean_box(0); -} -if (lean_is_scalar(x_375)) { - x_376 = lean_alloc_ctor(1, 1, 0); -} else { - x_376 = x_375; +return x_20; } -lean_ctor_set(x_376, 0, x_374); -return x_376; } +else +{ +lean_object* x_28; lean_object* x_29; +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec(x_1); +x_28 = lean_ctor_get(x_19, 0); +lean_inc(x_28); +lean_dec_ref(x_19); +x_29 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_29, 0, x_28); +return x_29; } } } +else +{ +uint8_t x_30; +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec(x_1); +x_30 = !lean_is_exclusive(x_6); +if (x_30 == 0) +{ +return x_6; } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: +else { -lean_object* x_12; -x_12 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg(x_1, x_4, x_5, x_7, x_8, x_9, x_10); -return x_12; +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_6, 0); +lean_inc(x_31); +lean_dec(x_6); +x_32 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_32, 0, x_31); +return x_32; } } -LEAN_EXPORT lean_object* l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__0___closed__0; -x_6 = lean_panic_fn(x_5, x_1); -x_7 = lean_apply_3(x_6, x_2, x_3, lean_box(0)); -return x_7; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__0(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -uint8_t x_6; -x_6 = lean_usize_dec_lt(x_3, x_2); -if (x_6 == 0) +lean_object* x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; +x_5 = lean_st_ref_get(x_3); +x_6 = lean_ctor_get(x_5, 0); +lean_inc_ref(x_6); +lean_dec(x_5); +x_7 = 0; +lean_inc(x_1); +x_8 = l_Lean_Environment_findConstVal_x3f(x_6, x_1, x_7); +if (lean_obj_tag(x_8) == 0) { -lean_object* x_7; -x_7 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_7, 0, x_4); -lean_ctor_set(x_7, 1, x_5); -return x_7; +lean_object* x_9; +x_9 = l_Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15___redArg(x_1, x_2, x_3); +return x_9; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_8 = lean_array_uget(x_4, x_3); -x_9 = lean_unsigned_to_nat(0u); -x_10 = lean_array_uset(x_4, x_3, x_9); -switch (lean_obj_tag(x_8)) { -case 1: +uint8_t x_10; +lean_dec_ref(x_2); +lean_dec(x_1); +x_10 = !lean_is_exclusive(x_8); +if (x_10 == 0) { -x_11 = x_8; -x_12 = x_5; -goto block_17; +lean_ctor_set_tag(x_8, 0); +return x_8; } -case 2: +else { -x_11 = x_8; -x_12 = x_5; -goto block_17; +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_8, 0); +lean_inc(x_11); +lean_dec(x_8); +x_12 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_12, 0, x_11); +return x_12; } -case 3: +} +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__13___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_18; lean_object* x_19; uint8_t x_20; -x_18 = lean_ctor_get(x_8, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_8, 2); -lean_inc(x_19); -x_20 = lean_nat_dec_eq(x_18, x_1); -if (x_20 == 0) +lean_object* x_15; lean_object* x_16; uint8_t x_21; +x_21 = lean_nat_dec_lt(x_8, x_1); +if (x_21 == 0) { -lean_dec(x_19); -lean_dec(x_18); -x_11 = x_8; -x_12 = x_5; -goto block_17; +lean_object* x_22; +lean_dec(x_13); +lean_dec_ref(x_12); +lean_dec(x_11); +lean_dec_ref(x_10); +lean_dec(x_8); +lean_dec_ref(x_4); +x_22 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_22, 0, x_9); +return x_22; } else { -uint8_t x_21; -x_21 = !lean_is_exclusive(x_8); -if (x_21 == 0) +lean_object* x_23; +x_23 = lean_array_fget_borrowed(x_2, x_8); +if (lean_obj_tag(x_23) == 1) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_22 = lean_ctor_get(x_8, 2); -lean_dec(x_22); -x_23 = lean_ctor_get(x_8, 1); -lean_dec(x_23); -x_24 = lean_ctor_get(x_8, 0); -lean_dec(x_24); -x_25 = lean_nat_add(x_5, x_18); -lean_ctor_set(x_8, 1, x_5); -x_11 = x_8; -x_12 = x_25; -goto block_17; +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_24 = lean_ctor_get(x_23, 0); +x_25 = lean_nat_add(x_8, x_3); +lean_inc_ref(x_4); +x_26 = lean_array_get_borrowed(x_4, x_5, x_25); +lean_dec(x_25); +x_27 = l_Lean_Expr_fvarId_x21(x_26); +lean_inc_ref(x_10); +x_28 = l_Lean_FVarId_getType___redArg(x_27, x_10, x_12, x_13); +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +lean_dec_ref(x_28); +lean_inc(x_13); +lean_inc_ref(x_12); +lean_inc(x_11); +lean_inc_ref(x_10); +x_30 = l_Lean_Compiler_LCNF_toLCNFType(x_29, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_30) == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +lean_dec_ref(x_30); +x_32 = l_Lean_Expr_replaceFVars(x_31, x_6, x_7); +lean_dec(x_31); +lean_inc(x_13); +lean_inc_ref(x_12); +x_33 = l_Lean_Compiler_LCNF_toMonoType(x_32, x_12, x_13); +if (lean_obj_tag(x_33) == 0) +{ +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +lean_dec_ref(x_33); +lean_inc(x_13); +lean_inc_ref(x_12); +x_35 = l_Lean_IR_toIRType(x_34, x_12, x_13); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; lean_object* x_37; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +lean_dec_ref(x_35); +x_37 = lean_array_set(x_9, x_24, x_36); +x_15 = x_37; +x_16 = lean_box(0); +goto block_20; } else { -lean_object* x_26; lean_object* x_27; +uint8_t x_38; +lean_dec(x_13); +lean_dec_ref(x_12); +lean_dec(x_11); +lean_dec_ref(x_10); +lean_dec_ref(x_9); lean_dec(x_8); -x_26 = lean_nat_add(x_5, x_18); -x_27 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_27, 0, x_18); -lean_ctor_set(x_27, 1, x_5); -lean_ctor_set(x_27, 2, x_19); -x_11 = x_27; -x_12 = x_26; -goto block_17; -} -} -} -default: +lean_dec_ref(x_4); +x_38 = !lean_is_exclusive(x_35); +if (x_38 == 0) { -x_11 = x_8; -x_12 = x_5; -goto block_17; -} +return x_35; } -block_17: +else { -size_t x_13; size_t x_14; lean_object* x_15; -x_13 = 1; -x_14 = lean_usize_add(x_3, x_13); -x_15 = lean_array_uset(x_10, x_3, x_11); -x_3 = x_14; -x_4 = x_15; -x_5 = x_12; -goto _start; +lean_object* x_39; lean_object* x_40; +x_39 = lean_ctor_get(x_35, 0); +lean_inc(x_39); +lean_dec(x_35); +x_40 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_40, 0, x_39); +return x_40; } } } +else +{ +uint8_t x_41; +lean_dec(x_13); +lean_dec_ref(x_12); +lean_dec(x_11); +lean_dec_ref(x_10); +lean_dec_ref(x_9); +lean_dec(x_8); +lean_dec_ref(x_4); +x_41 = !lean_is_exclusive(x_33); +if (x_41 == 0) +{ +return x_33; } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -size_t x_4; size_t x_5; lean_object* x_6; -x_4 = lean_array_size(x_1); -x_5 = 0; -x_6 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__0(x_2, x_4, x_5, x_1, x_3); -return x_6; +lean_object* x_42; lean_object* x_43; +x_42 = lean_ctor_get(x_33, 0); +lean_inc(x_42); +lean_dec(x_33); +x_43 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_43, 0, x_42); +return x_43; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__4(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4) { -_start: +} +else { -uint8_t x_5; -x_5 = lean_usize_dec_lt(x_2, x_1); -if (x_5 == 0) +uint8_t x_44; +lean_dec(x_13); +lean_dec_ref(x_12); +lean_dec(x_11); +lean_dec_ref(x_10); +lean_dec_ref(x_9); +lean_dec(x_8); +lean_dec_ref(x_4); +x_44 = !lean_is_exclusive(x_30); +if (x_44 == 0) { -lean_object* x_6; -x_6 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_6, 0, x_3); -lean_ctor_set(x_6, 1, x_4); -return x_6; +return x_30; } else { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_7 = lean_array_uget(x_3, x_2); -x_8 = lean_unsigned_to_nat(0u); -x_9 = lean_array_uset(x_3, x_2, x_8); -switch (lean_obj_tag(x_7)) { -case 1: -{ -x_10 = x_7; -x_11 = x_4; -goto block_16; +lean_object* x_45; lean_object* x_46; +x_45 = lean_ctor_get(x_30, 0); +lean_inc(x_45); +lean_dec(x_30); +x_46 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_46, 0, x_45); +return x_46; } -case 2: +} +} +else { -uint8_t x_17; -x_17 = !lean_is_exclusive(x_7); -if (x_17 == 0) +uint8_t x_47; +lean_dec(x_13); +lean_dec_ref(x_12); +lean_dec(x_11); +lean_dec_ref(x_10); +lean_dec_ref(x_9); +lean_dec(x_8); +lean_dec_ref(x_4); +x_47 = !lean_is_exclusive(x_28); +if (x_47 == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_7, 0); -lean_dec(x_18); -x_19 = lean_unsigned_to_nat(1u); -x_20 = lean_nat_add(x_4, x_19); -lean_ctor_set(x_7, 0, x_4); -x_10 = x_7; -x_11 = x_20; -goto block_16; +return x_28; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -lean_dec(x_7); -x_21 = lean_unsigned_to_nat(1u); -x_22 = lean_nat_add(x_4, x_21); -x_23 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_23, 0, x_4); -x_10 = x_23; -x_11 = x_22; -goto block_16; +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_28, 0); +lean_inc(x_48); +lean_dec(x_28); +x_49 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_49, 0, x_48); +return x_49; } } -case 3: -{ -x_10 = x_7; -x_11 = x_4; -goto block_16; } -default: +else { -x_10 = x_7; -x_11 = x_4; -goto block_16; +x_15 = x_9; +x_16 = lean_box(0); +goto block_20; } } -block_16: +block_20: { -size_t x_12; size_t x_13; lean_object* x_14; -x_12 = 1; -x_13 = lean_usize_add(x_2, x_12); -x_14 = lean_array_uset(x_9, x_2, x_10); -x_2 = x_13; -x_3 = x_14; -x_4 = x_11; +lean_object* x_17; lean_object* x_18; +x_17 = lean_unsigned_to_nat(1u); +x_18 = lean_nat_add(x_8, x_17); +lean_dec(x_8); +x_8 = x_18; +x_9 = x_15; goto _start; } } } -} -LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___lam__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_14 = lean_mk_empty_array_with_capacity(x_1); -x_15 = lean_unsigned_to_nat(0u); -x_16 = lean_nat_add(x_2, x_1); -x_17 = l_Array_toSubarray___redArg(x_7, x_2, x_16); -x_18 = lean_box(x_3); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_15); -x_20 = lean_box(x_3); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_19); -x_22 = lean_box(x_3); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_21); -x_24 = lean_box(x_3); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_23); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_14); +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_12 = lean_ctor_get(x_1, 0); +lean_inc_ref(x_12); +x_13 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_13); +lean_dec_ref(x_1); +x_14 = lean_ctor_get(x_12, 2); +lean_inc(x_14); +x_15 = lean_ctor_get(x_12, 3); +lean_inc(x_15); +x_16 = lean_ctor_get(x_12, 4); +lean_inc(x_16); +lean_dec_ref(x_12); +x_17 = lean_array_get_size(x_13); +x_18 = lean_unsigned_to_nat(0u); +x_19 = lean_box(0); +x_20 = lean_mk_array(x_14, x_19); +x_21 = lean_array_get_size(x_2); +x_22 = l_Array_extract___redArg(x_5, x_18, x_21); +x_23 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__13___redArg(x_17, x_13, x_21, x_3, x_5, x_22, x_2, x_18, x_20, x_7, x_8, x_9, x_10); +lean_dec_ref(x_22); +lean_dec_ref(x_13); +if (lean_obj_tag(x_23) == 0) +{ +uint8_t x_24; +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_23, 0); +x_26 = lean_alloc_ctor(10, 4, 0); +lean_ctor_set(x_26, 0, x_4); lean_ctor_set(x_26, 1, x_25); -x_27 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg(x_15, x_17, x_26, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_27) == 0) +lean_ctor_set(x_26, 2, x_15); +lean_ctor_set(x_26, 3, x_16); +lean_ctor_set(x_23, 0, x_26); +return x_23; +} +else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; size_t x_40; size_t x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_81; -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -if (lean_is_exclusive(x_27)) { - lean_ctor_release(x_27, 0); - x_29 = x_27; -} else { - lean_dec_ref(x_27); - x_29 = lean_box(0); +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_23, 0); +lean_inc(x_27); +lean_dec(x_23); +x_28 = lean_alloc_ctor(10, 4, 0); +lean_ctor_set(x_28, 0, x_4); +lean_ctor_set(x_28, 1, x_27); +lean_ctor_set(x_28, 2, x_15); +lean_ctor_set(x_28, 3, x_16); +x_29 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_29, 0, x_28); +return x_29; } -x_30 = lean_ctor_get(x_28, 1); -lean_inc(x_30); -x_31 = lean_ctor_get(x_30, 1); -lean_inc(x_31); -x_32 = lean_ctor_get(x_31, 1); -lean_inc(x_32); -x_33 = lean_ctor_get(x_32, 1); -lean_inc(x_33); -x_34 = lean_ctor_get(x_28, 0); -lean_inc(x_34); -lean_dec(x_28); -x_35 = lean_ctor_get(x_30, 0); -lean_inc(x_35); -lean_dec(x_30); -x_36 = lean_ctor_get(x_31, 0); -lean_inc(x_36); -lean_dec(x_31); -x_37 = lean_ctor_get(x_32, 0); -lean_inc(x_37); -lean_dec(x_32); -x_38 = lean_ctor_get(x_33, 0); -lean_inc(x_38); -x_39 = lean_ctor_get(x_33, 1); -lean_inc(x_39); -lean_dec(x_33); -x_40 = lean_array_size(x_34); -x_41 = 0; -lean_inc(x_39); -x_42 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__4(x_40, x_41, x_34, x_39); -x_43 = lean_ctor_get(x_42, 0); -lean_inc(x_43); -x_44 = lean_ctor_get(x_42, 1); -lean_inc(x_44); -if (lean_is_exclusive(x_42)) { - lean_ctor_release(x_42, 0); - lean_ctor_release(x_42, 1); - x_45 = x_42; -} else { - lean_dec_ref(x_42); - x_45 = lean_box(0); } -x_46 = lean_nat_sub(x_44, x_39); -lean_dec(x_44); -x_81 = lean_unbox(x_38); -lean_dec(x_38); -if (x_81 == 0) +else { -x_72 = x_43; -x_73 = x_15; -x_74 = lean_box(0); -goto block_80; +uint8_t x_30; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_4); +x_30 = !lean_is_exclusive(x_23); +if (x_30 == 0) +{ +return x_23; } else { -lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_82 = lean_unsigned_to_nat(8u); -lean_inc_ref(x_6); -x_83 = lean_apply_3(x_6, x_43, x_82, x_15); -x_84 = lean_ctor_get(x_83, 0); -lean_inc(x_84); -x_85 = lean_ctor_get(x_83, 1); -lean_inc(x_85); -lean_dec_ref(x_83); -x_72 = x_84; -x_73 = x_85; -x_74 = lean_box(0); -goto block_80; +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_23, 0); +lean_inc(x_31); +lean_dec(x_23); +x_32 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_32, 0, x_31); +return x_32; } -block_53: -{ -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_50, 0, x_4); -lean_ctor_set(x_50, 1, x_5); -lean_ctor_set(x_50, 2, x_39); -lean_ctor_set(x_50, 3, x_46); -lean_ctor_set(x_50, 4, x_48); -if (lean_is_scalar(x_45)) { - x_51 = lean_alloc_ctor(0, 2, 0); -} else { - x_51 = x_45; } -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_47); -if (lean_is_scalar(x_29)) { - x_52 = lean_alloc_ctor(0, 1, 0); -} else { - x_52 = x_29; } -lean_ctor_set(x_52, 0, x_51); -return x_52; } -block_62: +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec_ref(x_6); +lean_dec_ref(x_5); +lean_dec_ref(x_2); +return x_12; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_7; +lean_inc(x_5); +lean_inc_ref(x_4); +lean_inc(x_2); +x_7 = l_Lean_IR_getCtorLayout(x_2, x_4, x_5); +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +lean_dec_ref(x_7); +lean_inc_ref(x_4); +x_9 = l_Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14(x_2, x_4, x_5); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +lean_dec_ref(x_9); +x_11 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___closed__1; +x_12 = lean_st_mk_ref(x_11); +x_13 = lean_ctor_get(x_10, 2); +lean_inc_ref(x_13); +lean_dec(x_10); +x_14 = l_Lean_instInhabitedExpr; +x_15 = lean_alloc_closure((void*)(l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___lam__0___boxed), 11, 4); +lean_closure_set(x_15, 0, x_8); +lean_closure_set(x_15, 1, x_3); +lean_closure_set(x_15, 2, x_14); +lean_closure_set(x_15, 3, x_1); +x_16 = 0; +x_17 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___closed__3; +lean_inc(x_12); +x_18 = l_Lean_Meta_forallTelescope___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__2___redArg(x_13, x_15, x_16, x_17, x_12, x_4, x_5); +if (lean_obj_tag(x_18) == 0) { -uint8_t x_57; -x_57 = lean_unbox(x_35); -lean_dec(x_35); -if (x_57 == 0) +uint8_t x_19; +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) { -lean_dec_ref(x_6); -x_47 = x_54; -x_48 = x_55; -x_49 = lean_box(0); -goto block_53; +lean_object* x_20; +x_20 = lean_st_ref_get(x_12); +lean_dec(x_12); +lean_dec(x_20); +return x_18; } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_58 = lean_unsigned_to_nat(1u); -x_59 = lean_apply_3(x_6, x_54, x_58, x_55); -x_60 = lean_ctor_get(x_59, 0); -lean_inc(x_60); -x_61 = lean_ctor_get(x_59, 1); -lean_inc(x_61); -lean_dec_ref(x_59); -x_47 = x_60; -x_48 = x_61; -x_49 = lean_box(0); -goto block_53; -} +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_18, 0); +lean_inc(x_21); +lean_dec(x_18); +x_22 = lean_st_ref_get(x_12); +lean_dec(x_12); +lean_dec(x_22); +x_23 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_23, 0, x_21); +return x_23; } -block_71: -{ -uint8_t x_66; -x_66 = lean_unbox(x_36); -lean_dec(x_36); -if (x_66 == 0) -{ -x_54 = x_63; -x_55 = x_64; -x_56 = lean_box(0); -goto block_62; } else { -lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_67 = lean_unsigned_to_nat(2u); -lean_inc_ref(x_6); -x_68 = lean_apply_3(x_6, x_63, x_67, x_64); -x_69 = lean_ctor_get(x_68, 0); -lean_inc(x_69); -x_70 = lean_ctor_get(x_68, 1); -lean_inc(x_70); -lean_dec_ref(x_68); -x_54 = x_69; -x_55 = x_70; -x_56 = lean_box(0); -goto block_62; +lean_dec(x_12); +return x_18; } } -block_80: +else { -uint8_t x_75; -x_75 = lean_unbox(x_37); -lean_dec(x_37); -if (x_75 == 0) +uint8_t x_24; +lean_dec(x_8); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec_ref(x_3); +lean_dec(x_1); +x_24 = !lean_is_exclusive(x_9); +if (x_24 == 0) { -x_63 = x_72; -x_64 = x_73; -x_65 = lean_box(0); -goto block_71; +return x_9; } else { -lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_76 = lean_unsigned_to_nat(4u); -lean_inc_ref(x_6); -x_77 = lean_apply_3(x_6, x_72, x_76, x_73); -x_78 = lean_ctor_get(x_77, 0); -lean_inc(x_78); -x_79 = lean_ctor_get(x_77, 1); -lean_inc(x_79); -lean_dec_ref(x_77); -x_63 = x_78; -x_64 = x_79; -x_65 = lean_box(0); -goto block_71; +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_9, 0); +lean_inc(x_25); +lean_dec(x_9); +x_26 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_26, 0, x_25); +return x_26; } } } else { -uint8_t x_86; -lean_dec_ref(x_6); +uint8_t x_27; lean_dec(x_5); -lean_dec(x_4); -x_86 = !lean_is_exclusive(x_27); -if (x_86 == 0) +lean_dec_ref(x_4); +lean_dec_ref(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_27 = !lean_is_exclusive(x_7); +if (x_27 == 0) { -return x_27; +return x_7; } else { -lean_object* x_87; lean_object* x_88; -x_87 = lean_ctor_get(x_27, 0); -lean_inc(x_87); -lean_dec(x_27); -x_88 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_88, 0, x_87); -return x_88; +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_7, 0); +lean_inc(x_28); +lean_dec(x_7); +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, x_28); +return x_29; } } } } -static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__1() { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00Lean_IR_unboxedIRType_spec__10___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__2; -x_2 = lean_unsigned_to_nat(64u); -x_3 = lean_unsigned_to_nat(169u); -x_4 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__0; -x_5 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__0; -x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); -return x_6; -} +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_7; +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec_ref(x_1); +x_7 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_7, 0, x_3); +return x_7; } -static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__5; -x_2 = lean_unsigned_to_nat(0u); -x_3 = lean_alloc_ctor(0, 9, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_2); -lean_ctor_set(x_3, 2, x_2); -lean_ctor_set(x_3, 3, x_1); -lean_ctor_set(x_3, 4, x_1); -lean_ctor_set(x_3, 5, x_1); -lean_ctor_set(x_3, 6, x_1); -lean_ctor_set(x_3, 7, x_1); -lean_ctor_set(x_3, 8, x_1); -return x_3; +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_2, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_2, 1); +lean_inc(x_9); +lean_dec_ref(x_2); +x_10 = lean_box(0); +lean_inc(x_5); +lean_inc_ref(x_4); +lean_inc_ref(x_1); +x_11 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor(x_10, x_8, x_1, x_4, x_5); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +lean_dec_ref(x_11); +x_13 = lean_array_push(x_3, x_12); +x_2 = x_9; +x_3 = x_13; +goto _start; } +else +{ +uint8_t x_15; +lean_dec(x_9); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_1); +x_15 = !lean_is_exclusive(x_11); +if (x_15 == 0) +{ +return x_11; } -static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__3() { -_start: +else { -size_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = 5; -x_2 = lean_unsigned_to_nat(0u); -x_3 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__7; -x_4 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__8; -x_5 = lean_alloc_ctor(0, 4, sizeof(size_t)*1); -lean_ctor_set(x_5, 0, x_4); -lean_ctor_set(x_5, 1, x_3); -lean_ctor_set(x_5, 2, x_2); -lean_ctor_set(x_5, 3, x_2); -lean_ctor_set_usize(x_5, 4, x_1); -return x_5; +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_11, 0); +lean_inc(x_16); +lean_dec(x_11); +x_17 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_17, 0, x_16); +return x_17; +} } } -static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__4() { +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00Lean_IR_unboxedIRType_spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__9; -x_2 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__3; -x_3 = lean_box(1); -x_4 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__6; -x_5 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__2; -x_6 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_6, 0, x_5); -lean_ctor_set(x_6, 1, x_4); -lean_ctor_set(x_6, 2, x_3); -lean_ctor_set(x_6, 3, x_2); -lean_ctor_set(x_6, 4, x_1); -return x_6; +lean_object* x_9; +x_9 = l_List_forIn_x27_loop___at___00Lean_IR_unboxedIRType_spec__10___redArg(x_1, x_3, x_4, x_6, x_7); +return x_9; } } -static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__5() { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = lean_box(1); -x_2 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__3; -x_3 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__5; -x_4 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_4, 0, x_3); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_1); -return x_4; +lean_object* x_12; +x_12 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg(x_1, x_4, x_5, x_7, x_8, x_9, x_10); +return x_12; } } -static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__6() { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__13(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { -uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; lean_object* x_9; -x_1 = 1; -x_2 = lean_unsigned_to_nat(0u); -x_3 = lean_box(0); -x_4 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__11; -x_5 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__5; -x_6 = lean_box(1); -x_7 = 0; -x_8 = l_List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3___redArg___closed__10; -x_9 = lean_alloc_ctor(0, 7, 4); -lean_ctor_set(x_9, 0, x_8); -lean_ctor_set(x_9, 1, x_6); -lean_ctor_set(x_9, 2, x_5); -lean_ctor_set(x_9, 3, x_4); -lean_ctor_set(x_9, 4, x_3); -lean_ctor_set(x_9, 5, x_2); -lean_ctor_set(x_9, 6, x_3); -lean_ctor_set_uint8(x_9, sizeof(void*)*7, x_7); -lean_ctor_set_uint8(x_9, sizeof(void*)*7 + 1, x_7); -lean_ctor_set_uint8(x_9, sizeof(void*)*7 + 2, x_7); -lean_ctor_set_uint8(x_9, sizeof(void*)*7 + 3, x_1); -return x_9; +lean_object* x_18; +x_18 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__13___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_10, x_11, x_13, x_14, x_15, x_16); +return x_18; } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_4; -x_4 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___lam__0(x_1, x_2, x_3); -lean_dec(x_2); -return x_4; +uint8_t x_11; uint8_t x_12; lean_object* x_13; +x_11 = lean_unbox(x_4); +x_12 = lean_unbox(x_5); +x_13 = l_Lean_Meta_forallTelescopeReducing___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__4(x_1, x_2, x_3, x_11, x_12, x_6, x_7, x_8, x_9); +return x_13; } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -uint8_t x_14; lean_object* x_15; -x_14 = lean_unbox(x_3); -x_15 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___lam__1(x_1, x_2, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec_ref(x_8); -lean_dec(x_1); -return x_15; +lean_object* x_6; +x_6 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20(x_1, x_2, x_3, x_4); +lean_dec(x_4); +lean_dec_ref(x_3); +return x_6; } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; -x_11 = lean_st_ref_get(x_3); -x_12 = lean_ctor_get(x_11, 0); -lean_inc_ref(x_12); -lean_dec(x_11); -x_13 = 0; -lean_inc(x_1); -x_14 = l_Lean_Environment_find_x3f(x_12, x_1, x_13); -if (lean_obj_tag(x_14) == 1) -{ -lean_object* x_15; -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -lean_dec_ref(x_14); -if (lean_obj_tag(x_15) == 6) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; 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_16 = lean_ctor_get(x_15, 0); -lean_inc_ref(x_16); -lean_dec_ref(x_15); -x_17 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__4; -x_18 = lean_st_mk_ref(x_17); -x_19 = lean_ctor_get(x_16, 0); -lean_inc_ref(x_19); -x_20 = lean_ctor_get(x_16, 2); -lean_inc(x_20); -x_21 = lean_ctor_get(x_16, 3); -lean_inc(x_21); -x_22 = lean_ctor_get(x_16, 4); -lean_inc(x_22); -lean_dec_ref(x_16); -x_23 = lean_ctor_get(x_19, 2); -lean_inc_ref(x_23); -lean_dec_ref(x_19); -x_24 = lean_alloc_closure((void*)(l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___lam__0___boxed), 3, 0); -x_25 = lean_box(x_13); -x_26 = lean_alloc_closure((void*)(l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___lam__1___boxed), 13, 6); -lean_closure_set(x_26, 0, x_22); -lean_closure_set(x_26, 1, x_21); -lean_closure_set(x_26, 2, x_25); -lean_closure_set(x_26, 3, x_1); -lean_closure_set(x_26, 4, x_20); -lean_closure_set(x_26, 5, x_24); -x_27 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__6; -lean_inc(x_18); -x_28 = l_Lean_Meta_forallTelescopeReducing___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__5___redArg(x_23, x_26, x_13, x_13, x_27, x_18, x_2, x_3); -if (lean_obj_tag(x_28) == 0) -{ -uint8_t x_29; -x_29 = !lean_is_exclusive(x_28); -if (x_29 == 0) -{ -lean_object* x_30; -x_30 = lean_st_ref_get(x_18); -lean_dec(x_18); -lean_dec(x_30); -return x_28; +lean_object* x_6; +x_6 = l_Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11(x_1, x_2, x_3, x_4); +lean_dec(x_4); +lean_dec_ref(x_3); +return x_6; } -else +} +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__20___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_28, 0); -lean_inc(x_31); -lean_dec(x_28); -x_32 = lean_st_ref_get(x_18); -lean_dec(x_18); -lean_dec(x_32); -x_33 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_33, 0, x_31); -return x_33; +lean_object* x_7; +x_7 = l_Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__20(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec(x_2); +return x_7; } } -else +LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -lean_dec(x_18); -return x_28; +lean_object* x_6; +x_6 = l_Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15(x_1, x_2, x_3, x_4); +lean_dec(x_4); +return x_6; } } -else +LEAN_EXPORT lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -lean_dec(x_15); -lean_dec(x_1); -x_5 = x_2; -x_6 = x_3; -x_7 = lean_box(0); -goto block_10; +lean_object* x_7; +x_7 = l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec(x_2); +return x_7; } } -else +LEAN_EXPORT lean_object* l_Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: { -lean_dec(x_14); -lean_dec(x_1); -x_5 = x_2; -x_6 = x_3; -x_7 = lean_box(0); -goto block_10; +lean_object* x_8; +x_8 = l_Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_6); +lean_dec(x_2); +return x_8; } -block_10: +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00Lean_IR_unboxedIRType_spec__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -lean_object* x_8; lean_object* x_9; -x_8 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__1; -x_9 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__1(x_8, x_5, x_6); +lean_object* x_9; +x_9 = l_List_forIn_x27_loop___at___00Lean_IR_unboxedIRType_spec__10(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_2); return x_9; } } +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_12; +} } -LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__13___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; _start: { -uint8_t x_11; uint8_t x_12; lean_object* x_13; -x_11 = lean_unbox(x_4); -x_12 = lean_unbox(x_5); -x_13 = l_Lean_Meta_forallTelescopeReducing___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__5(x_1, x_2, x_3, x_11, x_12, x_6, x_7, x_8, x_9); -return x_13; +lean_object* x_18; +x_18 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__13(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec_ref(x_7); +lean_dec_ref(x_6); +lean_dec_ref(x_5); +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec(x_1); +return x_18; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_toIRType_spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_12; -x_12 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_12; +lean_object* x_5; +x_5 = l_panic___at___00Lean_IR_toIRType_spec__6(x_1, x_2, x_3); +return x_5; } } LEAN_EXPORT lean_object* l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { @@ -7862,25 +9584,56 @@ x_5 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLay return x_5; } } -LEAN_EXPORT lean_object* l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_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_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15___redArg___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_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15___redArg(x_1, x_2, x_3); +lean_dec(x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__20___redArg___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_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__20___redArg(x_1, x_2, x_3, x_4); +lean_dec(x_4); +lean_dec(x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18___redArg___boxed(lean_object* x_1, 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_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2(x_1, x_2, x_3, x_4, x_5); +x_7 = l_Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18___redArg(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__5___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__4___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { uint8_t x_10; uint8_t x_11; lean_object* x_12; x_10 = lean_unbox(x_3); x_11 = lean_unbox(x_4); -x_12 = l_Lean_Meta_forallTelescopeReducing___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__5___redArg(x_1, x_2, x_10, x_11, x_5, x_6, x_7, x_8); +x_12 = l_Lean_Meta_forallTelescopeReducing___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__4___redArg(x_1, x_2, x_10, x_11, x_5, x_6, x_7, x_8); return x_12; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19___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_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19(x_1, x_2, x_3, x_4); +lean_dec(x_4); +lean_dec_ref(x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { uint8_t x_13; uint8_t x_14; uint8_t x_15; uint8_t x_16; lean_object* x_17; @@ -7888,7 +9641,7 @@ x_13 = lean_unbox(x_2); x_14 = lean_unbox(x_3); x_15 = lean_unbox(x_4); x_16 = lean_unbox(x_5); -x_17 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_1, x_13, x_14, x_15, x_16, x_6, x_7, x_8, x_9, x_10, x_11); +x_17 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___lam__0(x_1, x_13, x_14, x_15, x_16, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_11); lean_dec_ref(x_10); lean_dec(x_9); @@ -7896,189 +9649,171 @@ lean_dec_ref(x_8); return x_17; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -size_t x_5; size_t x_6; lean_object* x_7; -x_5 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_6 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_7 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__4(x_5, x_6, x_3, x_4); -return x_7; +lean_object* x_5; +x_5 = l_Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11___redArg(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec_ref(x_2); +return x_5; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__0___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_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_7 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_8 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__0(x_1, x_6, x_7, x_4, x_5); +lean_object* x_6; +x_6 = l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg(x_1, x_2, x_3, x_4); +lean_dec(x_4); lean_dec(x_1); -return x_8; +return x_6; } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache(x_1, x_2, x_3); +x_5 = l_Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14(x_1, x_2, x_3); +lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -return x_9; -} -} -static lean_object* _init_l_Lean_IR_getCtorLayout___closed__0() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_IR_ctorLayoutExt; -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_IR_getCtorLayout(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00Lean_IR_unboxedIRType_spec__10___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_5 = l_Lean_IR_getCtorLayout___closed__0; -x_6 = l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg(x_5, x_1, x_3); -x_7 = !lean_is_exclusive(x_6); -if (x_7 == 0) -{ -lean_object* x_8; -x_8 = lean_ctor_get(x_6, 0); -if (lean_obj_tag(x_8) == 0) -{ -lean_object* x_9; -lean_free_object(x_6); -lean_inc(x_3); -lean_inc(x_1); -x_9 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache(x_1, x_2, x_3); -if (lean_obj_tag(x_9) == 0) -{ -lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -lean_dec_ref(x_9); -lean_inc(x_10); -x_11 = l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg(x_5, x_1, x_10, x_3); +lean_object* x_7; +x_7 = l_List_forIn_x27_loop___at___00Lean_IR_unboxedIRType_spec__10___redArg(x_1, x_2, x_3, x_4, x_5); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_getConstInfoInduct___at___00Lean_IR_unboxedIRType_spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Lean_getConstInfoInduct___at___00Lean_IR_unboxedIRType_spec__9(x_1, x_2, x_3); lean_dec(x_3); -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) +lean_dec_ref(x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_getCtorLayout___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_13; -x_13 = lean_ctor_get(x_11, 0); -lean_dec(x_13); -lean_ctor_set(x_11, 0, x_10); -return x_11; +lean_object* x_5; +x_5 = l_Lean_IR_getCtorLayout(x_1, x_2, x_3); +return x_5; } -else +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_14; -lean_dec(x_11); -x_14 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_14, 0, x_10); -return x_14; +size_t x_5; size_t x_6; lean_object* x_7; +x_5 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_6 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_7 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3(x_5, x_6, x_3, x_4); +return x_7; } } -else +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_7 = lean_unbox_usize(x_3); lean_dec(x_3); +x_8 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__0(x_1, x_6, x_7, x_4, x_5); lean_dec(x_1); -return x_9; +return x_8; } } -else +LEAN_EXPORT lean_object* l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_15; +lean_object* x_5; +x_5 = l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11(x_1, x_2, x_3); lean_dec(x_3); lean_dec_ref(x_2); -lean_dec(x_1); -x_15 = lean_ctor_get(x_8, 0); -lean_inc(x_15); -lean_dec_ref(x_8); -lean_ctor_set(x_6, 0, x_15); -return x_6; +return x_5; } } -else -{ -lean_object* x_16; -x_16 = lean_ctor_get(x_6, 0); -lean_inc(x_16); -lean_dec(x_6); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -lean_inc(x_3); -lean_inc(x_1); -x_17 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache(x_1, x_2, x_3); -if (lean_obj_tag(x_17) == 0) +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__13___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -lean_dec_ref(x_17); -lean_inc(x_18); -x_19 = l_Lean_Compiler_LCNF_CacheExtension_insert___at___00Lean_IR_nameToIRType_spec__1___redArg(x_5, x_1, x_18, x_3); +lean_object* x_15; +x_15 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__13___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec_ref(x_7); +lean_dec_ref(x_6); +lean_dec_ref(x_5); lean_dec(x_3); -if (lean_is_exclusive(x_19)) { - lean_ctor_release(x_19, 0); - x_20 = x_19; -} else { - lean_dec_ref(x_19); - x_20 = lean_box(0); +lean_dec_ref(x_2); +lean_dec(x_1); +return x_15; } -if (lean_is_scalar(x_20)) { - x_21 = lean_alloc_ctor(0, 1, 0); -} else { - x_21 = x_20; } -lean_ctor_set(x_21, 0, x_18); -return x_21; +LEAN_EXPORT lean_object* l_Lean_IR_toIRType___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_IR_toIRType(x_1, x_2, x_3); +return x_5; } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_toIRType_visitApp___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -lean_dec(x_3); -lean_dec(x_1); -return x_17; +lean_object* x_6; +x_6 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_toIRType_visitApp(x_1, x_2, x_3, x_4); +return x_6; } } -else +LEAN_EXPORT lean_object* l_Lean_IR_unboxedIRType___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_22; lean_object* x_23; -lean_dec(x_3); +lean_object* x_6; +x_6 = l_Lean_IR_unboxedIRType(x_1, x_2, x_3, x_4); lean_dec_ref(x_2); -lean_dec(x_1); -x_22 = lean_ctor_get(x_16, 0); -lean_inc(x_22); -lean_dec_ref(x_16); -x_23 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_23, 0, x_22); -return x_23; +return x_6; } } +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor(x_1, x_2, x_3, x_4, x_5); +return x_7; } } -LEAN_EXPORT lean_object* l_Lean_IR_getCtorLayout___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___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_IR_getCtorLayout(x_1, x_2, x_3); +x_5 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache(x_1, x_2, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___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_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg(x_1, x_2, x_3); +lean_dec(x_3); return x_5; } } +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +return x_9; +} +} lean_object* initialize_Lean_Compiler_IR_Format(uint8_t builtin); lean_object* initialize_Lean_Compiler_LCNF_MonoTypes(uint8_t builtin); +lean_object* initialize_Lean_Compiler_IR_UnboxResult(uint8_t builtin); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Compiler_IR_ToIRType(uint8_t builtin) { lean_object * res; @@ -8090,6 +9825,9 @@ lean_dec_ref(res); res = initialize_Lean_Compiler_LCNF_MonoTypes(builtin); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Compiler_IR_UnboxResult(builtin); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); l_Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0___redArg___lam__0___closed__0 = _init_l_Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0___redArg___lam__0___closed__0(); lean_mark_persistent(l_Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0___redArg___lam__0___closed__0); l_Lean_PersistentHashMap_findAux___at___00Lean_PersistentHashMap_find_x3f___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__1_spec__3___redArg___closed__0 = _init_l_Lean_PersistentHashMap_findAux___at___00Lean_PersistentHashMap_find_x3f___at___00Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0_spec__1_spec__3___redArg___closed__0(); @@ -8120,6 +9858,57 @@ if (lean_io_result_is_error(res)) return res; l_Lean_IR_trivialStructureInfoExt = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_IR_trivialStructureInfoExt); lean_dec_ref(res); +}l_Lean_IR_instInhabitedCtorFieldInfo_default = _init_l_Lean_IR_instInhabitedCtorFieldInfo_default(); +lean_mark_persistent(l_Lean_IR_instInhabitedCtorFieldInfo_default); +l_Lean_IR_instInhabitedCtorFieldInfo = _init_l_Lean_IR_instInhabitedCtorFieldInfo(); +lean_mark_persistent(l_Lean_IR_instInhabitedCtorFieldInfo); +l_Lean_IR_CtorFieldInfo_format___closed__0 = _init_l_Lean_IR_CtorFieldInfo_format___closed__0(); +lean_mark_persistent(l_Lean_IR_CtorFieldInfo_format___closed__0); +l_Lean_IR_CtorFieldInfo_format___closed__1 = _init_l_Lean_IR_CtorFieldInfo_format___closed__1(); +lean_mark_persistent(l_Lean_IR_CtorFieldInfo_format___closed__1); +l_Lean_IR_CtorFieldInfo_format___closed__2 = _init_l_Lean_IR_CtorFieldInfo_format___closed__2(); +lean_mark_persistent(l_Lean_IR_CtorFieldInfo_format___closed__2); +l_Lean_IR_CtorFieldInfo_format___closed__3 = _init_l_Lean_IR_CtorFieldInfo_format___closed__3(); +lean_mark_persistent(l_Lean_IR_CtorFieldInfo_format___closed__3); +l_Lean_IR_CtorFieldInfo_format___closed__4 = _init_l_Lean_IR_CtorFieldInfo_format___closed__4(); +lean_mark_persistent(l_Lean_IR_CtorFieldInfo_format___closed__4); +l_Lean_IR_CtorFieldInfo_format___closed__5 = _init_l_Lean_IR_CtorFieldInfo_format___closed__5(); +lean_mark_persistent(l_Lean_IR_CtorFieldInfo_format___closed__5); +l_Lean_IR_CtorFieldInfo_format___closed__6 = _init_l_Lean_IR_CtorFieldInfo_format___closed__6(); +lean_mark_persistent(l_Lean_IR_CtorFieldInfo_format___closed__6); +l_Lean_IR_CtorFieldInfo_format___closed__7 = _init_l_Lean_IR_CtorFieldInfo_format___closed__7(); +lean_mark_persistent(l_Lean_IR_CtorFieldInfo_format___closed__7); +l_Lean_IR_CtorFieldInfo_format___closed__8 = _init_l_Lean_IR_CtorFieldInfo_format___closed__8(); +lean_mark_persistent(l_Lean_IR_CtorFieldInfo_format___closed__8); +l_Lean_IR_CtorFieldInfo_format___closed__9 = _init_l_Lean_IR_CtorFieldInfo_format___closed__9(); +lean_mark_persistent(l_Lean_IR_CtorFieldInfo_format___closed__9); +l_Lean_IR_CtorFieldInfo_format___closed__10 = _init_l_Lean_IR_CtorFieldInfo_format___closed__10(); +lean_mark_persistent(l_Lean_IR_CtorFieldInfo_format___closed__10); +l_Lean_IR_CtorFieldInfo_format___closed__11 = _init_l_Lean_IR_CtorFieldInfo_format___closed__11(); +lean_mark_persistent(l_Lean_IR_CtorFieldInfo_format___closed__11); +l_Lean_IR_CtorFieldInfo_format___closed__12 = _init_l_Lean_IR_CtorFieldInfo_format___closed__12(); +lean_mark_persistent(l_Lean_IR_CtorFieldInfo_format___closed__12); +l_Lean_IR_CtorFieldInfo_format___closed__13 = _init_l_Lean_IR_CtorFieldInfo_format___closed__13(); +lean_mark_persistent(l_Lean_IR_CtorFieldInfo_format___closed__13); +l_Lean_IR_CtorFieldInfo_instToFormat___closed__0 = _init_l_Lean_IR_CtorFieldInfo_instToFormat___closed__0(); +lean_mark_persistent(l_Lean_IR_CtorFieldInfo_instToFormat___closed__0); +l_Lean_IR_CtorFieldInfo_instToFormat = _init_l_Lean_IR_CtorFieldInfo_instToFormat(); +lean_mark_persistent(l_Lean_IR_CtorFieldInfo_instToFormat); +l_Lean_IR_instInhabitedCtorLayout_default___closed__0 = _init_l_Lean_IR_instInhabitedCtorLayout_default___closed__0(); +lean_mark_persistent(l_Lean_IR_instInhabitedCtorLayout_default___closed__0); +l_Lean_IR_instInhabitedCtorLayout_default___closed__1 = _init_l_Lean_IR_instInhabitedCtorLayout_default___closed__1(); +lean_mark_persistent(l_Lean_IR_instInhabitedCtorLayout_default___closed__1); +l_Lean_IR_instInhabitedCtorLayout_default___closed__2 = _init_l_Lean_IR_instInhabitedCtorLayout_default___closed__2(); +lean_mark_persistent(l_Lean_IR_instInhabitedCtorLayout_default___closed__2); +l_Lean_IR_instInhabitedCtorLayout_default = _init_l_Lean_IR_instInhabitedCtorLayout_default(); +lean_mark_persistent(l_Lean_IR_instInhabitedCtorLayout_default); +l_Lean_IR_instInhabitedCtorLayout = _init_l_Lean_IR_instInhabitedCtorLayout(); +lean_mark_persistent(l_Lean_IR_instInhabitedCtorLayout); +if (builtin) {res = l_Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_895574536____hygCtx___hyg_2_(); +if (lean_io_result_is_error(res)) return res; +l_Lean_IR_ctorLayoutExt = lean_io_result_get_value(res); +lean_mark_persistent(l_Lean_IR_ctorLayoutExt); +lean_dec_ref(res); }l_Lean_IR_hasTrivialStructure_x3f___closed__0 = _init_l_Lean_IR_hasTrivialStructure_x3f___closed__0(); lean_mark_persistent(l_Lean_IR_hasTrivialStructure_x3f___closed__0); l_Lean_IR_hasTrivialStructure_x3f___lam__0___closed__0 = _init_l_Lean_IR_hasTrivialStructure_x3f___lam__0___closed__0(); @@ -8194,8 +9983,82 @@ l_Lean_IR_nameToIRType___closed__0 = _init_l_Lean_IR_nameToIRType___closed__0(); lean_mark_persistent(l_Lean_IR_nameToIRType___closed__0); l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_isAnyProducingType___closed__0 = _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_isAnyProducingType___closed__0(); lean_mark_persistent(l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_isAnyProducingType___closed__0); +l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__0 = _init_l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__0(); +lean_mark_persistent(l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__0); +l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__1 = _init_l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__1(); +lean_mark_persistent(l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__1); +l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__2 = _init_l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__2(); +lean_mark_persistent(l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__2); +l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__3 = _init_l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__3(); +lean_mark_persistent(l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__3); +l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__4 = _init_l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__4(); +lean_mark_persistent(l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_unboxedIRType_spec__11_spec__11___closed__4); +l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__0 = _init_l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__0(); +lean_mark_persistent(l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__0); +l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__1 = _init_l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__1(); +lean_mark_persistent(l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__1); +l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__2 = _init_l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__2(); +lean_mark_persistent(l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__2); +l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__3 = _init_l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__3(); +lean_mark_persistent(l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__3); +l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__4 = _init_l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__4(); +lean_mark_persistent(l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__4); +l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__5 = _init_l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__5(); +lean_mark_persistent(l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__5); +l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__6 = _init_l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__6(); +lean_mark_persistent(l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__6); +l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__7 = _init_l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__7(); +lean_mark_persistent(l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__7); +l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__8 = _init_l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__8(); +lean_mark_persistent(l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__8); +l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__9 = _init_l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__9(); +lean_mark_persistent(l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__9); +l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__10 = _init_l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__10(); +lean_mark_persistent(l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__10); +l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__11 = _init_l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__11(); +lean_mark_persistent(l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__11); +l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__12 = _init_l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__12(); +lean_mark_persistent(l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__12); +l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__13 = _init_l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__13(); +lean_mark_persistent(l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19_spec__20___redArg___closed__13); +l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg___closed__0 = _init_l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg___closed__0(); +lean_mark_persistent(l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg___closed__0); +l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg___closed__1 = _init_l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg___closed__1(); +lean_mark_persistent(l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg___closed__1); +l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg___closed__2 = _init_l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg___closed__2(); +lean_mark_persistent(l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg___closed__2); +l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg___closed__3 = _init_l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg___closed__3(); +lean_mark_persistent(l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17___redArg___closed__3); +l_Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19___closed__0 = _init_l_Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19___closed__0(); +lean_mark_persistent(l_Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor_spec__14_spec__15_spec__17_spec__18_spec__19___closed__0); +l_Lean_IR_getCtorLayout___closed__0 = _init_l_Lean_IR_getCtorLayout___closed__0(); +lean_mark_persistent(l_Lean_IR_getCtorLayout___closed__0); +l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__0 = _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__0(); +lean_mark_persistent(l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__0); +l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__1 = _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__1(); +lean_mark_persistent(l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__1); +l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___closed__0 = _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___closed__0(); +lean_mark_persistent(l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___closed__0); +l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___closed__1 = _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___closed__1(); +lean_mark_persistent(l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___closed__1); l_Lean_IR_toIRType___closed__0 = _init_l_Lean_IR_toIRType___closed__0(); lean_mark_persistent(l_Lean_IR_toIRType___closed__0); +l_Lean_getConstInfoInduct___at___00Lean_IR_unboxedIRType_spec__9___closed__0 = _init_l_Lean_getConstInfoInduct___at___00Lean_IR_unboxedIRType_spec__9___closed__0(); +lean_mark_persistent(l_Lean_getConstInfoInduct___at___00Lean_IR_unboxedIRType_spec__9___closed__0); +l_Lean_getConstInfoInduct___at___00Lean_IR_unboxedIRType_spec__9___closed__1 = _init_l_Lean_getConstInfoInduct___at___00Lean_IR_unboxedIRType_spec__9___closed__1(); +lean_mark_persistent(l_Lean_getConstInfoInduct___at___00Lean_IR_unboxedIRType_spec__9___closed__1); +l_Lean_IR_unboxedIRType___closed__0 = _init_l_Lean_IR_unboxedIRType___closed__0(); +lean_mark_persistent(l_Lean_IR_unboxedIRType___closed__0); +l_Lean_IR_unboxedIRType___closed__1 = _init_l_Lean_IR_unboxedIRType___closed__1(); +lean_mark_persistent(l_Lean_IR_unboxedIRType___closed__1); +l_Lean_IR_unboxedIRType___closed__2 = _init_l_Lean_IR_unboxedIRType___closed__2(); +lean_mark_persistent(l_Lean_IR_unboxedIRType___closed__2); +l_Lean_IR_unboxedIRType___closed__3 = _init_l_Lean_IR_unboxedIRType___closed__3(); +lean_mark_persistent(l_Lean_IR_unboxedIRType___closed__3); +l_Lean_IR_unboxedIRType___closed__4 = _init_l_Lean_IR_unboxedIRType___closed__4(); +lean_mark_persistent(l_Lean_IR_unboxedIRType___closed__4); +l_Lean_IR_unboxedIRType___closed__5 = _init_l_Lean_IR_unboxedIRType___closed__5(); +lean_mark_persistent(l_Lean_IR_unboxedIRType___closed__5); l_Lean_IR_toIRType___closed__1 = _init_l_Lean_IR_toIRType___closed__1(); lean_mark_persistent(l_Lean_IR_toIRType___closed__1); l_Lean_IR_toIRType___closed__2 = _init_l_Lean_IR_toIRType___closed__2(); @@ -8204,77 +10067,10 @@ l_Lean_IR_toIRType___closed__3 = _init_l_Lean_IR_toIRType___closed__3(); lean_mark_persistent(l_Lean_IR_toIRType___closed__3); l_Lean_IR_toIRType___closed__4 = _init_l_Lean_IR_toIRType___closed__4(); lean_mark_persistent(l_Lean_IR_toIRType___closed__4); -l_Lean_IR_instInhabitedCtorFieldInfo_default = _init_l_Lean_IR_instInhabitedCtorFieldInfo_default(); -lean_mark_persistent(l_Lean_IR_instInhabitedCtorFieldInfo_default); -l_Lean_IR_instInhabitedCtorFieldInfo = _init_l_Lean_IR_instInhabitedCtorFieldInfo(); -lean_mark_persistent(l_Lean_IR_instInhabitedCtorFieldInfo); -l_Lean_IR_CtorFieldInfo_format___closed__0 = _init_l_Lean_IR_CtorFieldInfo_format___closed__0(); -lean_mark_persistent(l_Lean_IR_CtorFieldInfo_format___closed__0); -l_Lean_IR_CtorFieldInfo_format___closed__1 = _init_l_Lean_IR_CtorFieldInfo_format___closed__1(); -lean_mark_persistent(l_Lean_IR_CtorFieldInfo_format___closed__1); -l_Lean_IR_CtorFieldInfo_format___closed__2 = _init_l_Lean_IR_CtorFieldInfo_format___closed__2(); -lean_mark_persistent(l_Lean_IR_CtorFieldInfo_format___closed__2); -l_Lean_IR_CtorFieldInfo_format___closed__3 = _init_l_Lean_IR_CtorFieldInfo_format___closed__3(); -lean_mark_persistent(l_Lean_IR_CtorFieldInfo_format___closed__3); -l_Lean_IR_CtorFieldInfo_format___closed__4 = _init_l_Lean_IR_CtorFieldInfo_format___closed__4(); -lean_mark_persistent(l_Lean_IR_CtorFieldInfo_format___closed__4); -l_Lean_IR_CtorFieldInfo_format___closed__5 = _init_l_Lean_IR_CtorFieldInfo_format___closed__5(); -lean_mark_persistent(l_Lean_IR_CtorFieldInfo_format___closed__5); -l_Lean_IR_CtorFieldInfo_format___closed__6 = _init_l_Lean_IR_CtorFieldInfo_format___closed__6(); -lean_mark_persistent(l_Lean_IR_CtorFieldInfo_format___closed__6); -l_Lean_IR_CtorFieldInfo_format___closed__7 = _init_l_Lean_IR_CtorFieldInfo_format___closed__7(); -lean_mark_persistent(l_Lean_IR_CtorFieldInfo_format___closed__7); -l_Lean_IR_CtorFieldInfo_format___closed__8 = _init_l_Lean_IR_CtorFieldInfo_format___closed__8(); -lean_mark_persistent(l_Lean_IR_CtorFieldInfo_format___closed__8); -l_Lean_IR_CtorFieldInfo_format___closed__9 = _init_l_Lean_IR_CtorFieldInfo_format___closed__9(); -lean_mark_persistent(l_Lean_IR_CtorFieldInfo_format___closed__9); -l_Lean_IR_CtorFieldInfo_format___closed__10 = _init_l_Lean_IR_CtorFieldInfo_format___closed__10(); -lean_mark_persistent(l_Lean_IR_CtorFieldInfo_format___closed__10); -l_Lean_IR_CtorFieldInfo_format___closed__11 = _init_l_Lean_IR_CtorFieldInfo_format___closed__11(); -lean_mark_persistent(l_Lean_IR_CtorFieldInfo_format___closed__11); -l_Lean_IR_CtorFieldInfo_format___closed__12 = _init_l_Lean_IR_CtorFieldInfo_format___closed__12(); -lean_mark_persistent(l_Lean_IR_CtorFieldInfo_format___closed__12); -l_Lean_IR_CtorFieldInfo_format___closed__13 = _init_l_Lean_IR_CtorFieldInfo_format___closed__13(); -lean_mark_persistent(l_Lean_IR_CtorFieldInfo_format___closed__13); -l_Lean_IR_CtorFieldInfo_instToFormat___closed__0 = _init_l_Lean_IR_CtorFieldInfo_instToFormat___closed__0(); -lean_mark_persistent(l_Lean_IR_CtorFieldInfo_instToFormat___closed__0); -l_Lean_IR_CtorFieldInfo_instToFormat = _init_l_Lean_IR_CtorFieldInfo_instToFormat(); -lean_mark_persistent(l_Lean_IR_CtorFieldInfo_instToFormat); -l_Lean_IR_instInhabitedCtorLayout_default___closed__0 = _init_l_Lean_IR_instInhabitedCtorLayout_default___closed__0(); -lean_mark_persistent(l_Lean_IR_instInhabitedCtorLayout_default___closed__0); -l_Lean_IR_instInhabitedCtorLayout_default___closed__1 = _init_l_Lean_IR_instInhabitedCtorLayout_default___closed__1(); -lean_mark_persistent(l_Lean_IR_instInhabitedCtorLayout_default___closed__1); -l_Lean_IR_instInhabitedCtorLayout_default___closed__2 = _init_l_Lean_IR_instInhabitedCtorLayout_default___closed__2(); -lean_mark_persistent(l_Lean_IR_instInhabitedCtorLayout_default___closed__2); -l_Lean_IR_instInhabitedCtorLayout_default = _init_l_Lean_IR_instInhabitedCtorLayout_default(); -lean_mark_persistent(l_Lean_IR_instInhabitedCtorLayout_default); -l_Lean_IR_instInhabitedCtorLayout = _init_l_Lean_IR_instInhabitedCtorLayout(); -lean_mark_persistent(l_Lean_IR_instInhabitedCtorLayout); -if (builtin) {res = l_Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_895574536____hygCtx___hyg_2_(); -if (lean_io_result_is_error(res)) return res; -l_Lean_IR_ctorLayoutExt = lean_io_result_get_value(res); -lean_mark_persistent(l_Lean_IR_ctorLayoutExt); -lean_dec_ref(res); -}l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__0 = _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__0(); -lean_mark_persistent(l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__0); -l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___closed__0 = _init_l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___closed__0(); -lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___closed__0); -l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___closed__0 = _init_l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___closed__0(); -lean_mark_persistent(l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2___closed__0); -l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__1 = _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__1(); -lean_mark_persistent(l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__1); -l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__2 = _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__2(); -lean_mark_persistent(l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__2); -l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__3 = _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__3(); -lean_mark_persistent(l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__3); -l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__4 = _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__4(); -lean_mark_persistent(l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__4); -l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__5 = _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__5(); -lean_mark_persistent(l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__5); -l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__6 = _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__6(); -lean_mark_persistent(l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache___closed__6); -l_Lean_IR_getCtorLayout___closed__0 = _init_l_Lean_IR_getCtorLayout___closed__0(); -lean_mark_persistent(l_Lean_IR_getCtorLayout___closed__0); +l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___closed__2 = _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___closed__2(); +lean_mark_persistent(l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___closed__2); +l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___closed__3 = _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___closed__3(); +lean_mark_persistent(l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_unboxedIRType_handleCtor___closed__3); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Elab/BuiltinCommand.c b/stage0/stdlib/Lean/Elab/BuiltinCommand.c index 97cd79c08fef..91e20595b215 100644 --- a/stage0/stdlib/Lean/Elab/BuiltinCommand.c +++ b/stage0/stdlib/Lean/Elab/BuiltinCommand.c @@ -22226,8 +22226,8 @@ goto block_42; block_29: { lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -lean_inc_ref(x_19); -x_22 = l_Array_append___redArg(x_19, x_21); +lean_inc_ref(x_15); +x_22 = l_Array_append___redArg(x_15, x_21); lean_dec_ref(x_21); lean_inc(x_18); lean_inc(x_14); @@ -22239,13 +22239,13 @@ lean_inc(x_14); x_24 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_24, 0, x_14); lean_ctor_set(x_24, 1, x_18); -lean_ctor_set(x_24, 2, x_19); +lean_ctor_set(x_24, 2, x_15); x_25 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation_spec__5___closed__1; lean_inc(x_14); x_26 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_26, 0, x_14); lean_ctor_set(x_26, 1, x_25); -x_27 = l_Lean_Syntax_node5(x_14, x_15, x_17, x_20, x_23, x_24, x_26); +x_27 = l_Lean_Syntax_node5(x_14, x_16, x_17, x_20, x_23, x_24, x_26); if (lean_is_scalar(x_12)) { x_28 = lean_alloc_ctor(0, 1, 0); } else { @@ -22279,11 +22279,11 @@ x_39 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_39, 0, x_14); lean_ctor_set(x_39, 1, x_38); x_40 = l_Array_mkArray2___redArg(x_39, x_37); -x_15 = x_31; -x_16 = lean_box(0); +x_15 = x_36; +x_16 = x_31; x_17 = x_33; x_18 = x_34; -x_19 = x_36; +x_19 = lean_box(0); x_20 = x_35; x_21 = x_40; goto block_29; @@ -22293,11 +22293,11 @@ else lean_object* x_41; lean_dec(x_2); x_41 = l_Lean_Elab_OpenDecl_elabOpenDecl___at___00Lean_Elab_Command_elabOpen_spec__0___closed__19; -x_15 = x_31; -x_16 = lean_box(0); +x_15 = x_36; +x_16 = x_31; x_17 = x_33; x_18 = x_34; -x_19 = x_36; +x_19 = lean_box(0); x_20 = x_35; x_21 = x_41; goto block_29; @@ -22420,19 +22420,19 @@ goto block_85; block_72: { lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_66 = l_Array_append___redArg(x_63, x_65); +x_66 = l_Array_append___redArg(x_59, x_65); lean_dec_ref(x_65); lean_inc(x_58); x_67 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_67, 0, x_58); -lean_ctor_set(x_67, 1, x_59); +lean_ctor_set(x_67, 1, x_62); lean_ctor_set(x_67, 2, x_66); x_68 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation_spec__5___closed__3; lean_inc(x_58); x_69 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_69, 0, x_58); lean_ctor_set(x_69, 1, x_68); -x_70 = l_Lean_Syntax_node4(x_58, x_62, x_64, x_61, x_67, x_69); +x_70 = l_Lean_Syntax_node4(x_58, x_61, x_63, x_64, x_67, x_69); if (lean_is_scalar(x_56)) { x_71 = lean_alloc_ctor(0, 1, 0); } else { @@ -22466,12 +22466,12 @@ x_82 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_82, 0, x_58); lean_ctor_set(x_82, 1, x_81); x_83 = l_Array_mkArray2___redArg(x_82, x_80); -x_59 = x_77; +x_59 = x_79; x_60 = lean_box(0); -x_61 = x_78; -x_62 = x_74; -x_63 = x_79; -x_64 = x_76; +x_61 = x_74; +x_62 = x_77; +x_63 = x_76; +x_64 = x_78; x_65 = x_83; goto block_72; } @@ -22480,12 +22480,12 @@ else lean_object* x_84; lean_dec(x_2); x_84 = l_Lean_Elab_OpenDecl_elabOpenDecl___at___00Lean_Elab_Command_elabOpen_spec__0___closed__19; -x_59 = x_77; +x_59 = x_79; x_60 = lean_box(0); -x_61 = x_78; -x_62 = x_74; -x_63 = x_79; -x_64 = x_76; +x_61 = x_74; +x_62 = x_77; +x_63 = x_76; +x_64 = x_78; x_65 = x_84; goto block_72; } @@ -22612,7 +22612,7 @@ lean_dec_ref(x_109); lean_inc(x_101); x_111 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_111, 0, x_101); -lean_ctor_set(x_111, 1, x_106); +lean_ctor_set(x_111, 1, x_105); lean_ctor_set(x_111, 2, x_110); x_112 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation_spec__5___closed__3; lean_inc(x_101); @@ -22621,8 +22621,8 @@ lean_ctor_set(x_113, 0, x_101); lean_ctor_set(x_113, 1, x_112); lean_inc_ref(x_113); lean_inc(x_101); -x_114 = l_Lean_Syntax_node2(x_101, x_103, x_113, x_113); -x_115 = l_Lean_Syntax_node4(x_101, x_102, x_104, x_108, x_111, x_114); +x_114 = l_Lean_Syntax_node2(x_101, x_102, x_113, x_113); +x_115 = l_Lean_Syntax_node4(x_101, x_108, x_103, x_106, x_111, x_114); if (lean_is_scalar(x_99)) { x_116 = lean_alloc_ctor(0, 1, 0); } else { @@ -22660,13 +22660,13 @@ x_129 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_129, 0, x_101); lean_ctor_set(x_129, 1, x_128); x_130 = l_Array_mkArray2___redArg(x_129, x_127); -x_102 = x_119; -x_103 = x_120; -x_104 = x_123; -x_105 = lean_box(0); -x_106 = x_124; +x_102 = x_120; +x_103 = x_123; +x_104 = lean_box(0); +x_105 = x_124; +x_106 = x_125; x_107 = x_126; -x_108 = x_125; +x_108 = x_119; x_109 = x_130; goto block_117; } @@ -22675,13 +22675,13 @@ else lean_object* x_131; lean_dec(x_2); x_131 = l_Lean_Elab_OpenDecl_elabOpenDecl___at___00Lean_Elab_Command_elabOpen_spec__0___closed__19; -x_102 = x_119; -x_103 = x_120; -x_104 = x_123; -x_105 = lean_box(0); -x_106 = x_124; +x_102 = x_120; +x_103 = x_123; +x_104 = lean_box(0); +x_105 = x_124; +x_106 = x_125; x_107 = x_126; -x_108 = x_125; +x_108 = x_119; x_109 = x_131; goto block_117; } @@ -24224,8 +24224,8 @@ if (lean_obj_tag(x_129) == 0) lean_object* x_130; x_130 = l_Lean_Syntax_getArgs(x_122); lean_dec(x_122); -x_39 = x_130; -x_40 = x_123; +x_39 = x_123; +x_40 = x_130; x_41 = x_124; x_42 = x_125; x_43 = lean_box(0); @@ -24239,8 +24239,8 @@ x_131 = l_Lean_Syntax_getArgs(x_122); lean_dec(x_122); if (x_52 == 0) { -x_39 = x_131; -x_40 = x_123; +x_39 = x_123; +x_40 = x_131; x_41 = x_124; x_42 = x_125; x_43 = lean_box(0); @@ -24257,8 +24257,8 @@ x_135 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___0 if (lean_obj_tag(x_135) == 0) { lean_dec_ref(x_135); -x_39 = x_131; -x_40 = x_123; +x_39 = x_123; +x_40 = x_131; x_41 = x_124; x_42 = x_125; x_43 = lean_box(0); @@ -24301,10 +24301,10 @@ uint8_t x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x x_44 = 0; x_45 = lean_box(x_44); x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_40); +lean_ctor_set(x_46, 0, x_39); lean_ctor_set(x_46, 1, x_45); x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_39); +lean_ctor_set(x_47, 0, x_40); lean_ctor_set(x_47, 1, x_46); x_48 = lean_unbox(x_36); lean_dec(x_36); @@ -24628,8 +24628,8 @@ if (lean_obj_tag(x_246) == 0) lean_object* x_247; x_247 = l_Lean_Syntax_getArgs(x_239); lean_dec(x_239); -x_152 = x_247; -x_153 = x_240; +x_152 = x_240; +x_153 = x_247; x_154 = x_241; x_155 = x_242; x_156 = lean_box(0); @@ -24643,8 +24643,8 @@ x_248 = l_Lean_Syntax_getArgs(x_239); lean_dec(x_239); if (x_165 == 0) { -x_152 = x_248; -x_153 = x_240; +x_152 = x_240; +x_153 = x_248; x_154 = x_241; x_155 = x_242; x_156 = lean_box(0); @@ -24661,8 +24661,8 @@ x_252 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___0 if (lean_obj_tag(x_252) == 0) { lean_dec_ref(x_252); -x_152 = x_248; -x_153 = x_240; +x_152 = x_240; +x_153 = x_248; x_154 = x_241; x_155 = x_242; x_156 = lean_box(0); @@ -24706,10 +24706,10 @@ uint8_t x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; uint8 x_157 = 0; x_158 = lean_box(x_157); x_159 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_159, 0, x_153); +lean_ctor_set(x_159, 0, x_152); lean_ctor_set(x_159, 1, x_158); x_160 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_160, 0, x_152); +lean_ctor_set(x_160, 0, x_153); lean_ctor_set(x_160, 1, x_159); x_161 = lean_unbox(x_149); lean_dec(x_149); @@ -25075,8 +25075,8 @@ if (lean_obj_tag(x_370) == 0) lean_object* x_371; x_371 = l_Lean_Syntax_getArgs(x_363); lean_dec(x_363); -x_272 = x_371; -x_273 = x_364; +x_272 = x_364; +x_273 = x_371; x_274 = x_365; x_275 = x_366; x_276 = lean_box(0); @@ -25090,8 +25090,8 @@ x_372 = l_Lean_Syntax_getArgs(x_363); lean_dec(x_363); if (x_285 == 0) { -x_272 = x_372; -x_273 = x_364; +x_272 = x_364; +x_273 = x_372; x_274 = x_365; x_275 = x_366; x_276 = lean_box(0); @@ -25108,8 +25108,8 @@ x_376 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___0 if (lean_obj_tag(x_376) == 0) { lean_dec_ref(x_376); -x_272 = x_372; -x_273 = x_364; +x_272 = x_364; +x_273 = x_372; x_274 = x_365; x_275 = x_366; x_276 = lean_box(0); @@ -25153,10 +25153,10 @@ uint8_t x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; uint8 x_277 = 0; x_278 = lean_box(x_277); x_279 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_279, 0, x_273); +lean_ctor_set(x_279, 0, x_272); lean_ctor_set(x_279, 1, x_278); x_280 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_280, 0, x_272); +lean_ctor_set(x_280, 0, x_273); lean_ctor_set(x_280, 1, x_279); x_281 = lean_unbox(x_268); lean_dec(x_268); @@ -28961,7 +28961,7 @@ return x_7; LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_logInfoAt___at___00Lean_Elab_Command_elabCheckCore_spec__1_spec__1___redArg(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_10; uint8_t x_11; uint8_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; uint8_t x_54; lean_object* x_55; uint8_t x_56; lean_object* x_57; lean_object* x_77; lean_object* x_78; uint8_t x_79; uint8_t x_80; lean_object* x_81; lean_object* x_82; uint8_t x_83; lean_object* x_84; lean_object* x_88; lean_object* x_89; lean_object* x_90; uint8_t x_91; lean_object* x_92; uint8_t x_93; uint8_t x_94; uint8_t x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; uint8_t x_105; uint8_t x_106; uint8_t x_107; uint8_t x_109; uint8_t x_125; +lean_object* x_10; uint8_t x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_50; lean_object* x_51; uint8_t x_52; uint8_t x_53; uint8_t x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_77; lean_object* x_78; uint8_t x_79; uint8_t x_80; lean_object* x_81; uint8_t x_82; lean_object* x_83; lean_object* x_84; lean_object* x_88; uint8_t x_89; lean_object* x_90; uint8_t x_91; lean_object* x_92; lean_object* x_93; uint8_t x_94; uint8_t x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; uint8_t x_104; lean_object* x_105; uint8_t x_106; uint8_t x_107; uint8_t x_109; uint8_t x_125; x_100 = 2; x_125 = l_Lean_instBEqMessageSeverity_beq(x_3, x_100); if (x_125 == 0) @@ -28996,15 +28996,15 @@ lean_ctor_set(x_25, 0, x_21); lean_ctor_set(x_25, 1, x_22); x_26 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_13); +lean_ctor_set(x_26, 1, x_16); x_27 = lean_alloc_ctor(0, 5, 3); -lean_ctor_set(x_27, 0, x_10); -lean_ctor_set(x_27, 1, x_14); -lean_ctor_set(x_27, 2, x_16); +lean_ctor_set(x_27, 0, x_14); +lean_ctor_set(x_27, 1, x_10); +lean_ctor_set(x_27, 2, x_12); lean_ctor_set(x_27, 3, x_15); lean_ctor_set(x_27, 4, x_26); lean_ctor_set_uint8(x_27, sizeof(void*)*5, x_11); -lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 1, x_12); +lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 1, x_13); lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 2, x_4); x_28 = l_Lean_MessageLog_add(x_27, x_24); lean_ctor_set(x_20, 6, x_28); @@ -29041,15 +29041,15 @@ lean_ctor_set(x_41, 0, x_21); lean_ctor_set(x_41, 1, x_22); x_42 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_13); +lean_ctor_set(x_42, 1, x_16); x_43 = lean_alloc_ctor(0, 5, 3); -lean_ctor_set(x_43, 0, x_10); -lean_ctor_set(x_43, 1, x_14); -lean_ctor_set(x_43, 2, x_16); +lean_ctor_set(x_43, 0, x_14); +lean_ctor_set(x_43, 1, x_10); +lean_ctor_set(x_43, 2, x_12); lean_ctor_set(x_43, 3, x_15); lean_ctor_set(x_43, 4, x_42); lean_ctor_set_uint8(x_43, sizeof(void*)*5, x_11); -lean_ctor_set_uint8(x_43, sizeof(void*)*5 + 1, x_12); +lean_ctor_set_uint8(x_43, sizeof(void*)*5 + 1, x_13); lean_ctor_set_uint8(x_43, sizeof(void*)*5 + 2, x_4); x_44 = l_Lean_MessageLog_add(x_43, x_38); x_45 = lean_alloc_ctor(0, 9, 0); @@ -29079,25 +29079,25 @@ if (x_60 == 0) { lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; x_61 = lean_ctor_get(x_59, 0); -lean_inc_ref(x_55); -x_62 = l_Lean_FileMap_toPosition(x_55, x_52); -lean_dec(x_52); -x_63 = l_Lean_FileMap_toPosition(x_55, x_57); +lean_inc_ref(x_56); +x_62 = l_Lean_FileMap_toPosition(x_56, x_51); +lean_dec(x_51); +x_63 = l_Lean_FileMap_toPosition(x_56, x_57); lean_dec(x_57); x_64 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_64, 0, x_63); x_65 = l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_innermostScopeName_x3f___closed__0; -if (x_56 == 0) +if (x_54 == 0) { lean_free_object(x_59); lean_dec_ref(x_50); -x_10 = x_51; -x_11 = x_53; -x_12 = x_54; -x_13 = x_61; -x_14 = x_62; +x_10 = x_62; +x_11 = x_52; +x_12 = x_64; +x_13 = x_53; +x_14 = x_55; x_15 = x_65; -x_16 = x_64; +x_16 = x_61; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -29114,7 +29114,7 @@ lean_object* x_67; lean_dec_ref(x_64); lean_dec_ref(x_62); lean_dec(x_61); -lean_dec_ref(x_51); +lean_dec_ref(x_55); lean_dec_ref(x_7); x_67 = lean_box(0); lean_ctor_set(x_59, 0, x_67); @@ -29123,13 +29123,13 @@ return x_59; else { lean_free_object(x_59); -x_10 = x_51; -x_11 = x_53; -x_12 = x_54; -x_13 = x_61; -x_14 = x_62; +x_10 = x_62; +x_11 = x_52; +x_12 = x_64; +x_13 = x_53; +x_14 = x_55; x_15 = x_65; -x_16 = x_64; +x_16 = x_61; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -29143,24 +29143,24 @@ lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean x_68 = lean_ctor_get(x_59, 0); lean_inc(x_68); lean_dec(x_59); -lean_inc_ref(x_55); -x_69 = l_Lean_FileMap_toPosition(x_55, x_52); -lean_dec(x_52); -x_70 = l_Lean_FileMap_toPosition(x_55, x_57); +lean_inc_ref(x_56); +x_69 = l_Lean_FileMap_toPosition(x_56, x_51); +lean_dec(x_51); +x_70 = l_Lean_FileMap_toPosition(x_56, x_57); lean_dec(x_57); x_71 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_71, 0, x_70); x_72 = l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_innermostScopeName_x3f___closed__0; -if (x_56 == 0) +if (x_54 == 0) { lean_dec_ref(x_50); -x_10 = x_51; -x_11 = x_53; -x_12 = x_54; -x_13 = x_68; -x_14 = x_69; +x_10 = x_69; +x_11 = x_52; +x_12 = x_71; +x_13 = x_53; +x_14 = x_55; x_15 = x_72; -x_16 = x_71; +x_16 = x_68; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -29177,7 +29177,7 @@ lean_object* x_74; lean_object* x_75; lean_dec_ref(x_71); lean_dec_ref(x_69); lean_dec(x_68); -lean_dec_ref(x_51); +lean_dec_ref(x_55); lean_dec_ref(x_7); x_74 = lean_box(0); x_75 = lean_alloc_ctor(0, 1, 0); @@ -29186,13 +29186,13 @@ return x_75; } else { -x_10 = x_51; -x_11 = x_53; -x_12 = x_54; -x_13 = x_68; -x_14 = x_69; +x_10 = x_69; +x_11 = x_52; +x_12 = x_71; +x_13 = x_53; +x_14 = x_55; x_15 = x_72; -x_16 = x_71; +x_16 = x_68; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -29204,16 +29204,16 @@ goto block_49; block_87: { lean_object* x_85; -x_85 = l_Lean_Syntax_getTailPos_x3f(x_82, x_79); -lean_dec(x_82); +x_85 = l_Lean_Syntax_getTailPos_x3f(x_78, x_79); +lean_dec(x_78); if (lean_obj_tag(x_85) == 0) { lean_inc(x_84); x_50 = x_77; -x_51 = x_78; -x_52 = x_84; -x_53 = x_79; -x_54 = x_80; +x_51 = x_84; +x_52 = x_79; +x_53 = x_80; +x_54 = x_82; x_55 = x_81; x_56 = x_83; x_57 = x_84; @@ -29226,10 +29226,10 @@ x_86 = lean_ctor_get(x_85, 0); lean_inc(x_86); lean_dec_ref(x_85); x_50 = x_77; -x_51 = x_78; -x_52 = x_84; -x_53 = x_79; -x_54 = x_80; +x_51 = x_84; +x_52 = x_79; +x_53 = x_80; +x_54 = x_82; x_55 = x_81; x_56 = x_83; x_57 = x_86; @@ -29241,17 +29241,17 @@ goto block_76; lean_object* x_95; lean_object* x_96; x_95 = l_Lean_replaceRef(x_1, x_90); lean_dec(x_90); -x_96 = l_Lean_Syntax_getPos_x3f(x_95, x_91); +x_96 = l_Lean_Syntax_getPos_x3f(x_95, x_89); if (lean_obj_tag(x_96) == 0) { lean_object* x_97; x_97 = lean_unsigned_to_nat(0u); x_77 = x_88; -x_78 = x_89; -x_79 = x_91; +x_78 = x_95; +x_79 = x_89; x_80 = x_94; x_81 = x_92; -x_82 = x_95; +x_82 = x_91; x_83 = x_93; x_84 = x_97; goto block_87; @@ -29263,11 +29263,11 @@ x_98 = lean_ctor_get(x_96, 0); lean_inc(x_98); lean_dec_ref(x_96); x_77 = x_88; -x_78 = x_89; -x_79 = x_91; +x_78 = x_95; +x_79 = x_89; x_80 = x_94; x_81 = x_92; -x_82 = x_95; +x_82 = x_91; x_83 = x_93; x_84 = x_98; goto block_87; @@ -29277,22 +29277,22 @@ goto block_87; { if (x_107 == 0) { -x_88 = x_103; -x_89 = x_102; -x_90 = x_101; -x_91 = x_106; -x_92 = x_104; +x_88 = x_101; +x_89 = x_106; +x_90 = x_102; +x_91 = x_104; +x_92 = x_103; x_93 = x_105; x_94 = x_3; goto block_99; } else { -x_88 = x_103; -x_89 = x_102; -x_90 = x_101; -x_91 = x_106; -x_92 = x_104; +x_88 = x_101; +x_89 = x_106; +x_90 = x_102; +x_91 = x_104; +x_92 = x_103; x_93 = x_105; x_94 = x_100; goto block_99; @@ -29320,11 +29320,11 @@ if (x_119 == 0) lean_inc_ref(x_111); lean_inc_ref(x_110); lean_inc(x_113); -x_101 = x_113; -x_102 = x_110; -x_103 = x_117; -x_104 = x_111; -x_105 = x_114; +x_101 = x_117; +x_102 = x_113; +x_103 = x_110; +x_104 = x_114; +x_105 = x_111; x_106 = x_109; x_107 = x_119; goto block_108; @@ -29337,11 +29337,11 @@ x_121 = l_Lean_Option_get___at___00Lean_Elab_addMacroStack___at___00Lean_throwEr lean_inc_ref(x_111); lean_inc_ref(x_110); lean_inc(x_113); -x_101 = x_113; -x_102 = x_110; -x_103 = x_117; -x_104 = x_111; -x_105 = x_114; +x_101 = x_117; +x_102 = x_113; +x_103 = x_110; +x_104 = x_114; +x_105 = x_111; x_106 = x_109; x_107 = x_121; goto block_108; @@ -32565,7 +32565,7 @@ return x_3; LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_logErrorAt___at___00Lean_Elab_logException___at___00Lean_Elab_Command_failIfSucceeds_spec__0_spec__0_spec__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; uint8_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_82; lean_object* x_83; lean_object* x_84; uint8_t x_85; uint8_t x_86; lean_object* x_87; uint8_t x_116; lean_object* x_117; lean_object* x_118; uint8_t x_119; uint8_t x_120; lean_object* x_121; uint8_t x_125; lean_object* x_126; uint8_t x_127; uint8_t x_128; uint8_t x_139; uint8_t x_140; lean_object* x_141; uint8_t x_142; uint8_t x_143; uint8_t x_145; uint8_t x_158; +lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_82; uint8_t x_83; uint8_t x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; uint8_t x_116; uint8_t x_117; lean_object* x_118; uint8_t x_119; lean_object* x_120; lean_object* x_121; uint8_t x_125; uint8_t x_126; lean_object* x_127; uint8_t x_128; uint8_t x_139; uint8_t x_140; uint8_t x_141; lean_object* x_142; uint8_t x_143; uint8_t x_145; uint8_t x_158; x_139 = 2; x_158 = l_Lean_instBEqMessageSeverity_beq(x_3, x_139); if (x_158 == 0) @@ -32617,15 +32617,15 @@ lean_ctor_set(x_27, 0, x_23); lean_ctor_set(x_27, 1, x_24); x_28 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_28, 1, x_8); +lean_ctor_set(x_28, 1, x_11); x_29 = lean_alloc_ctor(0, 5, 3); -lean_ctor_set(x_29, 0, x_10); -lean_ctor_set(x_29, 1, x_9); -lean_ctor_set(x_29, 2, x_14); -lean_ctor_set(x_29, 3, x_13); +lean_ctor_set(x_29, 0, x_13); +lean_ctor_set(x_29, 1, x_8); +lean_ctor_set(x_29, 2, x_10); +lean_ctor_set(x_29, 3, x_14); lean_ctor_set(x_29, 4, x_28); -lean_ctor_set_uint8(x_29, sizeof(void*)*5, x_11); -lean_ctor_set_uint8(x_29, sizeof(void*)*5 + 1, x_12); +lean_ctor_set_uint8(x_29, sizeof(void*)*5, x_12); +lean_ctor_set_uint8(x_29, sizeof(void*)*5 + 1, x_9); lean_ctor_set_uint8(x_29, sizeof(void*)*5 + 2, x_4); x_30 = l_Lean_MessageLog_add(x_29, x_26); lean_ctor_set(x_22, 1, x_30); @@ -32665,15 +32665,15 @@ lean_ctor_set(x_44, 0, x_23); lean_ctor_set(x_44, 1, x_24); x_45 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_45, 0, x_44); -lean_ctor_set(x_45, 1, x_8); +lean_ctor_set(x_45, 1, x_11); x_46 = lean_alloc_ctor(0, 5, 3); -lean_ctor_set(x_46, 0, x_10); -lean_ctor_set(x_46, 1, x_9); -lean_ctor_set(x_46, 2, x_14); -lean_ctor_set(x_46, 3, x_13); +lean_ctor_set(x_46, 0, x_13); +lean_ctor_set(x_46, 1, x_8); +lean_ctor_set(x_46, 2, x_10); +lean_ctor_set(x_46, 3, x_14); lean_ctor_set(x_46, 4, x_45); -lean_ctor_set_uint8(x_46, sizeof(void*)*5, x_11); -lean_ctor_set_uint8(x_46, sizeof(void*)*5 + 1, x_12); +lean_ctor_set_uint8(x_46, sizeof(void*)*5, x_12); +lean_ctor_set_uint8(x_46, sizeof(void*)*5 + 1, x_9); lean_ctor_set_uint8(x_46, sizeof(void*)*5 + 2, x_4); x_47 = l_Lean_MessageLog_add(x_46, x_34); x_48 = lean_alloc_ctor(0, 11, 0); @@ -32751,15 +32751,15 @@ lean_ctor_set(x_67, 0, x_53); lean_ctor_set(x_67, 1, x_54); x_68 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_68, 0, x_67); -lean_ctor_set(x_68, 1, x_8); +lean_ctor_set(x_68, 1, x_11); x_69 = lean_alloc_ctor(0, 5, 3); -lean_ctor_set(x_69, 0, x_10); -lean_ctor_set(x_69, 1, x_9); -lean_ctor_set(x_69, 2, x_14); -lean_ctor_set(x_69, 3, x_13); +lean_ctor_set(x_69, 0, x_13); +lean_ctor_set(x_69, 1, x_8); +lean_ctor_set(x_69, 2, x_10); +lean_ctor_set(x_69, 3, x_14); lean_ctor_set(x_69, 4, x_68); -lean_ctor_set_uint8(x_69, sizeof(void*)*5, x_11); -lean_ctor_set_uint8(x_69, sizeof(void*)*5 + 1, x_12); +lean_ctor_set_uint8(x_69, sizeof(void*)*5, x_12); +lean_ctor_set_uint8(x_69, sizeof(void*)*5 + 1, x_9); lean_ctor_set_uint8(x_69, sizeof(void*)*5 + 2, x_4); x_70 = l_Lean_MessageLog_add(x_69, x_56); if (lean_is_scalar(x_66)) { @@ -32789,10 +32789,10 @@ else { uint8_t x_75; lean_dec(x_18); -lean_dec(x_14); +lean_dec_ref(x_14); lean_dec_ref(x_13); -lean_dec_ref(x_10); -lean_dec_ref(x_9); +lean_dec_ref(x_11); +lean_dec(x_10); lean_dec_ref(x_8); x_75 = !lean_is_exclusive(x_19); if (x_75 == 0) @@ -32814,10 +32814,10 @@ return x_77; else { uint8_t x_78; -lean_dec(x_14); +lean_dec_ref(x_14); lean_dec_ref(x_13); -lean_dec_ref(x_10); -lean_dec_ref(x_9); +lean_dec_ref(x_11); +lean_dec(x_10); lean_dec_ref(x_8); x_78 = !lean_is_exclusive(x_17); if (x_78 == 0) @@ -32853,8 +32853,8 @@ if (x_93 == 0) lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; x_94 = lean_ctor_get(x_92, 0); lean_inc_ref(x_89); -x_95 = l_Lean_FileMap_toPosition(x_89, x_83); -lean_dec(x_83); +x_95 = l_Lean_FileMap_toPosition(x_89, x_85); +lean_dec(x_85); x_96 = l_Lean_FileMap_toPosition(x_89, x_87); lean_dec(x_87); x_97 = lean_alloc_ctor(1, 1, 0); @@ -32863,13 +32863,13 @@ x_98 = l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_innermostScopeN if (x_90 == 0) { lean_free_object(x_92); -x_8 = x_94; -x_9 = x_95; -x_10 = x_88; -x_11 = x_85; -x_12 = x_86; -x_13 = x_98; -x_14 = x_97; +x_8 = x_95; +x_9 = x_83; +x_10 = x_97; +x_11 = x_94; +x_12 = x_84; +x_13 = x_88; +x_14 = x_98; x_15 = x_6; x_16 = lean_box(0); goto block_81; @@ -32898,13 +32898,13 @@ return x_92; else { lean_free_object(x_92); -x_8 = x_94; -x_9 = x_95; -x_10 = x_88; -x_11 = x_85; -x_12 = x_86; -x_13 = x_98; -x_14 = x_97; +x_8 = x_95; +x_9 = x_83; +x_10 = x_97; +x_11 = x_94; +x_12 = x_84; +x_13 = x_88; +x_14 = x_98; x_15 = x_6; x_16 = lean_box(0); goto block_81; @@ -32918,8 +32918,8 @@ x_104 = lean_ctor_get(x_92, 0); lean_inc(x_104); lean_dec(x_92); lean_inc_ref(x_89); -x_105 = l_Lean_FileMap_toPosition(x_89, x_83); -lean_dec(x_83); +x_105 = l_Lean_FileMap_toPosition(x_89, x_85); +lean_dec(x_85); x_106 = l_Lean_FileMap_toPosition(x_89, x_87); lean_dec(x_87); x_107 = lean_alloc_ctor(1, 1, 0); @@ -32927,13 +32927,13 @@ lean_ctor_set(x_107, 0, x_106); x_108 = l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_innermostScopeName_x3f___closed__0; if (x_90 == 0) { -x_8 = x_104; -x_9 = x_105; -x_10 = x_88; -x_11 = x_85; -x_12 = x_86; -x_13 = x_108; -x_14 = x_107; +x_8 = x_105; +x_9 = x_83; +x_10 = x_107; +x_11 = x_104; +x_12 = x_84; +x_13 = x_88; +x_14 = x_108; x_15 = x_6; x_16 = lean_box(0); goto block_81; @@ -32962,13 +32962,13 @@ return x_114; } else { -x_8 = x_104; -x_9 = x_105; -x_10 = x_88; -x_11 = x_85; -x_12 = x_86; -x_13 = x_108; -x_14 = x_107; +x_8 = x_105; +x_9 = x_83; +x_10 = x_107; +x_11 = x_104; +x_12 = x_84; +x_13 = x_88; +x_14 = x_108; x_15 = x_6; x_16 = lean_box(0); goto block_81; @@ -32979,16 +32979,16 @@ goto block_81; block_124: { lean_object* x_122; -x_122 = l_Lean_Syntax_getTailPos_x3f(x_117, x_119); -lean_dec(x_117); +x_122 = l_Lean_Syntax_getTailPos_x3f(x_118, x_119); +lean_dec(x_118); if (lean_obj_tag(x_122) == 0) { lean_inc(x_121); x_82 = x_116; -x_83 = x_121; -x_84 = lean_box(0); -x_85 = x_119; -x_86 = x_120; +x_83 = x_117; +x_84 = x_119; +x_85 = x_121; +x_86 = lean_box(0); x_87 = x_121; goto block_115; } @@ -32999,10 +32999,10 @@ x_123 = lean_ctor_get(x_122, 0); lean_inc(x_123); lean_dec_ref(x_122); x_82 = x_116; -x_83 = x_121; -x_84 = lean_box(0); -x_85 = x_119; -x_86 = x_120; +x_83 = x_117; +x_84 = x_119; +x_85 = x_121; +x_86 = lean_box(0); x_87 = x_123; goto block_115; } @@ -33019,16 +33019,16 @@ lean_inc(x_130); lean_dec_ref(x_129); x_131 = l_Lean_replaceRef(x_1, x_130); lean_dec(x_130); -x_132 = l_Lean_Syntax_getPos_x3f(x_131, x_127); +x_132 = l_Lean_Syntax_getPos_x3f(x_131, x_126); if (lean_obj_tag(x_132) == 0) { lean_object* x_133; x_133 = lean_unsigned_to_nat(0u); x_116 = x_125; -x_117 = x_131; -x_118 = lean_box(0); -x_119 = x_127; -x_120 = x_128; +x_117 = x_128; +x_118 = x_131; +x_119 = x_126; +x_120 = lean_box(0); x_121 = x_133; goto block_124; } @@ -33039,10 +33039,10 @@ x_134 = lean_ctor_get(x_132, 0); lean_inc(x_134); lean_dec_ref(x_132); x_116 = x_125; -x_117 = x_131; -x_118 = lean_box(0); -x_119 = x_127; -x_120 = x_128; +x_117 = x_128; +x_118 = x_131; +x_119 = x_126; +x_120 = lean_box(0); x_121 = x_134; goto block_124; } @@ -33074,16 +33074,16 @@ return x_137; if (x_143 == 0) { x_125 = x_140; -x_126 = lean_box(0); -x_127 = x_142; +x_126 = x_141; +x_127 = lean_box(0); x_128 = x_3; goto block_138; } else { x_125 = x_140; -x_126 = lean_box(0); -x_127 = x_142; +x_126 = x_141; +x_127 = lean_box(0); x_128 = x_139; goto block_138; } @@ -33109,8 +33109,8 @@ if (x_152 == 0) { lean_dec(x_150); x_140 = x_145; -x_141 = lean_box(0); -x_142 = x_145; +x_141 = x_145; +x_142 = lean_box(0); x_143 = x_152; goto block_144; } @@ -33121,8 +33121,8 @@ x_153 = l_Lean_logAt___at___00Lean_log___at___00Lean_logWarning___at___00Lean_ch x_154 = l_Lean_Option_get___at___00Lean_Elab_addMacroStack___at___00Lean_throwError___at___00Lean_Elab_Command_elabModuleDoc_spec__2_spec__3_spec__4(x_150, x_153); lean_dec(x_150); x_140 = x_145; -x_141 = lean_box(0); -x_142 = x_145; +x_141 = x_145; +x_142 = lean_box(0); x_143 = x_154; goto block_144; } @@ -35187,13 +35187,13 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean x_12 = l_Lean_MessageData_ofFormat(x_11); x_13 = l_Lean_indentD(x_12); x_14 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_14, 0, x_9); +lean_ctor_set(x_14, 0, x_10); lean_ctor_set(x_14, 1, x_13); x_15 = l___private_Lean_Elab_SetOption_0__Lean_Elab_validateOptionValue_throwMistypedOptionValue___at___00Lean_Elab_validateOptionValue___at___00__private_Lean_Elab_SetOption_0__Lean_Elab_elabSetOption_setOption___at___00Lean_Elab_elabSetOption___at___00Lean_Elab_Command_elabSetOption_spec__0_spec__1_spec__3_spec__4___closed__1; x_16 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_16, 0, x_14); lean_ctor_set(x_16, 1, x_15); -x_17 = l_Lean_MessageData_ofExpr(x_10); +x_17 = l_Lean_MessageData_ofExpr(x_9); x_18 = l_Lean_indentD(x_17); x_19 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_19, 0, x_16); @@ -35233,8 +35233,8 @@ x_33 = lean_ctor_get(x_2, 0); x_34 = l_String_quote(x_33); lean_ctor_set_tag(x_2, 3); lean_ctor_set(x_2, 0, x_34); -x_9 = x_31; -x_10 = x_30; +x_9 = x_30; +x_10 = x_31; x_11 = x_2; goto block_29; } @@ -35247,8 +35247,8 @@ lean_dec(x_2); x_36 = l_String_quote(x_35); x_37 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_37, 0, x_36); -x_9 = x_31; -x_10 = x_30; +x_9 = x_30; +x_10 = x_31; x_11 = x_37; goto block_29; } @@ -35262,8 +35262,8 @@ if (x_38 == 0) { lean_object* x_39; x_39 = l___private_Lean_Elab_SetOption_0__Lean_Elab_validateOptionValue_throwMistypedOptionValue___at___00Lean_Elab_validateOptionValue___at___00__private_Lean_Elab_SetOption_0__Lean_Elab_elabSetOption_setOption___at___00Lean_Elab_elabSetOption___at___00Lean_Elab_Command_elabSetOption_spec__0_spec__1_spec__3_spec__4___closed__8; -x_9 = x_31; -x_10 = x_30; +x_9 = x_30; +x_10 = x_31; x_11 = x_39; goto block_29; } @@ -35271,8 +35271,8 @@ else { lean_object* x_40; x_40 = l___private_Lean_Elab_SetOption_0__Lean_Elab_validateOptionValue_throwMistypedOptionValue___at___00Lean_Elab_validateOptionValue___at___00__private_Lean_Elab_SetOption_0__Lean_Elab_elabSetOption_setOption___at___00Lean_Elab_elabSetOption___at___00Lean_Elab_Command_elabSetOption_spec__0_spec__1_spec__3_spec__4___closed__9; -x_9 = x_31; -x_10 = x_30; +x_9 = x_30; +x_10 = x_31; x_11 = x_40; goto block_29; } @@ -35293,8 +35293,8 @@ lean_ctor_set(x_2, 0, x_45); x_46 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_46, 0, x_43); lean_ctor_set(x_46, 1, x_2); -x_9 = x_31; -x_10 = x_30; +x_9 = x_30; +x_10 = x_31; x_11 = x_46; goto block_29; } @@ -35312,8 +35312,8 @@ lean_ctor_set(x_51, 0, x_50); x_52 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_52, 0, x_48); lean_ctor_set(x_52, 1, x_51); -x_9 = x_31; -x_10 = x_30; +x_9 = x_30; +x_10 = x_31; x_11 = x_52; goto block_29; } @@ -35328,8 +35328,8 @@ lean_object* x_54; lean_object* x_55; x_54 = lean_ctor_get(x_2, 0); x_55 = l_Nat_reprFast(x_54); lean_ctor_set(x_2, 0, x_55); -x_9 = x_31; -x_10 = x_30; +x_9 = x_30; +x_10 = x_31; x_11 = x_2; goto block_29; } @@ -35342,8 +35342,8 @@ lean_dec(x_2); x_57 = l_Nat_reprFast(x_56); x_58 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_58, 0, x_57); -x_9 = x_31; -x_10 = x_30; +x_9 = x_30; +x_10 = x_31; x_11 = x_58; goto block_29; } @@ -35366,8 +35366,8 @@ lean_dec(x_60); x_64 = l_Nat_reprFast(x_63); lean_ctor_set_tag(x_2, 3); lean_ctor_set(x_2, 0, x_64); -x_9 = x_31; -x_10 = x_30; +x_9 = x_30; +x_10 = x_31; x_11 = x_2; goto block_29; } @@ -35387,8 +35387,8 @@ x_71 = lean_string_append(x_68, x_70); lean_dec_ref(x_70); lean_ctor_set_tag(x_2, 3); lean_ctor_set(x_2, 0, x_71); -x_9 = x_31; -x_10 = x_30; +x_9 = x_30; +x_10 = x_31; x_11 = x_2; goto block_29; } @@ -35409,8 +35409,8 @@ lean_dec(x_72); x_76 = l_Nat_reprFast(x_75); x_77 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_77, 0, x_76); -x_9 = x_31; -x_10 = x_30; +x_9 = x_30; +x_10 = x_31; x_11 = x_77; goto block_29; } @@ -35430,8 +35430,8 @@ x_84 = lean_string_append(x_81, x_83); lean_dec_ref(x_83); x_85 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_85, 0, x_84); -x_9 = x_31; -x_10 = x_30; +x_9 = x_30; +x_10 = x_31; x_11 = x_85; goto block_29; } @@ -35446,8 +35446,8 @@ lean_dec_ref(x_2); x_87 = lean_box(0); x_88 = 0; x_89 = l_Lean_Syntax_formatStx(x_86, x_87, x_88); -x_9 = x_31; -x_10 = x_30; +x_9 = x_30; +x_10 = x_31; x_11 = x_89; goto block_29; } @@ -40962,9 +40962,9 @@ goto block_47; block_19: { lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_10 = l_List_appendTR___redArg(x_8, x_9); +x_10 = l_List_appendTR___redArg(x_4, x_9); x_11 = l_Lean_Elab_Command_elabVersion___redArg___closed__7; -x_12 = l_Lean_stringToMessageData(x_4); +x_12 = l_Lean_stringToMessageData(x_7); x_13 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_13, 0, x_11); lean_ctor_set(x_13, 1, x_12); @@ -40975,23 +40975,23 @@ x_16 = l_Lean_stringToMessageData(x_15); x_17 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_17, 0, x_13); lean_ctor_set(x_17, 1, x_16); -x_18 = l_Lean_logInfo___at___00Lean_Elab_Command_elabVersion_spec__1(x_17, x_5, x_6); +x_18 = l_Lean_logInfo___at___00Lean_Elab_Command_elabVersion_spec__1(x_17, x_5, x_8); return x_18; } block_30: { lean_object* x_26; uint8_t x_27; -x_26 = l_List_appendTR___redArg(x_22, x_25); +x_26 = l_List_appendTR___redArg(x_20, x_25); x_27 = l_Lean_Elab_Command_elabVersion___redArg___closed__8; if (x_27 == 0) { lean_object* x_28; x_28 = lean_box(0); -x_4 = x_20; +x_4 = x_26; x_5 = x_21; -x_6 = x_23; -x_7 = lean_box(0); -x_8 = x_26; +x_6 = lean_box(0); +x_7 = x_23; +x_8 = x_24; x_9 = x_28; goto block_19; } @@ -40999,11 +40999,11 @@ else { lean_object* x_29; x_29 = l_Lean_Elab_Command_elabVersion___redArg___closed__10; -x_4 = x_20; +x_4 = x_26; x_5 = x_21; -x_6 = x_23; -x_7 = lean_box(0); -x_8 = x_26; +x_6 = lean_box(0); +x_7 = x_23; +x_8 = x_24; x_9 = x_29; goto block_19; } @@ -41016,11 +41016,11 @@ if (x_36 == 0) { lean_object* x_37; x_37 = lean_box(0); -x_20 = x_31; -x_21 = x_32; -x_22 = x_35; +x_20 = x_35; +x_21 = x_31; +x_22 = lean_box(0); x_23 = x_33; -x_24 = lean_box(0); +x_24 = x_34; x_25 = x_37; goto block_30; } @@ -41028,11 +41028,11 @@ else { lean_object* x_38; x_38 = l_Lean_Elab_Command_elabVersion___redArg___closed__13; -x_20 = x_31; -x_21 = x_32; -x_22 = x_35; +x_20 = x_35; +x_21 = x_31; +x_22 = lean_box(0); x_23 = x_33; -x_24 = lean_box(0); +x_24 = x_34; x_25 = x_38; goto block_30; } @@ -41045,10 +41045,10 @@ if (x_44 == 0) { lean_object* x_45; x_45 = lean_box(0); -x_31 = x_40; -x_32 = x_41; -x_33 = x_42; -x_34 = lean_box(0); +x_31 = x_41; +x_32 = lean_box(0); +x_33 = x_40; +x_34 = x_42; x_35 = x_45; goto block_39; } @@ -41056,10 +41056,10 @@ else { lean_object* x_46; x_46 = l_Lean_Elab_Command_elabVersion___redArg___closed__16; -x_31 = x_40; -x_32 = x_41; -x_33 = x_42; -x_34 = lean_box(0); +x_31 = x_41; +x_32 = lean_box(0); +x_33 = x_40; +x_34 = x_42; x_35 = x_46; goto block_39; } @@ -42990,7 +42990,7 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_object* x_31; lean_object* x_43; uint8_t x_44; lean_object* x_45; lean_object* x_57; lean_object* x_58; uint8_t x_59; lean_object* x_60; lean_object* x_72; lean_object* x_73; uint8_t x_74; lean_object* x_75; lean_object* x_87; uint8_t x_88; lean_object* x_89; lean_object* x_103; uint8_t x_104; uint8_t x_105; +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; lean_object* x_31; uint8_t x_43; lean_object* x_44; lean_object* x_45; uint8_t x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; uint8_t x_87; lean_object* x_88; lean_object* x_89; lean_object* x_103; uint8_t x_104; uint8_t x_105; x_8 = lean_ctor_get(x_2, 0); lean_inc(x_8); x_9 = lean_ctor_get(x_2, 1); @@ -43112,30 +43112,30 @@ goto _start; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; x_32 = l_Lean_Elab_Command_elabSetOption___regBuiltin_Lean_Elab_Command_elabSetOption__1___closed__0; x_33 = l_Lean_Elab_Command_elabSetOption___regBuiltin_Lean_Elab_Command_elabSetOption__1___closed__1; -lean_inc(x_29); +lean_inc(x_30); if (lean_is_scalar(x_28)) { x_34 = lean_alloc_ctor(2, 2, 0); } else { x_34 = x_28; lean_ctor_set_tag(x_34, 2); } -lean_ctor_set(x_34, 0, x_29); +lean_ctor_set(x_34, 0, x_30); lean_ctor_set(x_34, 1, x_32); x_35 = lean_mk_syntax_ident(x_26); x_36 = l_Lean_Elab_Command_elabSection___closed__10; x_37 = l_Lean_Elab_Command_elabSection___closed__11; -lean_inc(x_29); +lean_inc(x_30); x_38 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_38, 0, x_29); +lean_ctor_set(x_38, 0, x_30); lean_ctor_set(x_38, 1, x_36); lean_ctor_set(x_38, 2, x_37); x_39 = l_Lean_Elab_elabSetOption___at___00Lean_Elab_Command_elabSetOption_spec__0___closed__5; -lean_inc(x_29); +lean_inc(x_30); x_40 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_40, 0, x_29); +lean_ctor_set(x_40, 0, x_30); lean_ctor_set(x_40, 1, x_39); -x_41 = l_Lean_Syntax_node4(x_29, x_33, x_34, x_35, x_38, x_40); -x_11 = x_30; +x_41 = l_Lean_Syntax_node4(x_30, x_33, x_34, x_35, x_38, x_40); +x_11 = x_29; x_12 = x_3; x_13 = x_41; x_14 = lean_box(0); @@ -43146,25 +43146,25 @@ goto block_25; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; x_46 = l_Lean_Elab_Command_elabSetOption___regBuiltin_Lean_Elab_Command_elabSetOption__1___closed__0; x_47 = l_Lean_Elab_Command_elabSetOption___regBuiltin_Lean_Elab_Command_elabSetOption__1___closed__1; -lean_inc(x_43); +lean_inc(x_44); x_48 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_48, 0, x_43); +lean_ctor_set(x_48, 0, x_44); lean_ctor_set(x_48, 1, x_46); x_49 = lean_mk_syntax_ident(x_26); x_50 = l_Lean_Elab_Command_elabSection___closed__10; x_51 = l_Lean_Elab_Command_elabSection___closed__11; -lean_inc(x_43); +lean_inc(x_44); x_52 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_52, 0, x_43); +lean_ctor_set(x_52, 0, x_44); lean_ctor_set(x_52, 1, x_50); lean_ctor_set(x_52, 2, x_51); x_53 = l_Lean_Elab_elabSetOption___at___00Lean_Elab_Command_elabSetOption_spec__0___closed__4; -lean_inc(x_43); +lean_inc(x_44); x_54 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_54, 0, x_43); +lean_ctor_set(x_54, 0, x_44); lean_ctor_set(x_54, 1, x_53); -x_55 = l_Lean_Syntax_node4(x_43, x_47, x_48, x_49, x_52, x_54); -x_11 = x_44; +x_55 = l_Lean_Syntax_node4(x_44, x_47, x_48, x_49, x_52, x_54); +x_11 = x_43; x_12 = x_3; x_13 = x_55; x_14 = lean_box(0); @@ -43175,22 +43175,22 @@ goto block_25; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; x_61 = l_Lean_Elab_Command_elabSetOption___regBuiltin_Lean_Elab_Command_elabSetOption__1___closed__0; x_62 = l_Lean_Elab_Command_elabSetOption___regBuiltin_Lean_Elab_Command_elabSetOption__1___closed__1; -lean_inc(x_57); +lean_inc(x_59); x_63 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_63, 0, x_57); +lean_ctor_set(x_63, 0, x_59); lean_ctor_set(x_63, 1, x_61); x_64 = lean_mk_syntax_ident(x_26); x_65 = l_Lean_Elab_Command_elabSection___closed__10; x_66 = l_Lean_Elab_Command_elabSection___closed__11; -lean_inc(x_57); +lean_inc(x_59); x_67 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_67, 0, x_57); +lean_ctor_set(x_67, 0, x_59); lean_ctor_set(x_67, 1, x_65); lean_ctor_set(x_67, 2, x_66); x_68 = lean_box(2); x_69 = l_Lean_Syntax_mkStrLit(x_58, x_68); -x_70 = l_Lean_Syntax_node4(x_57, x_62, x_63, x_64, x_67, x_69); -x_11 = x_59; +x_70 = l_Lean_Syntax_node4(x_59, x_62, x_63, x_64, x_67, x_69); +x_11 = x_57; x_12 = x_3; x_13 = x_70; x_14 = lean_box(0); @@ -43201,22 +43201,22 @@ goto block_25; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; x_76 = l_Lean_Elab_Command_elabSetOption___regBuiltin_Lean_Elab_Command_elabSetOption__1___closed__0; x_77 = l_Lean_Elab_Command_elabSetOption___regBuiltin_Lean_Elab_Command_elabSetOption__1___closed__1; -lean_inc(x_73); +lean_inc(x_74); x_78 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_78, 0, x_73); +lean_ctor_set(x_78, 0, x_74); lean_ctor_set(x_78, 1, x_76); x_79 = lean_mk_syntax_ident(x_26); x_80 = l_Lean_Elab_Command_elabSection___closed__10; x_81 = l_Lean_Elab_Command_elabSection___closed__11; -lean_inc(x_73); +lean_inc(x_74); x_82 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_82, 0, x_73); +lean_ctor_set(x_82, 0, x_74); lean_ctor_set(x_82, 1, x_80); lean_ctor_set(x_82, 2, x_81); x_83 = lean_box(2); -x_84 = l_Lean_Syntax_mkNatLit(x_72, x_83); -x_85 = l_Lean_Syntax_node4(x_73, x_77, x_78, x_79, x_82, x_84); -x_11 = x_74; +x_84 = l_Lean_Syntax_mkNatLit(x_73, x_83); +x_85 = l_Lean_Syntax_node4(x_74, x_77, x_78, x_79, x_82, x_84); +x_11 = x_72; x_12 = x_3; x_13 = x_85; x_14 = lean_box(0); @@ -43227,28 +43227,28 @@ goto block_25; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; x_90 = l_Lean_Elab_Command_elabSetOption___regBuiltin_Lean_Elab_Command_elabSetOption__1___closed__0; x_91 = l_Lean_Elab_Command_elabSetOption___regBuiltin_Lean_Elab_Command_elabSetOption__1___closed__1; -lean_inc(x_87); +lean_inc(x_88); x_92 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_92, 0, x_87); +lean_ctor_set(x_92, 0, x_88); lean_ctor_set(x_92, 1, x_90); x_93 = lean_mk_syntax_ident(x_26); x_94 = l_Lean_Elab_Command_elabSection___closed__10; x_95 = l_Lean_Elab_Command_elabSection___closed__11; -lean_inc(x_87); +lean_inc(x_88); x_96 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_96, 0, x_87); +lean_ctor_set(x_96, 0, x_88); lean_ctor_set(x_96, 1, x_94); lean_ctor_set(x_96, 2, x_95); x_97 = l_List_forIn_x27_loop___at___00__private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_elabWhere_describeOptions_spec__0___redArg___closed__5; x_98 = l_List_forIn_x27_loop___at___00__private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_elabWhere_describeOptions_spec__0___redArg___closed__6; -lean_inc(x_87); +lean_inc(x_88); x_99 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_99, 0, x_87); +lean_ctor_set(x_99, 0, x_88); lean_ctor_set(x_99, 1, x_98); -lean_inc(x_87); -x_100 = l_Lean_Syntax_node1(x_87, x_97, x_99); -x_101 = l_Lean_Syntax_node4(x_87, x_91, x_92, x_93, x_96, x_100); -x_11 = x_88; +lean_inc(x_88); +x_100 = l_Lean_Syntax_node1(x_88, x_97, x_99); +x_101 = l_Lean_Syntax_node4(x_88, x_91, x_92, x_93, x_96, x_100); +x_11 = x_87; x_12 = x_3; x_13 = x_101; x_14 = lean_box(0); @@ -43285,15 +43285,15 @@ if (lean_obj_tag(x_110) == 0) lean_object* x_112; x_112 = l_Lean_getMainModule___at___00Lean_Elab_Command_elabSection_spec__0___redArg(x_5); lean_dec_ref(x_112); -x_29 = x_111; -x_30 = x_105; +x_29 = x_105; +x_30 = x_111; x_31 = lean_box(0); goto block_42; } else { -x_29 = x_111; -x_30 = x_105; +x_29 = x_105; +x_30 = x_111; x_31 = lean_box(0); goto block_42; } @@ -43373,15 +43373,15 @@ if (lean_obj_tag(x_122) == 0) lean_object* x_124; x_124 = l_Lean_getMainModule___at___00Lean_Elab_Command_elabSection_spec__0___redArg(x_5); lean_dec_ref(x_124); -x_43 = x_123; -x_44 = x_105; +x_43 = x_105; +x_44 = x_123; x_45 = lean_box(0); goto block_56; } else { -x_43 = x_123; -x_44 = x_105; +x_43 = x_105; +x_44 = x_123; x_45 = lean_box(0); goto block_56; } @@ -43463,17 +43463,17 @@ if (lean_obj_tag(x_135) == 0) lean_object* x_137; x_137 = l_Lean_getMainModule___at___00Lean_Elab_Command_elabSection_spec__0___redArg(x_5); lean_dec_ref(x_137); -x_57 = x_136; +x_57 = x_105; x_58 = x_131; -x_59 = x_105; +x_59 = x_136; x_60 = lean_box(0); goto block_71; } else { -x_57 = x_136; +x_57 = x_105; x_58 = x_131; -x_59 = x_105; +x_59 = x_136; x_60 = lean_box(0); goto block_71; } @@ -43556,17 +43556,17 @@ if (lean_obj_tag(x_148) == 0) lean_object* x_150; x_150 = l_Lean_getMainModule___at___00Lean_Elab_Command_elabSection_spec__0___redArg(x_5); lean_dec_ref(x_150); -x_72 = x_144; -x_73 = x_149; -x_74 = x_105; +x_72 = x_105; +x_73 = x_144; +x_74 = x_149; x_75 = lean_box(0); goto block_86; } else { -x_72 = x_144; -x_73 = x_149; -x_74 = x_105; +x_72 = x_105; +x_73 = x_144; +x_74 = x_149; x_75 = lean_box(0); goto block_86; } @@ -43647,15 +43647,15 @@ if (lean_obj_tag(x_160) == 0) lean_object* x_162; x_162 = l_Lean_getMainModule___at___00Lean_Elab_Command_elabSection_spec__0___redArg(x_5); lean_dec_ref(x_162); -x_87 = x_161; -x_88 = x_105; +x_87 = x_105; +x_88 = x_161; x_89 = lean_box(0); goto block_102; } else { -x_87 = x_161; -x_88 = x_105; +x_87 = x_105; +x_88 = x_161; x_89 = lean_box(0); goto block_102; } @@ -44266,7 +44266,7 @@ lean_ctor_set(x_45, 0, x_40); lean_ctor_set(x_45, 1, x_43); x_46 = l_Lean_Elab_Command_elabSection___closed__10; x_47 = l_Lean_Elab_Command_elabSection___closed__11; -x_48 = lean_array_mk(x_39); +x_48 = lean_array_mk(x_37); x_49 = lean_array_size(x_48); x_50 = 0; x_51 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabWhere_spec__0(x_49, x_50, x_48); @@ -44279,10 +44279,10 @@ lean_ctor_set(x_53, 1, x_46); lean_ctor_set(x_53, 2, x_52); x_54 = l_Lean_Syntax_node2(x_40, x_44, x_45, x_53); x_55 = l_Lean_MessageData_ofSyntax(x_54); -x_56 = lean_array_push(x_41, x_55); +x_56 = lean_array_push(x_39, x_55); x_25 = x_56; -x_26 = x_38; -x_27 = x_37; +x_26 = x_41; +x_27 = x_38; x_28 = lean_box(0); goto block_36; } @@ -44313,21 +44313,21 @@ if (lean_obj_tag(x_66) == 0) lean_object* x_68; x_68 = l_Lean_getMainModule___at___00Lean_Elab_Command_elabSection_spec__0___redArg(x_60); lean_dec_ref(x_68); -x_37 = x_60; -x_38 = x_59; -x_39 = x_23; +x_37 = x_23; +x_38 = x_60; +x_39 = x_58; x_40 = x_67; -x_41 = x_58; +x_41 = x_59; x_42 = lean_box(0); goto block_57; } else { -x_37 = x_60; -x_38 = x_59; -x_39 = x_23; +x_37 = x_23; +x_38 = x_60; +x_39 = x_58; x_40 = x_67; -x_41 = x_58; +x_41 = x_59; x_42 = lean_box(0); goto block_57; } @@ -44396,25 +44396,25 @@ goto block_36; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; x_82 = l_Lean_Elab_Command_elabVariable___closed__0; x_83 = l_Lean_Elab_Command_elabVariable___closed__1; -lean_inc(x_80); +lean_inc(x_76); x_84 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_84, 0, x_80); +lean_ctor_set(x_84, 0, x_76); lean_ctor_set(x_84, 1, x_82); x_85 = l_Lean_Elab_Command_elabSection___closed__10; x_86 = l_Lean_Elab_Command_elabSection___closed__11; -x_87 = l_Array_append___redArg(x_86, x_76); -lean_dec_ref(x_76); -lean_inc(x_80); +x_87 = l_Array_append___redArg(x_86, x_77); +lean_dec_ref(x_77); +lean_inc(x_76); x_88 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_88, 0, x_80); +lean_ctor_set(x_88, 0, x_76); lean_ctor_set(x_88, 1, x_85); lean_ctor_set(x_88, 2, x_87); -x_89 = l_Lean_Syntax_node2(x_80, x_83, x_84, x_88); +x_89 = l_Lean_Syntax_node2(x_76, x_83, x_84, x_88); x_90 = l_Lean_MessageData_ofSyntax(x_89); -x_91 = lean_array_push(x_77, x_90); +x_91 = lean_array_push(x_80, x_90); x_58 = x_91; -x_59 = x_78; -x_60 = x_79; +x_59 = x_79; +x_60 = x_78; x_61 = lean_box(0); goto block_75; } @@ -44448,21 +44448,21 @@ if (lean_obj_tag(x_101) == 0) lean_object* x_106; x_106 = l_Lean_getMainModule___at___00Lean_Elab_Command_elabSection_spec__0___redArg(x_95); lean_dec_ref(x_106); -x_76 = x_104; -x_77 = x_93; -x_78 = x_94; -x_79 = x_95; -x_80 = x_105; +x_76 = x_105; +x_77 = x_104; +x_78 = x_95; +x_79 = x_94; +x_80 = x_93; x_81 = lean_box(0); goto block_92; } else { -x_76 = x_104; -x_77 = x_93; -x_78 = x_94; -x_79 = x_95; -x_80 = x_105; +x_76 = x_105; +x_77 = x_104; +x_78 = x_95; +x_79 = x_94; +x_80 = x_93; x_81 = lean_box(0); goto block_92; } @@ -44533,26 +44533,26 @@ goto block_75; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; x_120 = l_Lean_Elab_Command_elabUniverse___regBuiltin_Lean_Elab_Command_elabUniverse__1___closed__0; x_121 = l_Lean_Elab_Command_elabUniverse___regBuiltin_Lean_Elab_Command_elabUniverse__1___closed__1; -lean_inc(x_118); +lean_inc(x_117); x_122 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_122, 0, x_118); +lean_ctor_set(x_122, 0, x_117); lean_ctor_set(x_122, 1, x_120); x_123 = l_Lean_Elab_Command_elabSection___closed__10; x_124 = l_Lean_Elab_Command_elabSection___closed__11; -x_125 = lean_array_mk(x_115); +x_125 = lean_array_mk(x_114); x_126 = l_Array_append___redArg(x_124, x_125); lean_dec_ref(x_125); -lean_inc(x_118); +lean_inc(x_117); x_127 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_127, 0, x_118); +lean_ctor_set(x_127, 0, x_117); lean_ctor_set(x_127, 1, x_123); lean_ctor_set(x_127, 2, x_126); -x_128 = l_Lean_Syntax_node2(x_118, x_121, x_122, x_127); +x_128 = l_Lean_Syntax_node2(x_117, x_121, x_122, x_127); x_129 = l_Lean_MessageData_ofSyntax(x_128); -x_130 = lean_array_push(x_114, x_129); +x_130 = lean_array_push(x_118, x_129); x_93 = x_130; -x_94 = x_117; -x_95 = x_116; +x_94 = x_116; +x_95 = x_115; x_96 = lean_box(0); goto block_113; } @@ -44586,21 +44586,21 @@ if (lean_obj_tag(x_140) == 0) lean_object* x_145; x_145 = l_Lean_getMainModule___at___00Lean_Elab_Command_elabSection_spec__0___redArg(x_134); lean_dec_ref(x_145); -x_114 = x_132; -x_115 = x_143; -x_116 = x_134; -x_117 = x_133; -x_118 = x_144; +x_114 = x_143; +x_115 = x_134; +x_116 = x_133; +x_117 = x_144; +x_118 = x_132; x_119 = lean_box(0); goto block_131; } else { -x_114 = x_132; -x_115 = x_143; -x_116 = x_134; -x_117 = x_133; -x_118 = x_144; +x_114 = x_143; +x_115 = x_134; +x_116 = x_133; +x_117 = x_144; +x_118 = x_132; x_119 = lean_box(0); goto block_131; } @@ -44733,17 +44733,17 @@ return x_164; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; x_172 = l_Lean_Elab_Command_elabNamespace___closed__0; x_173 = l_Lean_Elab_Command_elabNamespace___closed__1; -lean_inc(x_169); +lean_inc(x_170); x_174 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_174, 0, x_169); +lean_ctor_set(x_174, 0, x_170); lean_ctor_set(x_174, 1, x_172); -x_175 = lean_mk_syntax_ident(x_168); -x_176 = l_Lean_Syntax_node2(x_169, x_173, x_174, x_175); +x_175 = lean_mk_syntax_ident(x_169); +x_176 = l_Lean_Syntax_node2(x_170, x_173, x_174, x_175); x_177 = l_Lean_MessageData_ofSyntax(x_176); -x_178 = lean_array_push(x_167, x_177); +x_178 = lean_array_push(x_166, x_177); x_153 = x_178; -x_154 = x_170; -x_155 = x_166; +x_154 = x_168; +x_155 = x_167; x_156 = lean_box(0); goto block_165; } @@ -44774,21 +44774,21 @@ if (lean_obj_tag(x_188) == 0) lean_object* x_190; x_190 = l_Lean_getMainModule___at___00Lean_Elab_Command_elabSection_spec__0___redArg(x_182); lean_dec_ref(x_190); -x_166 = x_182; -x_167 = x_180; -x_168 = x_19; -x_169 = x_189; -x_170 = x_181; +x_166 = x_180; +x_167 = x_182; +x_168 = x_181; +x_169 = x_19; +x_170 = x_189; x_171 = lean_box(0); goto block_179; } else { -x_166 = x_182; -x_167 = x_180; -x_168 = x_19; -x_169 = x_189; -x_170 = x_181; +x_166 = x_180; +x_167 = x_182; +x_168 = x_181; +x_169 = x_19; +x_170 = x_189; x_171 = lean_box(0); goto block_179; } diff --git a/stage0/stdlib/Lean/Elab/BuiltinTerm.c b/stage0/stdlib/Lean/Elab/BuiltinTerm.c index b944b058adfd..1f06d9dfbf24 100644 --- a/stage0/stdlib/Lean/Elab/BuiltinTerm.c +++ b/stage0/stdlib/Lean/Elab/BuiltinTerm.c @@ -10948,8 +10948,8 @@ x_19 = lean_ctor_get(x_18, 0); lean_inc(x_19); lean_dec_ref(x_18); x_20 = l_Lean_Elab_Term_elabNumLit___closed__1; -lean_inc(x_16); -x_21 = l_Nat_reprFast(x_16); +lean_inc(x_14); +x_21 = l_Nat_reprFast(x_14); x_22 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_22, 0, x_21); x_23 = l_Lean_MessageData_ofFormat(x_22); @@ -10960,8 +10960,8 @@ x_25 = l_Lean_Elab_Term_elabNumLit___closed__3; x_26 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_26, 0, x_24); lean_ctor_set(x_26, 1, x_25); -lean_inc_ref(x_13); -x_27 = l_Lean_indentExpr(x_13); +lean_inc_ref(x_11); +x_27 = l_Lean_indentExpr(x_11); x_28 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_28, 0, x_26); lean_ctor_set(x_28, 1, x_27); @@ -10976,14 +10976,14 @@ lean_ctor_set(x_33, 0, x_19); lean_ctor_set(x_33, 1, x_32); lean_inc_ref(x_33); x_34 = l_Lean_mkConst(x_31, x_33); -x_35 = l_Lean_mkRawNatLit(x_16); +x_35 = l_Lean_mkRawNatLit(x_14); lean_inc_ref(x_35); -lean_inc_ref(x_13); -x_36 = l_Lean_mkAppB(x_34, x_13, x_35); +lean_inc_ref(x_11); +x_36 = l_Lean_mkAppB(x_34, x_11, x_35); x_37 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_37, 0, x_30); -lean_inc_ref(x_10); -x_38 = l_Lean_Elab_Term_mkInstMVar(x_36, x_37, x_12, x_15, x_10, x_14, x_11, x_17); +lean_inc_ref(x_17); +x_38 = l_Lean_Elab_Term_mkInstMVar(x_36, x_37, x_13, x_16, x_17, x_12, x_15, x_10); if (lean_obj_tag(x_38) == 0) { lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; @@ -10993,12 +10993,12 @@ lean_dec_ref(x_38); x_40 = l_Lean_Elab_Term_elabNumLit___closed__9; x_41 = l_Lean_mkConst(x_40, x_33); lean_inc(x_39); -x_42 = l_Lean_mkApp3(x_41, x_13, x_35, x_39); +x_42 = l_Lean_mkApp3(x_41, x_11, x_35, x_39); x_43 = l_Lean_Expr_mvarId_x21(x_39); lean_dec(x_39); lean_inc_ref(x_42); -x_44 = l_Lean_Elab_Term_registerMVarErrorImplicitArgInfo___redArg(x_43, x_1, x_42, x_15, x_10); -lean_dec_ref(x_10); +x_44 = l_Lean_Elab_Term_registerMVarErrorImplicitArgInfo___redArg(x_43, x_1, x_42, x_16, x_17); +lean_dec_ref(x_17); if (lean_obj_tag(x_44) == 0) { uint8_t x_45; @@ -11045,8 +11045,8 @@ else { lean_dec_ref(x_35); lean_dec_ref(x_33); -lean_dec_ref(x_13); -lean_dec_ref(x_10); +lean_dec_ref(x_17); +lean_dec_ref(x_11); lean_dec(x_1); return x_38; } @@ -11054,13 +11054,13 @@ return x_38; else { uint8_t x_51; -lean_dec(x_17); -lean_dec(x_16); +lean_dec_ref(x_17); +lean_dec_ref(x_15); lean_dec(x_14); lean_dec_ref(x_13); -lean_dec_ref(x_12); +lean_dec(x_12); lean_dec_ref(x_11); -lean_dec_ref(x_10); +lean_dec(x_10); lean_dec(x_1); x_51 = !lean_is_exclusive(x_18); if (x_51 == 0) @@ -11085,30 +11085,30 @@ if (x_65 == 0) { if (lean_obj_tag(x_2) == 0) { -x_10 = x_55; -x_11 = x_58; -x_12 = x_57; -x_13 = x_56; -x_14 = x_62; -x_15 = x_61; -x_16 = x_63; +x_10 = x_56; +x_11 = x_55; +x_12 = x_58; +x_13 = x_61; +x_14 = x_60; +x_15 = x_63; +x_16 = x_62; x_17 = x_64; -x_18 = x_60; +x_18 = x_57; goto block_54; } else { lean_object* x_66; lean_object* x_67; -lean_dec_ref(x_60); +lean_dec_ref(x_57); x_66 = lean_ctor_get(x_2, 0); lean_inc(x_66); lean_dec_ref(x_2); -lean_inc(x_64); -lean_inc_ref(x_58); -lean_inc(x_62); -lean_inc_ref(x_55); +lean_inc(x_56); +lean_inc_ref(x_63); +lean_inc(x_58); +lean_inc_ref(x_64); lean_inc(x_66); -x_67 = l_Lean_Meta_isProp(x_66, x_55, x_62, x_58, x_64); +x_67 = l_Lean_Meta_isProp(x_66, x_64, x_58, x_63, x_56); if (lean_obj_tag(x_67) == 0) { lean_object* x_68; uint8_t x_69; @@ -11120,12 +11120,12 @@ lean_dec(x_68); if (x_69 == 0) { lean_object* x_70; -lean_inc(x_64); -lean_inc_ref(x_58); -lean_inc(x_62); -lean_inc_ref(x_55); +lean_inc(x_56); +lean_inc_ref(x_63); +lean_inc(x_58); +lean_inc_ref(x_64); lean_inc(x_66); -x_70 = lean_infer_type(x_66, x_55, x_62, x_58, x_64); +x_70 = lean_infer_type(x_66, x_64, x_58, x_63, x_56); if (lean_obj_tag(x_70) == 0) { lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; @@ -11145,15 +11145,15 @@ x_77 = l_Lean_MessageData_ofExpr(x_71); x_78 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_78, 0, x_76); lean_ctor_set(x_78, 1, x_77); -lean_inc_ref(x_57); -x_79 = l_Lean_throwError___at___00Lean_throwErrorAt___at___00Lean_Elab_Term_elabPipeCompletion_spec__0_spec__0___redArg(x_78, x_57, x_61, x_55, x_62, x_58, x_64); -x_10 = x_55; -x_11 = x_58; -x_12 = x_57; -x_13 = x_56; -x_14 = x_62; -x_15 = x_61; -x_16 = x_63; +lean_inc_ref(x_61); +x_79 = l_Lean_throwError___at___00Lean_throwErrorAt___at___00Lean_Elab_Term_elabPipeCompletion_spec__0_spec__0___redArg(x_78, x_61, x_62, x_64, x_58, x_63, x_56); +x_10 = x_56; +x_11 = x_55; +x_12 = x_58; +x_13 = x_61; +x_14 = x_60; +x_15 = x_63; +x_16 = x_62; x_17 = x_64; x_18 = x_79; goto block_54; @@ -11161,12 +11161,12 @@ goto block_54; else { lean_dec(x_66); -lean_dec(x_64); -lean_dec(x_63); -lean_dec(x_62); -lean_dec_ref(x_58); -lean_dec_ref(x_57); -lean_dec_ref(x_56); +lean_dec_ref(x_64); +lean_dec_ref(x_63); +lean_dec_ref(x_61); +lean_dec(x_60); +lean_dec(x_58); +lean_dec(x_56); lean_dec_ref(x_55); lean_dec(x_1); return x_70; @@ -11184,15 +11184,15 @@ x_83 = l_Lean_Elab_Term_elabNumLit___closed__17; x_84 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_84, 0, x_82); lean_ctor_set(x_84, 1, x_83); -lean_inc_ref(x_57); -x_85 = l_Lean_throwError___at___00Lean_throwErrorAt___at___00Lean_Elab_Term_elabPipeCompletion_spec__0_spec__0___redArg(x_84, x_57, x_61, x_55, x_62, x_58, x_64); -x_10 = x_55; -x_11 = x_58; -x_12 = x_57; -x_13 = x_56; -x_14 = x_62; -x_15 = x_61; -x_16 = x_63; +lean_inc_ref(x_61); +x_85 = l_Lean_throwError___at___00Lean_throwErrorAt___at___00Lean_Elab_Term_elabPipeCompletion_spec__0_spec__0___redArg(x_84, x_61, x_62, x_64, x_58, x_63, x_56); +x_10 = x_56; +x_11 = x_55; +x_12 = x_58; +x_13 = x_61; +x_14 = x_60; +x_15 = x_63; +x_16 = x_62; x_17 = x_64; x_18 = x_85; goto block_54; @@ -11202,12 +11202,12 @@ else { uint8_t x_86; lean_dec(x_66); -lean_dec(x_64); -lean_dec(x_63); -lean_dec(x_62); -lean_dec_ref(x_58); -lean_dec_ref(x_57); -lean_dec_ref(x_56); +lean_dec_ref(x_64); +lean_dec_ref(x_63); +lean_dec_ref(x_61); +lean_dec(x_60); +lean_dec(x_58); +lean_dec(x_56); lean_dec_ref(x_55); lean_dec(x_1); x_86 = !lean_is_exclusive(x_67); @@ -11231,15 +11231,15 @@ return x_88; else { lean_dec(x_2); -x_10 = x_55; -x_11 = x_58; -x_12 = x_57; -x_13 = x_56; -x_14 = x_62; -x_15 = x_61; -x_16 = x_63; +x_10 = x_56; +x_11 = x_55; +x_12 = x_58; +x_13 = x_61; +x_14 = x_60; +x_15 = x_63; +x_16 = x_62; x_17 = x_64; -x_18 = x_60; +x_18 = x_57; goto block_54; } } @@ -11267,14 +11267,14 @@ x_100 = l_Lean_Meta_getDecLevel(x_99, x_93, x_94, x_95, x_96); if (lean_obj_tag(x_100) == 0) { lean_dec(x_2); -x_10 = x_93; -x_11 = x_95; -x_12 = x_91; -x_13 = x_99; -x_14 = x_94; -x_15 = x_92; -x_16 = x_90; -x_17 = x_96; +x_10 = x_96; +x_11 = x_99; +x_12 = x_94; +x_13 = x_91; +x_14 = x_90; +x_15 = x_95; +x_16 = x_92; +x_17 = x_93; x_18 = x_100; goto block_54; } @@ -11288,32 +11288,32 @@ if (x_102 == 0) { uint8_t x_103; x_103 = l_Lean_Exception_isRuntime(x_101); -x_55 = x_93; -x_56 = x_99; -x_57 = x_91; -x_58 = x_95; +x_55 = x_99; +x_56 = x_96; +x_57 = x_100; +x_58 = x_94; x_59 = lean_box(0); -x_60 = x_100; -x_61 = x_92; -x_62 = x_94; -x_63 = x_90; -x_64 = x_96; +x_60 = x_90; +x_61 = x_91; +x_62 = x_92; +x_63 = x_95; +x_64 = x_93; x_65 = x_103; goto block_89; } else { lean_dec(x_101); -x_55 = x_93; -x_56 = x_99; -x_57 = x_91; -x_58 = x_95; +x_55 = x_99; +x_56 = x_96; +x_57 = x_100; +x_58 = x_94; x_59 = lean_box(0); -x_60 = x_100; -x_61 = x_92; -x_62 = x_94; -x_63 = x_90; -x_64 = x_96; +x_60 = x_90; +x_61 = x_91; +x_62 = x_92; +x_63 = x_95; +x_64 = x_93; x_65 = x_102; goto block_89; } @@ -23575,7 +23575,7 @@ return x_9; LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logWarning___at___00Lean_checkPrivateInPublic___at___00Lean_resolveGlobalName___at___00__private_Lean_ResolveName_0__Lean_resolveGlobalConstCore___at___00Lean_resolveGlobalConstNoOverloadCore___at___00Lean_Elab_OpenDecl_resolveId___at___00Lean_Elab_OpenDecl_elabOpenDecl___at___00Lean_Elab_Term_elabOpen_spec__2_spec__9_spec__12_spec__15_spec__29_spec__48_spec__56_spec__60_spec__64___redArg(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_50; lean_object* x_51; uint8_t x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; uint8_t x_56; lean_object* x_57; lean_object* x_77; lean_object* x_78; uint8_t x_79; lean_object* x_80; lean_object* x_81; uint8_t x_82; uint8_t x_83; lean_object* x_84; lean_object* x_88; lean_object* x_89; lean_object* x_90; uint8_t x_91; lean_object* x_92; uint8_t x_93; uint8_t x_94; uint8_t x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; uint8_t x_105; uint8_t x_106; uint8_t x_107; uint8_t x_109; uint8_t x_125; +uint8_t x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_50; lean_object* x_51; uint8_t x_52; uint8_t x_53; lean_object* x_54; uint8_t x_55; lean_object* x_56; lean_object* x_57; lean_object* x_77; uint8_t x_78; lean_object* x_79; uint8_t x_80; lean_object* x_81; lean_object* x_82; uint8_t x_83; lean_object* x_84; lean_object* x_88; lean_object* x_89; uint8_t x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; uint8_t x_94; uint8_t x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; uint8_t x_105; uint8_t x_106; uint8_t x_107; uint8_t x_109; uint8_t x_125; x_100 = 2; x_125 = l_Lean_instBEqMessageSeverity_beq(x_3, x_100); if (x_125 == 0) @@ -23610,15 +23610,15 @@ lean_ctor_set(x_25, 0, x_21); lean_ctor_set(x_25, 1, x_22); x_26 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_13); +lean_ctor_set(x_26, 1, x_14); x_27 = lean_alloc_ctor(0, 5, 3); lean_ctor_set(x_27, 0, x_11); -lean_ctor_set(x_27, 1, x_14); -lean_ctor_set(x_27, 2, x_15); -lean_ctor_set(x_27, 3, x_10); +lean_ctor_set(x_27, 1, x_15); +lean_ctor_set(x_27, 2, x_13); +lean_ctor_set(x_27, 3, x_16); lean_ctor_set(x_27, 4, x_26); lean_ctor_set_uint8(x_27, sizeof(void*)*5, x_12); -lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 1, x_16); +lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 1, x_10); lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 2, x_4); x_28 = l_Lean_MessageLog_add(x_27, x_24); lean_ctor_set(x_20, 6, x_28); @@ -23655,15 +23655,15 @@ lean_ctor_set(x_41, 0, x_21); lean_ctor_set(x_41, 1, x_22); x_42 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_13); +lean_ctor_set(x_42, 1, x_14); x_43 = lean_alloc_ctor(0, 5, 3); lean_ctor_set(x_43, 0, x_11); -lean_ctor_set(x_43, 1, x_14); -lean_ctor_set(x_43, 2, x_15); -lean_ctor_set(x_43, 3, x_10); +lean_ctor_set(x_43, 1, x_15); +lean_ctor_set(x_43, 2, x_13); +lean_ctor_set(x_43, 3, x_16); lean_ctor_set(x_43, 4, x_42); lean_ctor_set_uint8(x_43, sizeof(void*)*5, x_12); -lean_ctor_set_uint8(x_43, sizeof(void*)*5 + 1, x_16); +lean_ctor_set_uint8(x_43, sizeof(void*)*5 + 1, x_10); lean_ctor_set_uint8(x_43, sizeof(void*)*5 + 2, x_4); x_44 = l_Lean_MessageLog_add(x_43, x_38); x_45 = lean_alloc_ctor(0, 9, 0); @@ -23694,8 +23694,8 @@ if (x_60 == 0) lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; x_61 = lean_ctor_get(x_59, 0); lean_inc_ref(x_54); -x_62 = l_Lean_FileMap_toPosition(x_54, x_53); -lean_dec(x_53); +x_62 = l_Lean_FileMap_toPosition(x_54, x_56); +lean_dec(x_56); x_63 = l_Lean_FileMap_toPosition(x_54, x_57); lean_dec(x_57); x_64 = lean_alloc_ctor(1, 1, 0); @@ -23705,13 +23705,13 @@ if (x_55 == 0) { lean_free_object(x_59); lean_dec_ref(x_50); -x_10 = x_65; +x_10 = x_52; x_11 = x_51; -x_12 = x_52; -x_13 = x_61; -x_14 = x_62; -x_15 = x_64; -x_16 = x_56; +x_12 = x_53; +x_13 = x_64; +x_14 = x_61; +x_15 = x_62; +x_16 = x_65; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -23737,13 +23737,13 @@ return x_59; else { lean_free_object(x_59); -x_10 = x_65; +x_10 = x_52; x_11 = x_51; -x_12 = x_52; -x_13 = x_61; -x_14 = x_62; -x_15 = x_64; -x_16 = x_56; +x_12 = x_53; +x_13 = x_64; +x_14 = x_61; +x_15 = x_62; +x_16 = x_65; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -23758,8 +23758,8 @@ x_68 = lean_ctor_get(x_59, 0); lean_inc(x_68); lean_dec(x_59); lean_inc_ref(x_54); -x_69 = l_Lean_FileMap_toPosition(x_54, x_53); -lean_dec(x_53); +x_69 = l_Lean_FileMap_toPosition(x_54, x_56); +lean_dec(x_56); x_70 = l_Lean_FileMap_toPosition(x_54, x_57); lean_dec(x_57); x_71 = lean_alloc_ctor(1, 1, 0); @@ -23768,13 +23768,13 @@ x_72 = l_Lean_logAt___at___00Lean_log___at___00Lean_logWarning___at___00Lean_Ela if (x_55 == 0) { lean_dec_ref(x_50); -x_10 = x_72; +x_10 = x_52; x_11 = x_51; -x_12 = x_52; -x_13 = x_68; -x_14 = x_69; -x_15 = x_71; -x_16 = x_56; +x_12 = x_53; +x_13 = x_71; +x_14 = x_68; +x_15 = x_69; +x_16 = x_72; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -23800,13 +23800,13 @@ return x_75; } else { -x_10 = x_72; +x_10 = x_52; x_11 = x_51; -x_12 = x_52; -x_13 = x_68; -x_14 = x_69; -x_15 = x_71; -x_16 = x_56; +x_12 = x_53; +x_13 = x_71; +x_14 = x_68; +x_15 = x_69; +x_16 = x_72; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -23818,18 +23818,18 @@ goto block_49; block_87: { lean_object* x_85; -x_85 = l_Lean_Syntax_getTailPos_x3f(x_81, x_79); -lean_dec(x_81); +x_85 = l_Lean_Syntax_getTailPos_x3f(x_82, x_80); +lean_dec(x_82); if (lean_obj_tag(x_85) == 0) { lean_inc(x_84); x_50 = x_77; -x_51 = x_78; -x_52 = x_79; -x_53 = x_84; -x_54 = x_80; +x_51 = x_79; +x_52 = x_78; +x_53 = x_80; +x_54 = x_81; x_55 = x_83; -x_56 = x_82; +x_56 = x_84; x_57 = x_84; goto block_76; } @@ -23840,12 +23840,12 @@ x_86 = lean_ctor_get(x_85, 0); lean_inc(x_86); lean_dec_ref(x_85); x_50 = x_77; -x_51 = x_78; -x_52 = x_79; -x_53 = x_84; -x_54 = x_80; +x_51 = x_79; +x_52 = x_78; +x_53 = x_80; +x_54 = x_81; x_55 = x_83; -x_56 = x_82; +x_56 = x_84; x_57 = x_86; goto block_76; } @@ -23853,19 +23853,19 @@ goto block_76; block_99: { lean_object* x_95; lean_object* x_96; -x_95 = l_Lean_replaceRef(x_1, x_90); -lean_dec(x_90); -x_96 = l_Lean_Syntax_getPos_x3f(x_95, x_91); +x_95 = l_Lean_replaceRef(x_1, x_92); +lean_dec(x_92); +x_96 = l_Lean_Syntax_getPos_x3f(x_95, x_90); if (lean_obj_tag(x_96) == 0) { lean_object* x_97; x_97 = lean_unsigned_to_nat(0u); x_77 = x_88; -x_78 = x_89; -x_79 = x_91; -x_80 = x_92; -x_81 = x_95; -x_82 = x_94; +x_78 = x_94; +x_79 = x_89; +x_80 = x_90; +x_81 = x_91; +x_82 = x_95; x_83 = x_93; x_84 = x_97; goto block_87; @@ -23877,11 +23877,11 @@ x_98 = lean_ctor_get(x_96, 0); lean_inc(x_98); lean_dec_ref(x_96); x_77 = x_88; -x_78 = x_89; -x_79 = x_91; -x_80 = x_92; -x_81 = x_95; -x_82 = x_94; +x_78 = x_94; +x_79 = x_89; +x_80 = x_90; +x_81 = x_91; +x_82 = x_95; x_83 = x_93; x_84 = x_98; goto block_87; @@ -23891,22 +23891,22 @@ goto block_87; { if (x_107 == 0) { -x_88 = x_103; +x_88 = x_104; x_89 = x_101; -x_90 = x_102; -x_91 = x_106; -x_92 = x_104; +x_90 = x_106; +x_91 = x_102; +x_92 = x_103; x_93 = x_105; x_94 = x_3; goto block_99; } else { -x_88 = x_103; +x_88 = x_104; x_89 = x_101; -x_90 = x_102; -x_91 = x_106; -x_92 = x_104; +x_90 = x_106; +x_91 = x_102; +x_92 = x_103; x_93 = x_105; x_94 = x_100; goto block_99; @@ -23931,13 +23931,13 @@ x_118 = 1; x_119 = l_Lean_instBEqMessageSeverity_beq(x_3, x_118); if (x_119 == 0) { -lean_inc_ref(x_111); lean_inc(x_113); +lean_inc_ref(x_111); lean_inc_ref(x_110); x_101 = x_110; -x_102 = x_113; -x_103 = x_117; -x_104 = x_111; +x_102 = x_111; +x_103 = x_113; +x_104 = x_117; x_105 = x_114; x_106 = x_109; x_107 = x_119; @@ -23948,13 +23948,13 @@ else lean_object* x_120; uint8_t x_121; x_120 = l_Lean_logAt___at___00Lean_log___at___00Lean_logWarning___at___00Lean_Elab_Term_elabOmission_spec__0_spec__0_spec__1___redArg___closed__1; x_121 = l_Lean_Option_get___at___00Lean_Elab_addMacroStack___at___00Lean_throwError___at___00Lean_throwErrorAt___at___00Lean_Elab_Term_elabPipeCompletion_spec__0_spec__0_spec__2_spec__3(x_112, x_120); -lean_inc_ref(x_111); lean_inc(x_113); +lean_inc_ref(x_111); lean_inc_ref(x_110); x_101 = x_110; -x_102 = x_113; -x_103 = x_117; -x_104 = x_111; +x_102 = x_111; +x_103 = x_113; +x_104 = x_117; x_105 = x_114; x_106 = x_109; x_107 = x_121; @@ -32397,13 +32397,13 @@ lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean x_16 = l_Lean_MessageData_ofFormat(x_15); x_17 = l_Lean_indentD(x_16); x_18 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_18, 0, x_14); +lean_ctor_set(x_18, 0, x_13); lean_ctor_set(x_18, 1, x_17); x_19 = l___private_Lean_Elab_SetOption_0__Lean_Elab_validateOptionValue_throwMistypedOptionValue___at___00Lean_Elab_validateOptionValue___at___00__private_Lean_Elab_SetOption_0__Lean_Elab_elabSetOption_setOption___at___00Lean_Elab_elabSetOption___at___00Lean_Elab_Term_elabSetOption_spec__0_spec__1_spec__3_spec__4___closed__1; x_20 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_20, 0, x_18); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_MessageData_ofExpr(x_13); +x_21 = l_Lean_MessageData_ofExpr(x_14); x_22 = l_Lean_indentD(x_21); x_23 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_23, 0, x_20); @@ -32443,8 +32443,8 @@ x_37 = lean_ctor_get(x_2, 0); x_38 = l_String_quote(x_37); lean_ctor_set_tag(x_2, 3); lean_ctor_set(x_2, 0, x_38); -x_13 = x_34; -x_14 = x_35; +x_13 = x_35; +x_14 = x_34; x_15 = x_2; goto block_33; } @@ -32457,8 +32457,8 @@ lean_dec(x_2); x_40 = l_String_quote(x_39); x_41 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_41, 0, x_40); -x_13 = x_34; -x_14 = x_35; +x_13 = x_35; +x_14 = x_34; x_15 = x_41; goto block_33; } @@ -32472,8 +32472,8 @@ if (x_42 == 0) { lean_object* x_43; x_43 = l___private_Lean_Elab_SetOption_0__Lean_Elab_validateOptionValue_throwMistypedOptionValue___at___00Lean_Elab_validateOptionValue___at___00__private_Lean_Elab_SetOption_0__Lean_Elab_elabSetOption_setOption___at___00Lean_Elab_elabSetOption___at___00Lean_Elab_Term_elabSetOption_spec__0_spec__1_spec__3_spec__4___closed__8; -x_13 = x_34; -x_14 = x_35; +x_13 = x_35; +x_14 = x_34; x_15 = x_43; goto block_33; } @@ -32481,8 +32481,8 @@ else { lean_object* x_44; x_44 = l___private_Lean_Elab_SetOption_0__Lean_Elab_validateOptionValue_throwMistypedOptionValue___at___00Lean_Elab_validateOptionValue___at___00__private_Lean_Elab_SetOption_0__Lean_Elab_elabSetOption_setOption___at___00Lean_Elab_elabSetOption___at___00Lean_Elab_Term_elabSetOption_spec__0_spec__1_spec__3_spec__4___closed__9; -x_13 = x_34; -x_14 = x_35; +x_13 = x_35; +x_14 = x_34; x_15 = x_44; goto block_33; } @@ -32503,8 +32503,8 @@ lean_ctor_set(x_2, 0, x_49); x_50 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_50, 0, x_47); lean_ctor_set(x_50, 1, x_2); -x_13 = x_34; -x_14 = x_35; +x_13 = x_35; +x_14 = x_34; x_15 = x_50; goto block_33; } @@ -32522,8 +32522,8 @@ lean_ctor_set(x_55, 0, x_54); x_56 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_56, 0, x_52); lean_ctor_set(x_56, 1, x_55); -x_13 = x_34; -x_14 = x_35; +x_13 = x_35; +x_14 = x_34; x_15 = x_56; goto block_33; } @@ -32538,8 +32538,8 @@ lean_object* x_58; lean_object* x_59; x_58 = lean_ctor_get(x_2, 0); x_59 = l_Nat_reprFast(x_58); lean_ctor_set(x_2, 0, x_59); -x_13 = x_34; -x_14 = x_35; +x_13 = x_35; +x_14 = x_34; x_15 = x_2; goto block_33; } @@ -32552,8 +32552,8 @@ lean_dec(x_2); x_61 = l_Nat_reprFast(x_60); x_62 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_62, 0, x_61); -x_13 = x_34; -x_14 = x_35; +x_13 = x_35; +x_14 = x_34; x_15 = x_62; goto block_33; } @@ -32576,8 +32576,8 @@ lean_dec(x_64); x_68 = l_Nat_reprFast(x_67); lean_ctor_set_tag(x_2, 3); lean_ctor_set(x_2, 0, x_68); -x_13 = x_34; -x_14 = x_35; +x_13 = x_35; +x_14 = x_34; x_15 = x_2; goto block_33; } @@ -32597,8 +32597,8 @@ x_75 = lean_string_append(x_72, x_74); lean_dec_ref(x_74); lean_ctor_set_tag(x_2, 3); lean_ctor_set(x_2, 0, x_75); -x_13 = x_34; -x_14 = x_35; +x_13 = x_35; +x_14 = x_34; x_15 = x_2; goto block_33; } @@ -32619,8 +32619,8 @@ lean_dec(x_76); x_80 = l_Nat_reprFast(x_79); x_81 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_81, 0, x_80); -x_13 = x_34; -x_14 = x_35; +x_13 = x_35; +x_14 = x_34; x_15 = x_81; goto block_33; } @@ -32640,8 +32640,8 @@ x_88 = lean_string_append(x_85, x_87); lean_dec_ref(x_87); x_89 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_89, 0, x_88); -x_13 = x_34; -x_14 = x_35; +x_13 = x_35; +x_14 = x_34; x_15 = x_89; goto block_33; } @@ -32656,8 +32656,8 @@ lean_dec_ref(x_2); x_91 = lean_box(0); x_92 = 0; x_93 = l_Lean_Syntax_formatStx(x_90, x_91, x_92); -x_13 = x_34; -x_14 = x_35; +x_13 = x_35; +x_14 = x_34; x_15 = x_93; goto block_33; } diff --git a/stage0/stdlib/Lean/Elab/MutualDef.c b/stage0/stdlib/Lean/Elab/MutualDef.c index 3fedd0af7240..a248f6ff6718 100644 --- a/stage0/stdlib/Lean/Elab/MutualDef.c +++ b/stage0/stdlib/Lean/Elab/MutualDef.c @@ -61247,8 +61247,8 @@ lean_inc(x_11); lean_inc_ref(x_10); lean_inc(x_9); lean_inc_ref(x_8); -x_34 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withUsed___redArg(x_5, x_31, x_33, x_32, x_30, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec_ref(x_33); +x_34 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withUsed___redArg(x_5, x_31, x_32, x_30, x_29, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec_ref(x_32); lean_dec_ref(x_31); lean_dec_ref(x_5); x_15 = x_34; @@ -61263,15 +61263,15 @@ lean_dec(x_40); if (x_43 == 0) { lean_object* x_44; -lean_dec_ref(x_41); -lean_dec(x_39); +lean_dec_ref(x_39); +lean_dec(x_37); lean_inc(x_13); lean_inc_ref(x_12); lean_inc(x_11); lean_inc_ref(x_10); lean_inc(x_9); lean_inc_ref(x_8); -x_44 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars___redArg(x_5, x_6, x_38, x_37, x_43, x_8, x_9, x_10, x_11, x_12, x_13); +x_44 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars___redArg(x_5, x_6, x_38, x_36, x_43, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec_ref(x_38); lean_dec_ref(x_6); lean_dec_ref(x_5); @@ -61281,11 +61281,11 @@ goto block_28; else { lean_dec_ref(x_6); -x_29 = lean_box(0); +x_29 = x_36; x_30 = x_37; x_31 = x_38; x_32 = x_39; -x_33 = x_41; +x_33 = lean_box(0); goto block_35; } } @@ -61395,12 +61395,12 @@ x_78 = lean_nat_dec_lt(x_2, x_57); if (x_78 == 0) { lean_inc(x_74); -x_36 = lean_box(0); -x_37 = x_77; +x_36 = x_77; +x_37 = x_55; x_38 = x_50; -x_39 = x_55; +x_39 = x_46; x_40 = x_74; -x_41 = x_46; +x_41 = lean_box(0); goto block_45; } else @@ -61408,12 +61408,12 @@ else if (x_78 == 0) { lean_inc(x_74); -x_36 = lean_box(0); -x_37 = x_77; +x_36 = x_77; +x_37 = x_55; x_38 = x_50; -x_39 = x_55; +x_39 = x_46; x_40 = x_74; -x_41 = x_46; +x_41 = lean_box(0); goto block_45; } else @@ -61424,22 +61424,22 @@ x_80 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00__priv if (x_80 == 0) { lean_inc(x_74); -x_36 = lean_box(0); -x_37 = x_77; +x_36 = x_77; +x_37 = x_55; x_38 = x_50; -x_39 = x_55; +x_39 = x_46; x_40 = x_74; -x_41 = x_46; +x_41 = lean_box(0); goto block_45; } else { lean_dec_ref(x_6); -x_29 = lean_box(0); -x_30 = x_77; +x_29 = x_77; +x_30 = x_55; x_31 = x_50; -x_32 = x_55; -x_33 = x_46; +x_32 = x_46; +x_33 = lean_box(0); goto block_35; } } @@ -61630,7 +61630,7 @@ if (x_99 == 0) { lean_object* x_100; lean_inc_ref(x_12); -x_100 = l_Lean_Elab_logException___at___00Lean_Elab_withLogging___at___00__private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabMutualDef_processDeriving_spec__2_spec__3(x_98, x_8, x_9, x_10, x_11, x_12, x_13); +x_100 = l_Lean_Elab_logException___at___00Lean_Elab_withLogging___at___00__private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabMutualDef_processDeriving_spec__2_spec__3(x_97, x_8, x_9, x_10, x_11, x_12, x_13); if (lean_obj_tag(x_100) == 0) { size_t x_101; lean_object* x_102; @@ -61713,7 +61713,7 @@ lean_dec_ref(x_5); lean_dec(x_2); lean_dec_ref(x_1); x_107 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_107, 0, x_98); +lean_ctor_set(x_107, 0, x_97); return x_107; } } @@ -61726,15 +61726,15 @@ if (x_111 == 0) uint8_t x_112; lean_inc_ref(x_109); x_112 = l_Lean_Exception_isRuntime(x_109); -x_97 = lean_box(0); -x_98 = x_109; +x_97 = x_109; +x_98 = lean_box(0); x_99 = x_112; goto block_108; } else { -x_97 = lean_box(0); -x_98 = x_109; +x_97 = x_109; +x_98 = lean_box(0); x_99 = x_111; goto block_108; } @@ -61802,7 +61802,7 @@ return x_1; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabMutualDef_finishElab___lam__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; uint8_t x_52; lean_object* x_73; uint8_t x_79; +uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; lean_object* x_49; uint8_t x_50; lean_object* x_51; uint8_t x_52; lean_object* x_73; uint8_t x_79; x_35 = lean_st_ref_get(x_11); x_36 = lean_ctor_get(x_10, 0); lean_inc_ref(x_36); @@ -61870,11 +61870,11 @@ goto block_78; { lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; x_30 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabMutualDef_finishElab___lam__3___closed__0; -x_31 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00__private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_spec__5_spec__10(x_13, x_30); +x_31 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00__private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_spec__5_spec__10(x_14, x_30); x_32 = lean_alloc_ctor(0, 14, 2); lean_ctor_set(x_32, 0, x_15); lean_ctor_set(x_32, 1, x_16); -lean_ctor_set(x_32, 2, x_13); +lean_ctor_set(x_32, 2, x_14); lean_ctor_set(x_32, 3, x_17); lean_ctor_set(x_32, 4, x_31); lean_ctor_set(x_32, 5, x_18); @@ -61886,7 +61886,7 @@ lean_ctor_set(x_32, 10, x_23); lean_ctor_set(x_32, 11, x_24); lean_ctor_set(x_32, 12, x_25); lean_ctor_set(x_32, 13, x_27); -lean_ctor_set_uint8(x_32, sizeof(void*)*14, x_14); +lean_ctor_set_uint8(x_32, sizeof(void*)*14, x_13); lean_ctor_set_uint8(x_32, sizeof(void*)*14 + 1, x_26); x_33 = l_Lean_Elab_withSaveInfoContext___at___00__private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabMutualDef_finishElab_spec__14___redArg(x_1, x_6, x_7, x_8, x_9, x_32, x_28); return x_33; @@ -61904,7 +61904,7 @@ lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean x_55 = lean_ctor_get(x_53, 0); x_56 = lean_ctor_get(x_53, 5); lean_dec(x_56); -x_57 = l_Lean_Kernel_enableDiag(x_55, x_51); +x_57 = l_Lean_Kernel_enableDiag(x_55, x_50); x_58 = l___private_Lean_ExtraModUses_0__Lean_recordExtraModUseCore___at___00Lean_recordExtraModUseFromDecl___at___00Lean_Elab_liftMacroM___at___00__private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerm_spec__1_spec__2_spec__3___redArg___closed__6; lean_ctor_set(x_53, 5, x_58); lean_ctor_set(x_53, 0, x_57); @@ -61948,7 +61948,7 @@ lean_inc(x_62); lean_inc(x_61); lean_inc(x_60); lean_dec(x_53); -x_68 = l_Lean_Kernel_enableDiag(x_60, x_51); +x_68 = l_Lean_Kernel_enableDiag(x_60, x_50); x_69 = l___private_Lean_ExtraModUses_0__Lean_recordExtraModUseCore___at___00Lean_recordExtraModUseFromDecl___at___00Lean_Elab_liftMacroM___at___00__private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerm_spec__1_spec__2_spec__3___redArg___closed__6; x_70 = lean_alloc_ctor(0, 9, 0); lean_ctor_set(x_70, 0, x_68); @@ -62017,8 +62017,8 @@ if (x_77 == 0) { if (x_76 == 0) { -x_13 = x_73; -x_14 = x_76; +x_13 = x_76; +x_14 = x_73; x_15 = x_36; x_16 = x_37; x_17 = x_39; @@ -62038,16 +62038,16 @@ goto block_34; } else { -x_50 = x_73; -x_51 = x_76; +x_50 = x_76; +x_51 = x_73; x_52 = x_77; goto block_72; } } else { -x_50 = x_73; -x_51 = x_76; +x_50 = x_76; +x_51 = x_73; x_52 = x_76; goto block_72; } @@ -62308,7 +62308,7 @@ uint8_t x_12; x_12 = lean_usize_dec_eq(x_3, x_4); if (x_12 == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; uint8_t x_18; lean_object* x_19; uint8_t x_26; lean_object* x_27; lean_object* x_31; lean_object* x_35; uint8_t x_36; uint8_t x_37; uint8_t x_69; lean_object* x_85; lean_object* x_86; uint8_t x_87; +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; uint8_t x_18; lean_object* x_19; uint8_t x_26; lean_object* x_27; lean_object* x_31; uint8_t x_35; lean_object* x_36; uint8_t x_37; uint8_t x_69; lean_object* x_85; lean_object* x_86; uint8_t x_87; x_13 = lean_array_uget(x_2, x_3); x_14 = lean_ctor_get(x_13, 0); lean_inc_ref(x_14); @@ -62436,7 +62436,7 @@ goto block_34; } else { -x_26 = x_36; +x_26 = x_35; x_27 = lean_box(0); goto block_30; } @@ -62488,15 +62488,15 @@ x_53 = lean_unbox(x_52); lean_dec(x_52); if (x_53 == 0) { -x_35 = lean_box(0); -x_36 = x_45; +x_35 = x_45; +x_36 = lean_box(0); x_37 = x_47; goto block_40; } else { -x_35 = lean_box(0); -x_36 = x_45; +x_35 = x_45; +x_36 = lean_box(0); x_37 = x_45; goto block_40; } @@ -62910,8 +62910,8 @@ size_t x_26; size_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_3 x_26 = lean_array_size(x_1); x_27 = 0; lean_inc_ref(x_1); -x_28 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabMutualDef_finishElab_spec__4(x_18, x_2, x_26, x_27, x_1); -lean_dec(x_18); +x_28 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabMutualDef_finishElab_spec__4(x_22, x_2, x_26, x_27, x_1); +lean_dec(x_22); x_29 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabMutualDef_finishElab___lam__4___boxed__const__1; lean_inc(x_16); x_30 = lean_alloc_closure((void*)(l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabMutualDef_finishElab___lam__2___boxed), 14, 7); @@ -62929,7 +62929,7 @@ lean_closure_set(x_32, 1, x_16); lean_closure_set(x_32, 2, x_15); lean_closure_set(x_32, 3, x_1); lean_closure_set(x_32, 4, x_31); -x_33 = l_Lean_withoutExporting___at___00Lean_Meta_abstractProof___at___00__private_Lean_Elab_MutualDef_0__Lean_Elab_Term_fillHolesFromWhereFinally_spec__4_spec__5___redArg(x_32, x_24, x_20, x_23, x_17, x_19, x_21, x_22); +x_33 = l_Lean_withoutExporting___at___00Lean_Meta_abstractProof___at___00__private_Lean_Elab_MutualDef_0__Lean_Elab_Term_fillHolesFromWhereFinally_spec__4_spec__5___redArg(x_32, x_24, x_18, x_19, x_21, x_23, x_20, x_17); return x_33; } block_49: @@ -62942,16 +62942,16 @@ lean_inc(x_44); lean_dec_ref(x_43); x_45 = lean_unbox(x_44); lean_dec(x_44); -lean_inc(x_42); -x_15 = x_37; -x_16 = x_42; -x_17 = x_36; -x_18 = x_42; -x_19 = x_41; +lean_inc(x_41); +x_15 = x_35; +x_16 = x_41; +x_17 = x_42; +x_18 = x_37; +x_19 = x_40; x_20 = x_39; -x_21 = x_35; -x_22 = x_38; -x_23 = x_40; +x_21 = x_36; +x_22 = x_41; +x_23 = x_38; x_24 = x_45; x_25 = lean_box(0); goto block_34; @@ -62964,9 +62964,9 @@ lean_dec(x_41); lean_dec(x_40); lean_dec_ref(x_39); lean_dec(x_38); -lean_dec(x_37); +lean_dec_ref(x_37); lean_dec_ref(x_36); -lean_dec_ref(x_35); +lean_dec(x_35); lean_dec_ref(x_7); lean_dec_ref(x_5); lean_dec_ref(x_4); @@ -63007,14 +63007,14 @@ lean_inc_ref(x_52); lean_inc(x_51); lean_inc_ref(x_50); x_61 = lean_apply_8(x_6, x_60, x_50, x_51, x_52, x_53, x_54, x_55, lean_box(0)); -x_35 = x_54; +x_35 = x_58; x_36 = x_52; -x_37 = x_58; -x_38 = x_55; -x_39 = x_50; +x_37 = x_50; +x_38 = x_53; +x_39 = x_54; x_40 = x_51; -x_41 = x_53; -x_42 = x_57; +x_41 = x_57; +x_42 = x_55; x_43 = x_61; goto block_49; } @@ -63025,13 +63025,13 @@ if (x_59 == 0) lean_dec_ref(x_6); x_15 = x_58; x_16 = x_57; -x_17 = x_52; -x_18 = x_57; -x_19 = x_53; -x_20 = x_50; -x_21 = x_54; -x_22 = x_55; -x_23 = x_51; +x_17 = x_55; +x_18 = x_50; +x_19 = x_51; +x_20 = x_54; +x_21 = x_52; +x_22 = x_57; +x_23 = x_53; x_24 = x_59; x_25 = lean_box(0); goto block_34; @@ -63060,28 +63060,28 @@ lean_inc_ref(x_52); lean_inc(x_51); lean_inc_ref(x_50); x_66 = lean_apply_8(x_6, x_65, x_50, x_51, x_52, x_53, x_54, x_55, lean_box(0)); -x_35 = x_54; +x_35 = x_58; x_36 = x_52; -x_37 = x_58; -x_38 = x_55; -x_39 = x_50; +x_37 = x_50; +x_38 = x_53; +x_39 = x_54; x_40 = x_51; -x_41 = x_53; -x_42 = x_57; +x_41 = x_57; +x_42 = x_55; x_43 = x_66; goto block_49; } else { lean_dec_ref(x_6); -x_35 = x_54; +x_35 = x_58; x_36 = x_52; -x_37 = x_58; -x_38 = x_55; -x_39 = x_50; +x_37 = x_50; +x_38 = x_53; +x_39 = x_54; x_40 = x_51; -x_41 = x_53; -x_42 = x_57; +x_41 = x_57; +x_42 = x_55; x_43 = x_64; goto block_49; } @@ -75004,7 +75004,7 @@ return x_7; LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_logErrorAt___at___00Lean_Elab_logException___at___00Lean_Elab_elabAttrs___at___00Lean_Elab_elabDeclAttrs___at___00Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabMutualDef_spec__4_spec__5_spec__7_spec__10_spec__13_spec__15(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_8; uint8_t x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_82; uint8_t x_83; uint8_t x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; uint8_t x_116; lean_object* x_117; uint8_t x_118; uint8_t x_119; lean_object* x_120; lean_object* x_121; uint8_t x_125; uint8_t x_126; lean_object* x_127; uint8_t x_128; uint8_t x_139; uint8_t x_140; uint8_t x_141; lean_object* x_142; uint8_t x_143; uint8_t x_145; uint8_t x_158; +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; uint8_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_82; lean_object* x_83; lean_object* x_84; uint8_t x_85; uint8_t x_86; lean_object* x_87; uint8_t x_116; lean_object* x_117; lean_object* x_118; uint8_t x_119; uint8_t x_120; lean_object* x_121; uint8_t x_125; lean_object* x_126; uint8_t x_127; uint8_t x_128; uint8_t x_139; uint8_t x_140; lean_object* x_141; uint8_t x_142; uint8_t x_143; uint8_t x_145; uint8_t x_158; x_139 = 2; x_158 = l_Lean_instBEqMessageSeverity_beq(x_3, x_139); if (x_158 == 0) @@ -75056,15 +75056,15 @@ lean_ctor_set(x_27, 0, x_23); lean_ctor_set(x_27, 1, x_24); x_28 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_28, 1, x_14); +lean_ctor_set(x_28, 1, x_9); x_29 = lean_alloc_ctor(0, 5, 3); -lean_ctor_set(x_29, 0, x_13); -lean_ctor_set(x_29, 1, x_11); -lean_ctor_set(x_29, 2, x_12); -lean_ctor_set(x_29, 3, x_8); +lean_ctor_set(x_29, 0, x_10); +lean_ctor_set(x_29, 1, x_8); +lean_ctor_set(x_29, 2, x_14); +lean_ctor_set(x_29, 3, x_13); lean_ctor_set(x_29, 4, x_28); -lean_ctor_set_uint8(x_29, sizeof(void*)*5, x_9); -lean_ctor_set_uint8(x_29, sizeof(void*)*5 + 1, x_10); +lean_ctor_set_uint8(x_29, sizeof(void*)*5, x_11); +lean_ctor_set_uint8(x_29, sizeof(void*)*5 + 1, x_12); lean_ctor_set_uint8(x_29, sizeof(void*)*5 + 2, x_4); x_30 = l_Lean_MessageLog_add(x_29, x_26); lean_ctor_set(x_22, 1, x_30); @@ -75104,15 +75104,15 @@ lean_ctor_set(x_44, 0, x_23); lean_ctor_set(x_44, 1, x_24); x_45 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_45, 0, x_44); -lean_ctor_set(x_45, 1, x_14); +lean_ctor_set(x_45, 1, x_9); x_46 = lean_alloc_ctor(0, 5, 3); -lean_ctor_set(x_46, 0, x_13); -lean_ctor_set(x_46, 1, x_11); -lean_ctor_set(x_46, 2, x_12); -lean_ctor_set(x_46, 3, x_8); +lean_ctor_set(x_46, 0, x_10); +lean_ctor_set(x_46, 1, x_8); +lean_ctor_set(x_46, 2, x_14); +lean_ctor_set(x_46, 3, x_13); lean_ctor_set(x_46, 4, x_45); -lean_ctor_set_uint8(x_46, sizeof(void*)*5, x_9); -lean_ctor_set_uint8(x_46, sizeof(void*)*5 + 1, x_10); +lean_ctor_set_uint8(x_46, sizeof(void*)*5, x_11); +lean_ctor_set_uint8(x_46, sizeof(void*)*5 + 1, x_12); lean_ctor_set_uint8(x_46, sizeof(void*)*5 + 2, x_4); x_47 = l_Lean_MessageLog_add(x_46, x_34); x_48 = lean_alloc_ctor(0, 11, 0); @@ -75190,15 +75190,15 @@ lean_ctor_set(x_67, 0, x_53); lean_ctor_set(x_67, 1, x_54); x_68 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_68, 0, x_67); -lean_ctor_set(x_68, 1, x_14); +lean_ctor_set(x_68, 1, x_9); x_69 = lean_alloc_ctor(0, 5, 3); -lean_ctor_set(x_69, 0, x_13); -lean_ctor_set(x_69, 1, x_11); -lean_ctor_set(x_69, 2, x_12); -lean_ctor_set(x_69, 3, x_8); +lean_ctor_set(x_69, 0, x_10); +lean_ctor_set(x_69, 1, x_8); +lean_ctor_set(x_69, 2, x_14); +lean_ctor_set(x_69, 3, x_13); lean_ctor_set(x_69, 4, x_68); -lean_ctor_set_uint8(x_69, sizeof(void*)*5, x_9); -lean_ctor_set_uint8(x_69, sizeof(void*)*5 + 1, x_10); +lean_ctor_set_uint8(x_69, sizeof(void*)*5, x_11); +lean_ctor_set_uint8(x_69, sizeof(void*)*5 + 1, x_12); lean_ctor_set_uint8(x_69, sizeof(void*)*5 + 2, x_4); x_70 = l_Lean_MessageLog_add(x_69, x_56); if (lean_is_scalar(x_66)) { @@ -75228,10 +75228,10 @@ else { uint8_t x_75; lean_dec(x_18); -lean_dec_ref(x_14); +lean_dec(x_14); lean_dec_ref(x_13); -lean_dec(x_12); -lean_dec_ref(x_11); +lean_dec_ref(x_10); +lean_dec_ref(x_9); lean_dec_ref(x_8); x_75 = !lean_is_exclusive(x_19); if (x_75 == 0) @@ -75253,10 +75253,10 @@ return x_77; else { uint8_t x_78; -lean_dec_ref(x_14); +lean_dec(x_14); lean_dec_ref(x_13); -lean_dec(x_12); -lean_dec_ref(x_11); +lean_dec_ref(x_10); +lean_dec_ref(x_9); lean_dec_ref(x_8); x_78 = !lean_is_exclusive(x_17); if (x_78 == 0) @@ -75292,8 +75292,8 @@ if (x_93 == 0) lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; x_94 = lean_ctor_get(x_92, 0); lean_inc_ref(x_89); -x_95 = l_Lean_FileMap_toPosition(x_89, x_85); -lean_dec(x_85); +x_95 = l_Lean_FileMap_toPosition(x_89, x_83); +lean_dec(x_83); x_96 = l_Lean_FileMap_toPosition(x_89, x_87); lean_dec(x_87); x_97 = lean_alloc_ctor(1, 1, 0); @@ -75302,13 +75302,13 @@ x_98 = l___private_Lean_Elab_ErrorUtils_0__List_toOxford___at___00__private_Lean if (x_90 == 0) { lean_free_object(x_92); -x_8 = x_98; -x_9 = x_83; -x_10 = x_84; -x_11 = x_95; -x_12 = x_97; -x_13 = x_88; -x_14 = x_94; +x_8 = x_95; +x_9 = x_94; +x_10 = x_88; +x_11 = x_85; +x_12 = x_86; +x_13 = x_98; +x_14 = x_97; x_15 = x_6; x_16 = lean_box(0); goto block_81; @@ -75337,13 +75337,13 @@ return x_92; else { lean_free_object(x_92); -x_8 = x_98; -x_9 = x_83; -x_10 = x_84; -x_11 = x_95; -x_12 = x_97; -x_13 = x_88; -x_14 = x_94; +x_8 = x_95; +x_9 = x_94; +x_10 = x_88; +x_11 = x_85; +x_12 = x_86; +x_13 = x_98; +x_14 = x_97; x_15 = x_6; x_16 = lean_box(0); goto block_81; @@ -75357,8 +75357,8 @@ x_104 = lean_ctor_get(x_92, 0); lean_inc(x_104); lean_dec(x_92); lean_inc_ref(x_89); -x_105 = l_Lean_FileMap_toPosition(x_89, x_85); -lean_dec(x_85); +x_105 = l_Lean_FileMap_toPosition(x_89, x_83); +lean_dec(x_83); x_106 = l_Lean_FileMap_toPosition(x_89, x_87); lean_dec(x_87); x_107 = lean_alloc_ctor(1, 1, 0); @@ -75366,13 +75366,13 @@ lean_ctor_set(x_107, 0, x_106); x_108 = l___private_Lean_Elab_ErrorUtils_0__List_toOxford___at___00__private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getPendingMVarErrorMessage_interpretedAsParameters_spec__1___closed__0; if (x_90 == 0) { -x_8 = x_108; -x_9 = x_83; -x_10 = x_84; -x_11 = x_105; -x_12 = x_107; -x_13 = x_88; -x_14 = x_104; +x_8 = x_105; +x_9 = x_104; +x_10 = x_88; +x_11 = x_85; +x_12 = x_86; +x_13 = x_108; +x_14 = x_107; x_15 = x_6; x_16 = lean_box(0); goto block_81; @@ -75401,13 +75401,13 @@ return x_114; } else { -x_8 = x_108; -x_9 = x_83; -x_10 = x_84; -x_11 = x_105; -x_12 = x_107; -x_13 = x_88; -x_14 = x_104; +x_8 = x_105; +x_9 = x_104; +x_10 = x_88; +x_11 = x_85; +x_12 = x_86; +x_13 = x_108; +x_14 = x_107; x_15 = x_6; x_16 = lean_box(0); goto block_81; @@ -75418,16 +75418,16 @@ goto block_81; block_124: { lean_object* x_122; -x_122 = l_Lean_Syntax_getTailPos_x3f(x_117, x_118); +x_122 = l_Lean_Syntax_getTailPos_x3f(x_117, x_119); lean_dec(x_117); if (lean_obj_tag(x_122) == 0) { lean_inc(x_121); x_82 = x_116; -x_83 = x_118; -x_84 = x_119; -x_85 = x_121; -x_86 = lean_box(0); +x_83 = x_121; +x_84 = lean_box(0); +x_85 = x_119; +x_86 = x_120; x_87 = x_121; goto block_115; } @@ -75438,10 +75438,10 @@ x_123 = lean_ctor_get(x_122, 0); lean_inc(x_123); lean_dec_ref(x_122); x_82 = x_116; -x_83 = x_118; -x_84 = x_119; -x_85 = x_121; -x_86 = lean_box(0); +x_83 = x_121; +x_84 = lean_box(0); +x_85 = x_119; +x_86 = x_120; x_87 = x_123; goto block_115; } @@ -75458,16 +75458,16 @@ lean_inc(x_130); lean_dec_ref(x_129); x_131 = l_Lean_replaceRef(x_1, x_130); lean_dec(x_130); -x_132 = l_Lean_Syntax_getPos_x3f(x_131, x_126); +x_132 = l_Lean_Syntax_getPos_x3f(x_131, x_127); if (lean_obj_tag(x_132) == 0) { lean_object* x_133; x_133 = lean_unsigned_to_nat(0u); x_116 = x_125; x_117 = x_131; -x_118 = x_126; -x_119 = x_128; -x_120 = lean_box(0); +x_118 = lean_box(0); +x_119 = x_127; +x_120 = x_128; x_121 = x_133; goto block_124; } @@ -75479,9 +75479,9 @@ lean_inc(x_134); lean_dec_ref(x_132); x_116 = x_125; x_117 = x_131; -x_118 = x_126; -x_119 = x_128; -x_120 = lean_box(0); +x_118 = lean_box(0); +x_119 = x_127; +x_120 = x_128; x_121 = x_134; goto block_124; } @@ -75513,16 +75513,16 @@ return x_137; if (x_143 == 0) { x_125 = x_140; -x_126 = x_141; -x_127 = lean_box(0); +x_126 = lean_box(0); +x_127 = x_142; x_128 = x_3; goto block_138; } else { x_125 = x_140; -x_126 = x_141; -x_127 = lean_box(0); +x_126 = lean_box(0); +x_127 = x_142; x_128 = x_139; goto block_138; } @@ -75548,8 +75548,8 @@ if (x_152 == 0) { lean_dec(x_150); x_140 = x_145; -x_141 = x_145; -x_142 = lean_box(0); +x_141 = lean_box(0); +x_142 = x_145; x_143 = x_152; goto block_144; } @@ -75560,8 +75560,8 @@ x_153 = l_Lean_logAt___at___00Lean_logWarningAt___at___00Lean_Linter_logLint___a x_154 = l_Lean_Option_get___at___00Lean_Elab_addMacroStack___at___00Lean_throwError___at___00__private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkModifiers_spec__0_spec__1_spec__2(x_150, x_153); lean_dec(x_150); x_140 = x_145; -x_141 = x_145; -x_142 = lean_box(0); +x_141 = lean_box(0); +x_142 = x_145; x_143 = x_154; goto block_144; } @@ -76406,7 +76406,7 @@ return x_5; LEAN_EXPORT lean_object* l_Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabMutualDef_spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_5; lean_object* x_6; uint8_t x_7; uint8_t x_8; lean_object* x_9; uint8_t x_10; uint8_t x_11; uint8_t x_12; uint8_t x_16; lean_object* x_17; uint8_t x_18; uint8_t x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; uint8_t x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; uint8_t x_40; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; uint8_t x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_66; uint8_t x_67; lean_object* x_68; uint8_t x_69; lean_object* x_70; lean_object* x_86; uint8_t x_87; uint8_t x_88; uint8_t x_109; lean_object* x_123; lean_object* x_124; uint8_t x_125; +lean_object* x_5; lean_object* x_6; uint8_t x_7; uint8_t x_8; uint8_t x_9; lean_object* x_10; uint8_t x_11; uint8_t x_12; lean_object* x_16; lean_object* x_17; uint8_t x_18; uint8_t x_19; uint8_t x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; uint8_t x_35; uint8_t x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; uint8_t x_57; uint8_t x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_66; lean_object* x_67; uint8_t x_68; uint8_t x_69; lean_object* x_70; lean_object* x_86; uint8_t x_87; uint8_t x_88; uint8_t x_109; lean_object* x_123; lean_object* x_124; uint8_t x_125; x_28 = lean_unsigned_to_nat(0u); x_29 = l_Lean_Syntax_getArg(x_1, x_28); x_30 = lean_unsigned_to_nat(1u); @@ -76455,11 +76455,11 @@ goto block_122; lean_object* x_13; lean_object* x_14; x_13 = lean_alloc_ctor(0, 3, 5); lean_ctor_set(x_13, 0, x_1); -lean_ctor_set(x_13, 1, x_9); -lean_ctor_set(x_13, 2, x_5); -lean_ctor_set_uint8(x_13, sizeof(void*)*3, x_8); -lean_ctor_set_uint8(x_13, sizeof(void*)*3 + 1, x_10); -lean_ctor_set_uint8(x_13, sizeof(void*)*3 + 2, x_11); +lean_ctor_set(x_13, 1, x_6); +lean_ctor_set(x_13, 2, x_10); +lean_ctor_set_uint8(x_13, sizeof(void*)*3, x_9); +lean_ctor_set_uint8(x_13, sizeof(void*)*3 + 1, x_11); +lean_ctor_set_uint8(x_13, sizeof(void*)*3 + 2, x_8); lean_ctor_set_uint8(x_13, sizeof(void*)*3 + 3, x_7); lean_ctor_set_uint8(x_13, sizeof(void*)*3 + 4, x_12); x_14 = lean_alloc_ctor(0, 1, 0); @@ -76469,18 +76469,18 @@ return x_14; block_27: { uint8_t x_24; -x_24 = l_Lean_Syntax_isNone(x_20); -lean_dec(x_20); +x_24 = l_Lean_Syntax_isNone(x_16); +lean_dec(x_16); if (x_24 == 0) { uint8_t x_25; x_25 = 1; -x_5 = x_22; -x_6 = lean_box(0); -x_7 = x_16; -x_8 = x_18; -x_9 = x_17; -x_10 = x_19; +x_5 = lean_box(0); +x_6 = x_17; +x_7 = x_18; +x_8 = x_20; +x_9 = x_19; +x_10 = x_22; x_11 = x_21; x_12 = x_25; goto block_15; @@ -76489,12 +76489,12 @@ else { uint8_t x_26; x_26 = 0; -x_5 = x_22; -x_6 = lean_box(0); -x_7 = x_16; -x_8 = x_18; -x_9 = x_17; -x_10 = x_19; +x_5 = lean_box(0); +x_6 = x_17; +x_7 = x_18; +x_8 = x_20; +x_9 = x_19; +x_10 = x_22; x_11 = x_21; x_12 = x_26; goto block_15; @@ -76508,15 +76508,15 @@ lean_dec(x_31); if (lean_obj_tag(x_41) == 0) { lean_object* x_42; -lean_dec(x_34); -lean_dec_ref(x_32); +lean_dec(x_38); +lean_dec_ref(x_37); x_42 = l_Lean_Elab_Term_MutualClosure_getModifiersForLetRecs___closed__0; -x_16 = x_35; -x_17 = x_37; -x_18 = x_36; -x_19 = x_40; -x_20 = x_38; -x_21 = x_39; +x_16 = x_33; +x_17 = x_32; +x_18 = x_34; +x_19 = x_36; +x_20 = x_35; +x_21 = x_40; x_22 = x_42; x_23 = lean_box(0); goto block_27; @@ -76527,7 +76527,7 @@ lean_object* x_43; lean_object* x_44; x_43 = lean_ctor_get(x_41, 0); lean_inc(x_43); lean_dec_ref(x_41); -x_44 = l_Lean_Elab_elabDeclAttrs___at___00Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabMutualDef_spec__4_spec__5(x_43, x_32, x_34); +x_44 = l_Lean_Elab_elabDeclAttrs___at___00Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabMutualDef_spec__4_spec__5(x_43, x_37, x_38); lean_dec(x_43); if (lean_obj_tag(x_44) == 0) { @@ -76535,12 +76535,12 @@ lean_object* x_45; x_45 = lean_ctor_get(x_44, 0); lean_inc(x_45); lean_dec_ref(x_44); -x_16 = x_35; -x_17 = x_37; -x_18 = x_36; -x_19 = x_40; -x_20 = x_38; -x_21 = x_39; +x_16 = x_33; +x_17 = x_32; +x_18 = x_34; +x_19 = x_36; +x_20 = x_35; +x_21 = x_40; x_22 = x_45; x_23 = lean_box(0); goto block_27; @@ -76548,8 +76548,8 @@ goto block_27; else { uint8_t x_46; -lean_dec(x_38); -lean_dec(x_37); +lean_dec(x_33); +lean_dec(x_32); lean_dec(x_1); x_46 = !lean_is_exclusive(x_44); if (x_46 == 0) @@ -76578,14 +76578,14 @@ if (x_62 == 0) { uint8_t x_63; x_63 = 1; -x_32 = x_59; -x_33 = lean_box(0); -x_34 = x_60; -x_35 = x_54; +x_32 = x_55; +x_33 = x_54; +x_34 = x_56; +x_35 = x_57; x_36 = x_58; -x_37 = x_55; -x_38 = x_56; -x_39 = x_57; +x_37 = x_59; +x_38 = x_60; +x_39 = lean_box(0); x_40 = x_63; goto block_49; } @@ -76593,14 +76593,14 @@ else { uint8_t x_64; x_64 = 0; -x_32 = x_59; -x_33 = lean_box(0); -x_34 = x_60; -x_35 = x_54; +x_32 = x_55; +x_33 = x_54; +x_34 = x_56; +x_35 = x_57; x_36 = x_58; -x_37 = x_55; -x_38 = x_56; -x_39 = x_57; +x_37 = x_59; +x_38 = x_60; +x_39 = lean_box(0); x_40 = x_64; goto block_49; } @@ -76643,7 +76643,7 @@ if (x_77 == 0) { lean_object* x_78; lean_object* x_79; uint8_t x_80; lean_dec(x_70); -lean_dec(x_68); +lean_dec(x_67); lean_dec(x_53); lean_dec(x_31); lean_dec(x_1); @@ -76721,8 +76721,8 @@ lean_object* x_95; lean_dec(x_93); x_95 = lean_box(0); x_66 = lean_box(0); -x_67 = x_88; -x_68 = x_86; +x_67 = x_86; +x_68 = x_88; x_69 = x_87; x_70 = x_95; goto block_85; @@ -76744,8 +76744,8 @@ lean_ctor_set(x_101, 0, x_97); lean_ctor_set(x_101, 1, x_100); lean_ctor_set(x_94, 0, x_101); x_66 = lean_box(0); -x_67 = x_88; -x_68 = x_86; +x_67 = x_86; +x_68 = x_88; x_69 = x_87; x_70 = x_94; goto block_85; @@ -76766,8 +76766,8 @@ lean_ctor_set(x_106, 1, x_105); x_107 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_107, 0, x_106); x_66 = lean_box(0); -x_67 = x_88; -x_68 = x_86; +x_67 = x_86; +x_68 = x_88; x_69 = x_87; x_70 = x_107; goto block_85; @@ -76878,7 +76878,7 @@ return x_2; LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Command_elabMutualDef_spec__7___redArg___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, uint8_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, uint8_t x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21, lean_object* x_22) { _start: { -lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; uint8_t x_70; lean_object* x_71; lean_object* x_72; lean_object* x_81; uint8_t x_82; lean_object* x_83; uint8_t x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_158; lean_object* x_159; uint8_t x_160; lean_object* x_161; lean_object* x_162; uint8_t x_163; lean_object* x_164; +lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; uint8_t x_70; lean_object* x_71; lean_object* x_72; lean_object* x_81; uint8_t x_82; lean_object* x_83; uint8_t x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_158; lean_object* x_159; uint8_t x_160; lean_object* x_161; lean_object* x_162; uint8_t x_163; lean_object* x_164; x_158 = lean_st_ref_get(x_22); x_159 = lean_ctor_get(x_158, 0); lean_inc_ref(x_159); @@ -77186,33 +77186,33 @@ lean_ctor_set(x_51, 1, x_50); x_52 = lean_array_push(x_41, x_51); if (x_42 == 0) { -lean_dec(x_38); +lean_dec(x_39); x_24 = x_52; x_25 = x_42; -x_26 = x_39; +x_26 = x_38; x_27 = x_40; x_28 = lean_box(0); goto block_36; } else { -if (lean_obj_tag(x_38) == 0) +if (lean_obj_tag(x_39) == 0) { uint8_t x_53; x_53 = 0; x_24 = x_52; x_25 = x_53; -x_26 = x_39; +x_26 = x_38; x_27 = x_40; x_28 = lean_box(0); goto block_36; } else { -lean_dec_ref(x_38); +lean_dec_ref(x_39); x_24 = x_52; x_25 = x_42; -x_26 = x_39; +x_26 = x_38; x_27 = x_40; x_28 = lean_box(0); goto block_36; @@ -77230,26 +77230,26 @@ lean_ctor_set(x_73, 1, x_5); x_74 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_74, 0, x_73); x_75 = lean_alloc_ctor(0, 10, 1); -lean_ctor_set(x_75, 0, x_68); -lean_ctor_set(x_75, 1, x_65); -lean_ctor_set(x_75, 2, x_59); -lean_ctor_set(x_75, 3, x_55); -lean_ctor_set(x_75, 4, x_62); -lean_ctor_set(x_75, 5, x_57); -lean_ctor_set(x_75, 6, x_64); -lean_ctor_set(x_75, 7, x_66); +lean_ctor_set(x_75, 0, x_64); +lean_ctor_set(x_75, 1, x_63); +lean_ctor_set(x_75, 2, x_56); +lean_ctor_set(x_75, 3, x_68); +lean_ctor_set(x_75, 4, x_67); +lean_ctor_set(x_75, 5, x_60); +lean_ctor_set(x_75, 6, x_57); +lean_ctor_set(x_75, 7, x_58); lean_ctor_set(x_75, 8, x_74); lean_ctor_set(x_75, 9, x_71); -lean_ctor_set_uint8(x_75, sizeof(void*)*10, x_61); -if (lean_obj_tag(x_60) == 0) +lean_ctor_set_uint8(x_75, sizeof(void*)*10, x_66); +if (lean_obj_tag(x_62) == 0) { -x_37 = x_56; -x_38 = x_72; -x_39 = x_75; -x_40 = x_63; -x_41 = x_69; +x_37 = x_55; +x_38 = x_75; +x_39 = x_72; +x_40 = x_61; +x_41 = x_65; x_42 = x_70; -x_43 = x_67; +x_43 = x_69; x_44 = lean_box(0); goto block_54; } @@ -77258,34 +77258,34 @@ else if (lean_obj_tag(x_72) == 0) { lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_76 = lean_ctor_get(x_60, 0); +x_76 = lean_ctor_get(x_62, 0); lean_inc(x_76); -lean_dec_ref(x_60); +lean_dec_ref(x_62); x_77 = lean_ctor_get(x_76, 1); lean_inc(x_77); lean_dec(x_76); x_78 = l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Command_elabMutualDef_spec__7___redArg___lam__1___closed__0; x_79 = l_Lean_Language_SnapshotTask_cancelRec___redArg(x_78, x_77); -x_37 = x_56; -x_38 = x_72; -x_39 = x_75; -x_40 = x_63; -x_41 = x_69; +x_37 = x_55; +x_38 = x_75; +x_39 = x_72; +x_40 = x_61; +x_41 = x_65; x_42 = x_70; -x_43 = x_67; +x_43 = x_69; x_44 = lean_box(0); goto block_54; } else { -lean_dec_ref(x_60); -x_37 = x_56; -x_38 = x_72; -x_39 = x_75; -x_40 = x_63; -x_41 = x_69; +lean_dec_ref(x_62); +x_37 = x_55; +x_38 = x_75; +x_39 = x_72; +x_40 = x_61; +x_41 = x_65; x_42 = x_70; -x_43 = x_67; +x_43 = x_69; x_44 = lean_box(0); goto block_54; } @@ -77317,21 +77317,21 @@ x_104 = lean_ctor_get(x_97, 0); lean_inc(x_104); lean_dec(x_97); x_105 = lean_box(0); -x_55 = x_88; -x_56 = x_103; -x_57 = x_90; -x_58 = lean_box(0); -x_59 = x_87; -x_60 = x_104; -x_61 = x_84; -x_62 = x_89; -x_63 = x_94; -x_64 = x_91; -x_65 = x_86; -x_66 = x_92; -x_67 = x_95; -x_68 = x_85; -x_69 = x_81; +x_55 = x_103; +x_56 = x_87; +x_57 = x_91; +x_58 = x_92; +x_59 = lean_box(0); +x_60 = x_90; +x_61 = x_94; +x_62 = x_104; +x_63 = x_86; +x_64 = x_85; +x_65 = x_81; +x_66 = x_84; +x_67 = x_89; +x_68 = x_88; +x_69 = x_95; x_70 = x_82; x_71 = x_93; x_72 = x_105; @@ -77347,21 +77347,21 @@ if (lean_obj_tag(x_106) == 0) { lean_object* x_107; x_107 = lean_box(0); -x_55 = x_88; -x_56 = x_103; -x_57 = x_90; -x_58 = lean_box(0); -x_59 = x_87; -x_60 = x_106; -x_61 = x_84; -x_62 = x_89; -x_63 = x_94; -x_64 = x_91; -x_65 = x_86; -x_66 = x_92; -x_67 = x_95; -x_68 = x_85; -x_69 = x_81; +x_55 = x_103; +x_56 = x_87; +x_57 = x_91; +x_58 = x_92; +x_59 = lean_box(0); +x_60 = x_90; +x_61 = x_94; +x_62 = x_106; +x_63 = x_86; +x_64 = x_85; +x_65 = x_81; +x_66 = x_84; +x_67 = x_89; +x_68 = x_88; +x_69 = x_95; x_70 = x_82; x_71 = x_93; x_72 = x_107; @@ -77387,21 +77387,21 @@ if (lean_obj_tag(x_113) == 0) lean_object* x_114; lean_free_object(x_108); x_114 = lean_box(0); -x_55 = x_88; -x_56 = x_103; -x_57 = x_90; -x_58 = lean_box(0); -x_59 = x_87; -x_60 = x_106; -x_61 = x_84; -x_62 = x_89; -x_63 = x_94; -x_64 = x_91; -x_65 = x_86; -x_66 = x_92; -x_67 = x_95; -x_68 = x_85; -x_69 = x_81; +x_55 = x_103; +x_56 = x_87; +x_57 = x_91; +x_58 = x_92; +x_59 = lean_box(0); +x_60 = x_90; +x_61 = x_94; +x_62 = x_106; +x_63 = x_86; +x_64 = x_85; +x_65 = x_81; +x_66 = x_84; +x_67 = x_89; +x_68 = x_88; +x_69 = x_95; x_70 = x_82; x_71 = x_93; x_72 = x_114; @@ -77427,21 +77427,21 @@ lean_dec_ref(x_117); lean_free_object(x_113); lean_free_object(x_108); x_120 = lean_box(0); -x_55 = x_88; -x_56 = x_103; -x_57 = x_90; -x_58 = lean_box(0); -x_59 = x_87; -x_60 = x_106; -x_61 = x_84; -x_62 = x_89; -x_63 = x_94; -x_64 = x_91; -x_65 = x_86; -x_66 = x_92; -x_67 = x_95; -x_68 = x_85; -x_69 = x_81; +x_55 = x_103; +x_56 = x_87; +x_57 = x_91; +x_58 = x_92; +x_59 = lean_box(0); +x_60 = x_90; +x_61 = x_94; +x_62 = x_106; +x_63 = x_86; +x_64 = x_85; +x_65 = x_81; +x_66 = x_84; +x_67 = x_89; +x_68 = x_88; +x_69 = x_95; x_70 = x_82; x_71 = x_93; x_72 = x_120; @@ -77466,21 +77466,21 @@ lean_dec_ref(x_123); lean_free_object(x_113); lean_free_object(x_108); x_125 = lean_box(0); -x_55 = x_88; -x_56 = x_103; -x_57 = x_90; -x_58 = lean_box(0); -x_59 = x_87; -x_60 = x_106; -x_61 = x_84; -x_62 = x_89; -x_63 = x_94; -x_64 = x_91; -x_65 = x_86; -x_66 = x_92; -x_67 = x_95; -x_68 = x_85; -x_69 = x_81; +x_55 = x_103; +x_56 = x_87; +x_57 = x_91; +x_58 = x_92; +x_59 = lean_box(0); +x_60 = x_90; +x_61 = x_94; +x_62 = x_106; +x_63 = x_86; +x_64 = x_85; +x_65 = x_81; +x_66 = x_84; +x_67 = x_89; +x_68 = x_88; +x_69 = x_95; x_70 = x_82; x_71 = x_93; x_72 = x_125; @@ -77493,21 +77493,21 @@ x_126 = lean_box(0); lean_ctor_set(x_108, 1, x_123); lean_ctor_set(x_108, 0, x_126); lean_ctor_set(x_113, 0, x_108); -x_55 = x_88; -x_56 = x_103; -x_57 = x_90; -x_58 = lean_box(0); -x_59 = x_87; -x_60 = x_106; -x_61 = x_84; -x_62 = x_89; -x_63 = x_94; -x_64 = x_91; -x_65 = x_86; -x_66 = x_92; -x_67 = x_95; -x_68 = x_85; -x_69 = x_81; +x_55 = x_103; +x_56 = x_87; +x_57 = x_91; +x_58 = x_92; +x_59 = lean_box(0); +x_60 = x_90; +x_61 = x_94; +x_62 = x_106; +x_63 = x_86; +x_64 = x_85; +x_65 = x_81; +x_66 = x_84; +x_67 = x_89; +x_68 = x_88; +x_69 = x_95; x_70 = x_82; x_71 = x_93; x_72 = x_113; @@ -77532,21 +77532,21 @@ lean_object* x_131; lean_dec_ref(x_128); lean_free_object(x_108); x_131 = lean_box(0); -x_55 = x_88; -x_56 = x_103; -x_57 = x_90; -x_58 = lean_box(0); -x_59 = x_87; -x_60 = x_106; -x_61 = x_84; -x_62 = x_89; -x_63 = x_94; -x_64 = x_91; -x_65 = x_86; -x_66 = x_92; -x_67 = x_95; -x_68 = x_85; -x_69 = x_81; +x_55 = x_103; +x_56 = x_87; +x_57 = x_91; +x_58 = x_92; +x_59 = lean_box(0); +x_60 = x_90; +x_61 = x_94; +x_62 = x_106; +x_63 = x_86; +x_64 = x_85; +x_65 = x_81; +x_66 = x_84; +x_67 = x_89; +x_68 = x_88; +x_69 = x_95; x_70 = x_82; x_71 = x_93; x_72 = x_131; @@ -77570,21 +77570,21 @@ lean_object* x_136; lean_dec_ref(x_134); lean_free_object(x_108); x_136 = lean_box(0); -x_55 = x_88; -x_56 = x_103; -x_57 = x_90; -x_58 = lean_box(0); -x_59 = x_87; -x_60 = x_106; -x_61 = x_84; -x_62 = x_89; -x_63 = x_94; -x_64 = x_91; -x_65 = x_86; -x_66 = x_92; -x_67 = x_95; -x_68 = x_85; -x_69 = x_81; +x_55 = x_103; +x_56 = x_87; +x_57 = x_91; +x_58 = x_92; +x_59 = lean_box(0); +x_60 = x_90; +x_61 = x_94; +x_62 = x_106; +x_63 = x_86; +x_64 = x_85; +x_65 = x_81; +x_66 = x_84; +x_67 = x_89; +x_68 = x_88; +x_69 = x_95; x_70 = x_82; x_71 = x_93; x_72 = x_136; @@ -77598,21 +77598,21 @@ lean_ctor_set(x_108, 1, x_134); lean_ctor_set(x_108, 0, x_137); x_138 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_138, 0, x_108); -x_55 = x_88; -x_56 = x_103; -x_57 = x_90; -x_58 = lean_box(0); -x_59 = x_87; -x_60 = x_106; -x_61 = x_84; -x_62 = x_89; -x_63 = x_94; -x_64 = x_91; -x_65 = x_86; -x_66 = x_92; -x_67 = x_95; -x_68 = x_85; -x_69 = x_81; +x_55 = x_103; +x_56 = x_87; +x_57 = x_91; +x_58 = x_92; +x_59 = lean_box(0); +x_60 = x_90; +x_61 = x_94; +x_62 = x_106; +x_63 = x_86; +x_64 = x_85; +x_65 = x_81; +x_66 = x_84; +x_67 = x_89; +x_68 = x_88; +x_69 = x_95; x_70 = x_82; x_71 = x_93; x_72 = x_138; @@ -77635,21 +77635,21 @@ if (lean_obj_tag(x_141) == 0) { lean_object* x_142; x_142 = lean_box(0); -x_55 = x_88; -x_56 = x_103; -x_57 = x_90; -x_58 = lean_box(0); -x_59 = x_87; -x_60 = x_106; -x_61 = x_84; -x_62 = x_89; -x_63 = x_94; -x_64 = x_91; -x_65 = x_86; -x_66 = x_92; -x_67 = x_95; -x_68 = x_85; -x_69 = x_81; +x_55 = x_103; +x_56 = x_87; +x_57 = x_91; +x_58 = x_92; +x_59 = lean_box(0); +x_60 = x_90; +x_61 = x_94; +x_62 = x_106; +x_63 = x_86; +x_64 = x_85; +x_65 = x_81; +x_66 = x_84; +x_67 = x_89; +x_68 = x_88; +x_69 = x_95; x_70 = x_82; x_71 = x_93; x_72 = x_142; @@ -77678,21 +77678,21 @@ lean_object* x_148; lean_dec_ref(x_145); lean_dec(x_144); x_148 = lean_box(0); -x_55 = x_88; -x_56 = x_103; -x_57 = x_90; -x_58 = lean_box(0); -x_59 = x_87; -x_60 = x_106; -x_61 = x_84; -x_62 = x_89; -x_63 = x_94; -x_64 = x_91; -x_65 = x_86; -x_66 = x_92; -x_67 = x_95; -x_68 = x_85; -x_69 = x_81; +x_55 = x_103; +x_56 = x_87; +x_57 = x_91; +x_58 = x_92; +x_59 = lean_box(0); +x_60 = x_90; +x_61 = x_94; +x_62 = x_106; +x_63 = x_86; +x_64 = x_85; +x_65 = x_81; +x_66 = x_84; +x_67 = x_89; +x_68 = x_88; +x_69 = x_95; x_70 = x_82; x_71 = x_93; x_72 = x_148; @@ -77716,21 +77716,21 @@ lean_object* x_153; lean_dec_ref(x_151); lean_dec(x_144); x_153 = lean_box(0); -x_55 = x_88; -x_56 = x_103; -x_57 = x_90; -x_58 = lean_box(0); -x_59 = x_87; -x_60 = x_106; -x_61 = x_84; -x_62 = x_89; -x_63 = x_94; -x_64 = x_91; -x_65 = x_86; -x_66 = x_92; -x_67 = x_95; -x_68 = x_85; -x_69 = x_81; +x_55 = x_103; +x_56 = x_87; +x_57 = x_91; +x_58 = x_92; +x_59 = lean_box(0); +x_60 = x_90; +x_61 = x_94; +x_62 = x_106; +x_63 = x_86; +x_64 = x_85; +x_65 = x_81; +x_66 = x_84; +x_67 = x_89; +x_68 = x_88; +x_69 = x_95; x_70 = x_82; x_71 = x_93; x_72 = x_153; @@ -77749,21 +77749,21 @@ if (lean_is_scalar(x_144)) { x_156 = x_144; } lean_ctor_set(x_156, 0, x_155); -x_55 = x_88; -x_56 = x_103; -x_57 = x_90; -x_58 = lean_box(0); -x_59 = x_87; -x_60 = x_106; -x_61 = x_84; -x_62 = x_89; -x_63 = x_94; -x_64 = x_91; -x_65 = x_86; -x_66 = x_92; -x_67 = x_95; -x_68 = x_85; -x_69 = x_81; +x_55 = x_103; +x_56 = x_87; +x_57 = x_91; +x_58 = x_92; +x_59 = lean_box(0); +x_60 = x_90; +x_61 = x_94; +x_62 = x_106; +x_63 = x_86; +x_64 = x_85; +x_65 = x_81; +x_66 = x_84; +x_67 = x_89; +x_68 = x_88; +x_69 = x_95; x_70 = x_82; x_71 = x_93; x_72 = x_156; @@ -79353,8 +79353,8 @@ if (x_52 == 0) lean_dec_ref(x_49); x_5 = lean_box(0); x_6 = x_42; -x_7 = x_47; -x_8 = x_43; +x_7 = x_43; +x_8 = x_47; x_9 = x_50; goto block_14; } @@ -79369,8 +79369,8 @@ if (x_52 == 0) lean_dec_ref(x_49); x_5 = lean_box(0); x_6 = x_42; -x_7 = x_47; -x_8 = x_43; +x_7 = x_43; +x_8 = x_47; x_9 = x_50; goto block_14; } @@ -79382,8 +79382,8 @@ x_55 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lea lean_dec_ref(x_49); x_5 = lean_box(0); x_6 = x_42; -x_7 = x_47; -x_8 = x_43; +x_7 = x_43; +x_8 = x_47; x_9 = x_55; goto block_14; } @@ -79396,8 +79396,8 @@ x_57 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lea lean_dec_ref(x_49); x_5 = lean_box(0); x_6 = x_42; -x_7 = x_47; -x_8 = x_43; +x_7 = x_43; +x_8 = x_47; x_9 = x_57; goto block_14; } @@ -79483,8 +79483,8 @@ if (x_79 == 0) lean_dec_ref(x_76); x_5 = lean_box(0); x_6 = x_69; -x_7 = x_74; -x_8 = x_70; +x_7 = x_70; +x_8 = x_74; x_9 = x_77; goto block_14; } @@ -79499,8 +79499,8 @@ if (x_79 == 0) lean_dec_ref(x_76); x_5 = lean_box(0); x_6 = x_69; -x_7 = x_74; -x_8 = x_70; +x_7 = x_70; +x_8 = x_74; x_9 = x_77; goto block_14; } @@ -79512,8 +79512,8 @@ x_82 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lea lean_dec_ref(x_76); x_5 = lean_box(0); x_6 = x_69; -x_7 = x_74; -x_8 = x_70; +x_7 = x_70; +x_8 = x_74; x_9 = x_82; goto block_14; } @@ -79526,8 +79526,8 @@ x_84 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lea lean_dec_ref(x_76); x_5 = lean_box(0); x_6 = x_69; -x_7 = x_74; -x_8 = x_70; +x_7 = x_70; +x_8 = x_74; x_9 = x_84; goto block_14; } @@ -79613,8 +79613,8 @@ x_11 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_11, 0, x_10); x_12 = lean_alloc_closure((void*)(l_Lean_Elab_Command_runTermElabM___boxed), 5, 2); lean_closure_set(x_12, 0, lean_box(0)); -lean_closure_set(x_12, 1, x_7); -x_13 = l_Lean_Elab_Command_withInitQuotContext(x_11, x_12, x_6, x_8); +lean_closure_set(x_12, 1, x_8); +x_13 = l_Lean_Elab_Command_withInitQuotContext(x_11, x_12, x_6, x_7); return x_13; } } diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/WF/Main.c b/stage0/stdlib/Lean/Elab/PreDefinition/WF/Main.c index 2a254cdcf825..d32f05931dcf 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/WF/Main.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/WF/Main.c @@ -1315,7 +1315,7 @@ return x_1; LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_logWarningAt___at___00Lean_Elab_wfRecursion_spec__12_spec__14___redArg(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -uint8_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_50; uint8_t x_51; lean_object* x_52; uint8_t x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; lean_object* x_57; lean_object* x_77; lean_object* x_78; uint8_t x_79; lean_object* x_80; uint8_t x_81; lean_object* x_82; uint8_t x_83; lean_object* x_84; lean_object* x_88; lean_object* x_89; lean_object* x_90; uint8_t x_91; lean_object* x_92; uint8_t x_93; uint8_t x_94; uint8_t x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; uint8_t x_105; uint8_t x_106; uint8_t x_107; uint8_t x_109; uint8_t x_125; +uint8_t x_10; uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_50; uint8_t x_51; uint8_t x_52; uint8_t x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_77; uint8_t x_78; lean_object* x_79; uint8_t x_80; lean_object* x_81; uint8_t x_82; lean_object* x_83; lean_object* x_84; lean_object* x_88; uint8_t x_89; uint8_t x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; uint8_t x_94; uint8_t x_100; lean_object* x_101; lean_object* x_102; uint8_t x_103; lean_object* x_104; lean_object* x_105; uint8_t x_106; uint8_t x_107; uint8_t x_109; uint8_t x_125; x_100 = 2; x_125 = l_Lean_instBEqMessageSeverity_beq(x_3, x_100); if (x_125 == 0) @@ -1350,15 +1350,15 @@ lean_ctor_set(x_25, 0, x_21); lean_ctor_set(x_25, 1, x_22); x_26 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_14); +lean_ctor_set(x_26, 1, x_15); x_27 = lean_alloc_ctor(0, 5, 3); -lean_ctor_set(x_27, 0, x_13); -lean_ctor_set(x_27, 1, x_16); -lean_ctor_set(x_27, 2, x_12); -lean_ctor_set(x_27, 3, x_11); +lean_ctor_set(x_27, 0, x_16); +lean_ctor_set(x_27, 1, x_13); +lean_ctor_set(x_27, 2, x_14); +lean_ctor_set(x_27, 3, x_12); lean_ctor_set(x_27, 4, x_26); -lean_ctor_set_uint8(x_27, sizeof(void*)*5, x_15); -lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 1, x_10); +lean_ctor_set_uint8(x_27, sizeof(void*)*5, x_10); +lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 1, x_11); lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 2, x_4); x_28 = l_Lean_MessageLog_add(x_27, x_24); lean_ctor_set(x_20, 6, x_28); @@ -1395,15 +1395,15 @@ lean_ctor_set(x_41, 0, x_21); lean_ctor_set(x_41, 1, x_22); x_42 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_14); +lean_ctor_set(x_42, 1, x_15); x_43 = lean_alloc_ctor(0, 5, 3); -lean_ctor_set(x_43, 0, x_13); -lean_ctor_set(x_43, 1, x_16); -lean_ctor_set(x_43, 2, x_12); -lean_ctor_set(x_43, 3, x_11); +lean_ctor_set(x_43, 0, x_16); +lean_ctor_set(x_43, 1, x_13); +lean_ctor_set(x_43, 2, x_14); +lean_ctor_set(x_43, 3, x_12); lean_ctor_set(x_43, 4, x_42); -lean_ctor_set_uint8(x_43, sizeof(void*)*5, x_15); -lean_ctor_set_uint8(x_43, sizeof(void*)*5 + 1, x_10); +lean_ctor_set_uint8(x_43, sizeof(void*)*5, x_10); +lean_ctor_set_uint8(x_43, sizeof(void*)*5 + 1, x_11); lean_ctor_set_uint8(x_43, sizeof(void*)*5 + 2, x_4); x_44 = l_Lean_MessageLog_add(x_43, x_38); x_45 = lean_alloc_ctor(0, 9, 0); @@ -1433,10 +1433,10 @@ if (x_60 == 0) { lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; x_61 = lean_ctor_get(x_59, 0); -lean_inc_ref(x_52); -x_62 = l_Lean_FileMap_toPosition(x_52, x_55); +lean_inc_ref(x_54); +x_62 = l_Lean_FileMap_toPosition(x_54, x_55); lean_dec(x_55); -x_63 = l_Lean_FileMap_toPosition(x_52, x_57); +x_63 = l_Lean_FileMap_toPosition(x_54, x_57); lean_dec(x_57); x_64 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_64, 0, x_63); @@ -1446,12 +1446,12 @@ if (x_53 == 0) lean_free_object(x_59); lean_dec_ref(x_50); x_10 = x_51; -x_11 = x_65; -x_12 = x_64; -x_13 = x_54; -x_14 = x_61; -x_15 = x_56; -x_16 = x_62; +x_11 = x_52; +x_12 = x_65; +x_13 = x_62; +x_14 = x_64; +x_15 = x_61; +x_16 = x_56; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -1468,7 +1468,7 @@ lean_object* x_67; lean_dec_ref(x_64); lean_dec_ref(x_62); lean_dec(x_61); -lean_dec_ref(x_54); +lean_dec_ref(x_56); lean_dec_ref(x_7); x_67 = lean_box(0); lean_ctor_set(x_59, 0, x_67); @@ -1478,12 +1478,12 @@ else { lean_free_object(x_59); x_10 = x_51; -x_11 = x_65; -x_12 = x_64; -x_13 = x_54; -x_14 = x_61; -x_15 = x_56; -x_16 = x_62; +x_11 = x_52; +x_12 = x_65; +x_13 = x_62; +x_14 = x_64; +x_15 = x_61; +x_16 = x_56; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -1497,10 +1497,10 @@ lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean x_68 = lean_ctor_get(x_59, 0); lean_inc(x_68); lean_dec(x_59); -lean_inc_ref(x_52); -x_69 = l_Lean_FileMap_toPosition(x_52, x_55); +lean_inc_ref(x_54); +x_69 = l_Lean_FileMap_toPosition(x_54, x_55); lean_dec(x_55); -x_70 = l_Lean_FileMap_toPosition(x_52, x_57); +x_70 = l_Lean_FileMap_toPosition(x_54, x_57); lean_dec(x_57); x_71 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_71, 0, x_70); @@ -1509,12 +1509,12 @@ if (x_53 == 0) { lean_dec_ref(x_50); x_10 = x_51; -x_11 = x_72; -x_12 = x_71; -x_13 = x_54; -x_14 = x_68; -x_15 = x_56; -x_16 = x_69; +x_11 = x_52; +x_12 = x_72; +x_13 = x_69; +x_14 = x_71; +x_15 = x_68; +x_16 = x_56; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -1531,7 +1531,7 @@ lean_object* x_74; lean_object* x_75; lean_dec_ref(x_71); lean_dec_ref(x_69); lean_dec(x_68); -lean_dec_ref(x_54); +lean_dec_ref(x_56); lean_dec_ref(x_7); x_74 = lean_box(0); x_75 = lean_alloc_ctor(0, 1, 0); @@ -1541,12 +1541,12 @@ return x_75; else { x_10 = x_51; -x_11 = x_72; -x_12 = x_71; -x_13 = x_54; -x_14 = x_68; -x_15 = x_56; -x_16 = x_69; +x_11 = x_52; +x_12 = x_72; +x_13 = x_69; +x_14 = x_71; +x_15 = x_68; +x_16 = x_56; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -1558,16 +1558,16 @@ goto block_49; block_87: { lean_object* x_85; -x_85 = l_Lean_Syntax_getTailPos_x3f(x_82, x_83); -lean_dec(x_82); +x_85 = l_Lean_Syntax_getTailPos_x3f(x_79, x_78); +lean_dec(x_79); if (lean_obj_tag(x_85) == 0) { lean_inc(x_84); x_50 = x_77; -x_51 = x_79; -x_52 = x_78; -x_53 = x_81; -x_54 = x_80; +x_51 = x_78; +x_52 = x_80; +x_53 = x_82; +x_54 = x_81; x_55 = x_84; x_56 = x_83; x_57 = x_84; @@ -1580,10 +1580,10 @@ x_86 = lean_ctor_get(x_85, 0); lean_inc(x_86); lean_dec_ref(x_85); x_50 = x_77; -x_51 = x_79; -x_52 = x_78; -x_53 = x_81; -x_54 = x_80; +x_51 = x_78; +x_52 = x_80; +x_53 = x_82; +x_54 = x_81; x_55 = x_84; x_56 = x_83; x_57 = x_86; @@ -1593,19 +1593,19 @@ goto block_76; block_99: { lean_object* x_95; lean_object* x_96; -x_95 = l_Lean_replaceRef(x_1, x_90); -lean_dec(x_90); -x_96 = l_Lean_Syntax_getPos_x3f(x_95, x_93); +x_95 = l_Lean_replaceRef(x_1, x_92); +lean_dec(x_92); +x_96 = l_Lean_Syntax_getPos_x3f(x_95, x_89); if (lean_obj_tag(x_96) == 0) { lean_object* x_97; x_97 = lean_unsigned_to_nat(0u); x_77 = x_88; x_78 = x_89; -x_79 = x_94; -x_80 = x_92; +x_79 = x_95; +x_80 = x_94; x_81 = x_91; -x_82 = x_95; +x_82 = x_90; x_83 = x_93; x_84 = x_97; goto block_87; @@ -1618,10 +1618,10 @@ lean_inc(x_98); lean_dec_ref(x_96); x_77 = x_88; x_78 = x_89; -x_79 = x_94; -x_80 = x_92; +x_79 = x_95; +x_80 = x_94; x_81 = x_91; -x_82 = x_95; +x_82 = x_90; x_83 = x_93; x_84 = x_98; goto block_87; @@ -1631,23 +1631,23 @@ goto block_87; { if (x_107 == 0) { -x_88 = x_103; -x_89 = x_101; -x_90 = x_102; -x_91 = x_105; +x_88 = x_101; +x_89 = x_106; +x_90 = x_103; +x_91 = x_102; x_92 = x_104; -x_93 = x_106; +x_93 = x_105; x_94 = x_3; goto block_99; } else { -x_88 = x_103; -x_89 = x_101; -x_90 = x_102; -x_91 = x_105; +x_88 = x_101; +x_89 = x_106; +x_90 = x_103; +x_91 = x_102; x_92 = x_104; -x_93 = x_106; +x_93 = x_105; x_94 = x_100; goto block_99; } @@ -1674,11 +1674,11 @@ if (x_119 == 0) lean_inc_ref(x_110); lean_inc(x_113); lean_inc_ref(x_111); -x_101 = x_111; -x_102 = x_113; -x_103 = x_117; -x_104 = x_110; -x_105 = x_114; +x_101 = x_117; +x_102 = x_111; +x_103 = x_114; +x_104 = x_113; +x_105 = x_110; x_106 = x_109; x_107 = x_119; goto block_108; @@ -1691,11 +1691,11 @@ x_121 = l_Lean_Option_get___at___00Lean_Elab_addMacroStack___at___00Lean_throwEr lean_inc_ref(x_110); lean_inc(x_113); lean_inc_ref(x_111); -x_101 = x_111; -x_102 = x_113; -x_103 = x_117; -x_104 = x_110; -x_105 = x_114; +x_101 = x_117; +x_102 = x_111; +x_103 = x_114; +x_104 = x_113; +x_105 = x_110; x_106 = x_109; x_107 = x_121; goto block_108; @@ -3958,13 +3958,13 @@ return x_224; block_36: { lean_object* x_32; uint8_t x_33; -lean_dec(x_28); -lean_dec_ref(x_27); -lean_dec_ref(x_25); -lean_dec_ref(x_24); -x_32 = l_Lean_setEnv___at___00Lean_Elab_wfRecursion_spec__10___redArg(x_23, x_29, x_26); +lean_dec_ref(x_29); +lean_dec_ref(x_28); +lean_dec(x_25); +lean_dec_ref(x_23); +x_32 = l_Lean_setEnv___at___00Lean_Elab_wfRecursion_spec__10___redArg(x_24, x_27, x_26); lean_dec(x_26); -lean_dec(x_29); +lean_dec(x_27); x_33 = !lean_is_exclusive(x_32); if (x_33 == 0) { @@ -3989,21 +3989,21 @@ return x_35; if (lean_obj_tag(x_44) == 0) { lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -lean_dec(x_42); -lean_dec_ref(x_41); -lean_dec_ref(x_38); +lean_dec_ref(x_42); +lean_dec(x_39); +lean_dec_ref(x_37); x_45 = lean_ctor_get(x_44, 0); lean_inc(x_45); lean_dec_ref(x_44); x_46 = lean_st_ref_get(x_40); -x_47 = l_Lean_setEnv___at___00Lean_Elab_wfRecursion_spec__10___redArg(x_37, x_43, x_40); +x_47 = l_Lean_setEnv___at___00Lean_Elab_wfRecursion_spec__10___redArg(x_38, x_41, x_40); lean_dec_ref(x_47); x_48 = lean_ctor_get(x_46, 0); lean_inc_ref(x_48); lean_dec(x_46); lean_inc(x_40); lean_inc_ref(x_48); -x_49 = l_Lean_Meta_unfoldDeclsFrom(x_48, x_45, x_39, x_40); +x_49 = l_Lean_Meta_unfoldDeclsFrom(x_48, x_45, x_43, x_40); if (lean_obj_tag(x_49) == 0) { uint8_t x_50; @@ -4026,7 +4026,7 @@ lean_ctor_set(x_52, 5, x_57); lean_ctor_set(x_52, 0, x_56); x_58 = lean_st_ref_set(x_40, x_52); lean_dec(x_40); -x_59 = lean_st_ref_take(x_43); +x_59 = lean_st_ref_take(x_41); x_60 = !lean_is_exclusive(x_59); if (x_60 == 0) { @@ -4035,8 +4035,8 @@ x_61 = lean_ctor_get(x_59, 1); lean_dec(x_61); x_62 = l_Lean_setEnv___at___00Lean_Elab_wfRecursion_spec__10___redArg___closed__3; lean_ctor_set(x_59, 1, x_62); -x_63 = lean_st_ref_set(x_43, x_59); -lean_dec(x_43); +x_63 = lean_st_ref_set(x_41, x_59); +lean_dec(x_41); x_64 = !lean_is_exclusive(x_2); if (x_64 == 0) { @@ -4102,8 +4102,8 @@ lean_ctor_set(x_81, 1, x_80); lean_ctor_set(x_81, 2, x_77); lean_ctor_set(x_81, 3, x_78); lean_ctor_set(x_81, 4, x_79); -x_82 = lean_st_ref_set(x_43, x_81); -lean_dec(x_43); +x_82 = lean_st_ref_set(x_41, x_81); +lean_dec(x_41); x_83 = lean_ctor_get(x_2, 0); lean_inc(x_83); x_84 = lean_ctor_get_uint8(x_2, sizeof(void*)*9); @@ -4189,7 +4189,7 @@ lean_ctor_set(x_104, 7, x_100); lean_ctor_set(x_104, 8, x_101); x_105 = lean_st_ref_set(x_40, x_104); lean_dec(x_40); -x_106 = lean_st_ref_take(x_43); +x_106 = lean_st_ref_take(x_41); x_107 = lean_ctor_get(x_106, 0); lean_inc_ref(x_107); x_108 = lean_ctor_get(x_106, 2); @@ -4220,8 +4220,8 @@ lean_ctor_set(x_113, 1, x_112); lean_ctor_set(x_113, 2, x_108); lean_ctor_set(x_113, 3, x_109); lean_ctor_set(x_113, 4, x_110); -x_114 = lean_st_ref_set(x_43, x_113); -lean_dec(x_43); +x_114 = lean_st_ref_set(x_41, x_113); +lean_dec(x_41); x_115 = lean_ctor_get(x_2, 0); lean_inc(x_115); x_116 = lean_ctor_get_uint8(x_2, sizeof(void*)*9); @@ -4329,7 +4329,7 @@ lean_ctor_set(x_139, 7, x_134); lean_ctor_set(x_139, 8, x_135); x_140 = lean_st_ref_set(x_40, x_139); lean_dec(x_40); -x_141 = lean_st_ref_take(x_43); +x_141 = lean_st_ref_take(x_41); x_142 = lean_ctor_get(x_141, 0); lean_inc_ref(x_142); x_143 = lean_ctor_get(x_141, 2); @@ -4360,8 +4360,8 @@ lean_ctor_set(x_148, 1, x_147); lean_ctor_set(x_148, 2, x_143); lean_ctor_set(x_148, 3, x_144); lean_ctor_set(x_148, 4, x_145); -x_149 = lean_st_ref_set(x_43, x_148); -lean_dec(x_43); +x_149 = lean_st_ref_set(x_41, x_148); +lean_dec(x_41); x_150 = lean_ctor_get(x_2, 0); lean_inc(x_150); x_151 = lean_ctor_get_uint8(x_2, sizeof(void*)*9); @@ -4418,7 +4418,7 @@ else { uint8_t x_162; lean_dec_ref(x_48); -lean_dec(x_43); +lean_dec(x_41); lean_dec(x_40); lean_dec_ref(x_2); x_162 = !lean_is_exclusive(x_49); @@ -4489,25 +4489,25 @@ lean_dec_ref(x_178); lean_inc(x_172); lean_inc_ref(x_171); x_180 = l_Lean_Elab_eraseRecAppSyntaxExpr(x_179, x_171, x_172); -x_37 = x_175; -x_38 = x_167; -x_39 = x_171; +x_37 = x_167; +x_38 = x_175; +x_39 = x_168; x_40 = x_172; -x_41 = x_169; -x_42 = x_168; -x_43 = x_170; +x_41 = x_170; +x_42 = x_169; +x_43 = x_171; x_44 = x_180; goto block_166; } else { -x_37 = x_175; -x_38 = x_167; -x_39 = x_171; +x_37 = x_167; +x_38 = x_175; +x_39 = x_168; x_40 = x_172; -x_41 = x_169; -x_42 = x_168; -x_43 = x_170; +x_41 = x_170; +x_42 = x_169; +x_43 = x_171; x_44 = x_178; goto block_166; } @@ -4524,13 +4524,13 @@ lean_dec_ref(x_2); x_181 = lean_ctor_get(x_176, 0); lean_inc(x_181); lean_dec_ref(x_176); -x_23 = x_175; -x_24 = x_167; -x_25 = x_171; +x_23 = x_167; +x_24 = x_175; +x_25 = x_168; x_26 = x_172; -x_27 = x_169; -x_28 = x_168; -x_29 = x_170; +x_27 = x_170; +x_28 = x_169; +x_29 = x_171; x_30 = x_181; x_31 = lean_box(0); goto block_36; @@ -4541,24 +4541,24 @@ goto block_36; if (lean_obj_tag(x_189) == 0) { lean_dec_ref(x_189); -x_167 = x_188; -x_168 = x_185; -x_169 = x_186; -x_170 = x_184; -x_171 = x_187; -x_172 = x_183; +x_167 = x_185; +x_168 = x_188; +x_169 = x_184; +x_170 = x_186; +x_171 = x_183; +x_172 = x_187; x_173 = lean_box(0); goto block_182; } else { uint8_t x_190; -lean_dec_ref(x_188); -lean_dec_ref(x_187); -lean_dec_ref(x_186); -lean_dec(x_185); -lean_dec(x_184); -lean_dec(x_183); +lean_dec(x_188); +lean_dec(x_187); +lean_dec(x_186); +lean_dec_ref(x_185); +lean_dec_ref(x_184); +lean_dec_ref(x_183); lean_dec_ref(x_15); lean_dec_ref(x_9); lean_dec_ref(x_8); @@ -4637,12 +4637,12 @@ size_t x_207; lean_object* x_208; x_207 = lean_usize_of_nat(x_204); lean_inc_ref(x_198); x_208 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Elab_wfRecursion_spec__14(x_10, x_11, x_12, x_13, x_14, x_205, x_6, x_5, x_207, x_14, x_194, x_195, x_196, x_197, x_198, x_199); -x_183 = x_199; -x_184 = x_197; -x_185 = x_195; -x_186 = x_196; -x_187 = x_198; -x_188 = x_194; +x_183 = x_198; +x_184 = x_196; +x_185 = x_194; +x_186 = x_197; +x_187 = x_199; +x_188 = x_195; x_189 = x_208; goto block_193; } @@ -4653,12 +4653,12 @@ size_t x_209; lean_object* x_210; x_209 = lean_usize_of_nat(x_204); lean_inc_ref(x_198); x_210 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Elab_wfRecursion_spec__14(x_10, x_11, x_12, x_13, x_14, x_205, x_6, x_5, x_209, x_14, x_194, x_195, x_196, x_197, x_198, x_199); -x_183 = x_199; -x_184 = x_197; -x_185 = x_195; -x_186 = x_196; -x_187 = x_198; -x_188 = x_194; +x_183 = x_198; +x_184 = x_196; +x_185 = x_194; +x_186 = x_197; +x_187 = x_199; +x_188 = x_195; x_189 = x_210; goto block_193; } @@ -5616,7 +5616,7 @@ lean_inc_ref(x_4); x_38 = l_Lean_withEnv___at___00Lean_Elab_wfRecursion_spec__7___redArg(x_37, x_36, x_4, x_5, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_38) == 0) { -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_129; size_t x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; uint8_t x_190; +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_129; size_t x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; uint8_t x_190; x_39 = lean_ctor_get(x_38, 0); lean_inc(x_39); lean_dec_ref(x_38); @@ -5727,11 +5727,11 @@ lean_inc(x_54); lean_inc_ref(x_53); lean_inc(x_52); lean_inc_ref(x_51); -lean_inc_ref(x_47); +lean_inc_ref(x_46); lean_inc(x_14); lean_inc(x_43); lean_inc(x_41); -x_56 = l_Lean_Elab_WF_preDefsFromUnaryNonRec(x_41, x_43, x_14, x_47, x_51, x_52, x_53, x_54); +x_56 = l_Lean_Elab_WF_preDefsFromUnaryNonRec(x_41, x_43, x_14, x_46, x_51, x_52, x_53, x_54); if (lean_obj_tag(x_56) == 0) { lean_object* x_57; lean_object* x_58; @@ -5744,10 +5744,10 @@ lean_inc(x_52); lean_inc_ref(x_51); lean_inc(x_50); lean_inc_ref(x_49); -lean_inc_ref(x_47); +lean_inc_ref(x_46); lean_inc(x_14); lean_inc_ref(x_1); -x_58 = l_Lean_Elab_Mutual_addPreDefsFromUnary(x_1, x_14, x_57, x_47, x_46, x_49, x_50, x_51, x_52, x_53, x_54); +x_58 = l_Lean_Elab_Mutual_addPreDefsFromUnary(x_1, x_14, x_57, x_46, x_48, x_49, x_50, x_51, x_52, x_53, x_54); lean_dec(x_57); if (lean_obj_tag(x_58) == 0) { @@ -5769,7 +5769,7 @@ lean_inc(x_54); lean_inc_ref(x_53); lean_inc(x_52); lean_inc_ref(x_51); -x_60 = l_Lean_Elab_Mutual_cleanPreDef(x_44, x_46, x_51, x_52, x_53, x_54); +x_60 = l_Lean_Elab_Mutual_cleanPreDef(x_44, x_48, x_51, x_52, x_53, x_54); if (lean_obj_tag(x_60) == 0) { lean_object* x_61; lean_object* x_62; @@ -5780,16 +5780,16 @@ lean_inc(x_54); lean_inc_ref(x_53); lean_inc(x_52); lean_inc_ref(x_51); -x_62 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_wfRecursion_spec__17___redArg(x_46, x_33, x_12, x_14, x_51, x_52, x_53, x_54); +x_62 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_wfRecursion_spec__17___redArg(x_48, x_33, x_12, x_14, x_51, x_52, x_53, x_54); if (lean_obj_tag(x_62) == 0) { lean_object* x_63; lean_object* x_64; lean_object* x_65; x_63 = lean_ctor_get(x_62, 0); lean_inc(x_63); lean_dec_ref(x_62); -x_64 = lean_ctor_get(x_47, 3); +x_64 = lean_ctor_get(x_46, 3); lean_inc(x_64); -lean_dec_ref(x_47); +lean_dec_ref(x_46); lean_inc(x_54); lean_inc_ref(x_53); lean_inc(x_52); @@ -5831,12 +5831,12 @@ lean_inc_ref(x_53); lean_inc(x_52); lean_inc_ref(x_51); lean_inc(x_64); -x_72 = l_Lean_Elab_WF_mkUnfoldEq(x_61, x_64, x_48, x_51, x_52, x_53, x_54); +x_72 = l_Lean_Elab_WF_mkUnfoldEq(x_61, x_64, x_47, x_51, x_52, x_53, x_54); if (lean_obj_tag(x_72) == 0) { lean_dec_ref(x_72); -x_19 = x_63; -x_20 = x_64; +x_19 = x_64; +x_20 = x_63; x_21 = x_49; x_22 = x_50; x_23 = x_51; @@ -5862,9 +5862,9 @@ return x_72; else { lean_dec(x_61); -lean_dec_ref(x_48); -x_19 = x_63; -x_20 = x_64; +lean_dec_ref(x_47); +x_19 = x_64; +x_20 = x_63; x_21 = x_49; x_22 = x_50; x_23 = x_51; @@ -5887,7 +5887,7 @@ lean_dec(x_52); lean_dec_ref(x_51); lean_dec(x_50); lean_dec_ref(x_49); -lean_dec_ref(x_48); +lean_dec_ref(x_47); x_73 = !lean_is_exclusive(x_69); if (x_73 == 0) { @@ -5916,7 +5916,7 @@ lean_dec(x_52); lean_dec_ref(x_51); lean_dec(x_50); lean_dec_ref(x_49); -lean_dec_ref(x_48); +lean_dec_ref(x_47); return x_68; } } @@ -5931,7 +5931,7 @@ lean_dec(x_52); lean_dec_ref(x_51); lean_dec(x_50); lean_dec_ref(x_49); -lean_dec_ref(x_48); +lean_dec_ref(x_47); return x_65; } } @@ -5945,8 +5945,8 @@ lean_dec(x_52); lean_dec_ref(x_51); lean_dec(x_50); lean_dec_ref(x_49); -lean_dec_ref(x_48); lean_dec_ref(x_47); +lean_dec_ref(x_46); lean_dec(x_43); lean_dec(x_41); x_76 = !lean_is_exclusive(x_62); @@ -5975,8 +5975,8 @@ lean_dec(x_52); lean_dec_ref(x_51); lean_dec(x_50); lean_dec_ref(x_49); -lean_dec_ref(x_48); lean_dec_ref(x_47); +lean_dec_ref(x_46); lean_dec(x_43); lean_dec(x_41); lean_dec(x_14); @@ -6005,8 +6005,8 @@ lean_dec(x_52); lean_dec_ref(x_51); lean_dec(x_50); lean_dec_ref(x_49); -lean_dec_ref(x_48); lean_dec_ref(x_47); +lean_dec_ref(x_46); lean_dec(x_44); lean_dec(x_43); lean_dec(x_41); @@ -6022,8 +6022,8 @@ lean_dec(x_52); lean_dec_ref(x_51); lean_dec(x_50); lean_dec_ref(x_49); -lean_dec_ref(x_48); lean_dec_ref(x_47); +lean_dec_ref(x_46); lean_dec(x_44); lean_dec(x_43); lean_dec(x_41); @@ -6041,8 +6041,8 @@ lean_dec(x_52); lean_dec_ref(x_51); lean_dec(x_50); lean_dec_ref(x_49); -lean_dec_ref(x_48); lean_dec_ref(x_47); +lean_dec_ref(x_46); lean_dec(x_44); lean_dec(x_43); lean_dec(x_41); @@ -6127,9 +6127,9 @@ if (x_114 == 0) { lean_dec(x_45); lean_dec(x_42); -x_46 = x_109; -x_47 = x_111; -x_48 = x_93; +x_46 = x_111; +x_47 = x_93; +x_48 = x_109; x_49 = x_95; x_50 = x_96; x_51 = x_97; @@ -6173,9 +6173,9 @@ x_124 = l_Lean_addTrace___at___00Lean_Elab_wfRecursion_spec__15___redArg(x_86, x if (lean_obj_tag(x_124) == 0) { lean_dec_ref(x_124); -x_46 = x_109; -x_47 = x_111; -x_48 = x_93; +x_46 = x_111; +x_47 = x_93; +x_48 = x_109; x_49 = x_95; x_50 = x_96; x_51 = x_97; @@ -6243,12 +6243,12 @@ return x_127; if (lean_obj_tag(x_131) == 1) { lean_object* x_142; -lean_dec_ref(x_133); +lean_dec_ref(x_134); x_142 = lean_ctor_get(x_131, 0); lean_inc(x_142); lean_dec_ref(x_131); x_92 = x_132; -x_93 = x_134; +x_93 = x_133; x_94 = x_142; x_95 = x_135; x_96 = x_136; @@ -6270,7 +6270,7 @@ lean_inc(x_138); lean_inc_ref(x_137); lean_inc(x_136); lean_inc_ref(x_135); -x_144 = l_Lean_withoutExporting___at___00Lean_Elab_wfRecursion_spec__19___redArg(x_133, x_143, x_135, x_136, x_137, x_138, x_139, x_140); +x_144 = l_Lean_withoutExporting___at___00Lean_Elab_wfRecursion_spec__19___redArg(x_134, x_143, x_135, x_136, x_137, x_138, x_139, x_140); if (lean_obj_tag(x_144) == 0) { lean_object* x_145; @@ -6278,7 +6278,7 @@ x_145 = lean_ctor_get(x_144, 0); lean_inc(x_145); lean_dec_ref(x_144); x_92 = x_132; -x_93 = x_134; +x_93 = x_133; x_94 = x_145; x_95 = x_135; x_96 = x_136; @@ -6298,7 +6298,7 @@ lean_dec(x_138); lean_dec_ref(x_137); lean_dec(x_136); lean_dec_ref(x_135); -lean_dec_ref(x_134); +lean_dec_ref(x_133); lean_dec_ref(x_132); lean_dec(x_89); lean_dec(x_45); @@ -6372,8 +6372,8 @@ if (x_168 == 0) { lean_free_object(x_161); x_132 = x_163; -x_133 = x_167; -x_134 = x_164; +x_133 = x_164; +x_134 = x_167; x_135 = x_150; x_136 = x_151; x_137 = x_152; @@ -6399,8 +6399,8 @@ if (lean_obj_tag(x_173) == 0) { lean_dec_ref(x_173); x_132 = x_163; -x_133 = x_167; -x_134 = x_164; +x_133 = x_164; +x_134 = x_167; x_135 = x_150; x_136 = x_151; x_137 = x_152; @@ -6460,8 +6460,8 @@ lean_dec(x_177); if (x_179 == 0) { x_132 = x_174; -x_133 = x_178; -x_134 = x_175; +x_133 = x_175; +x_134 = x_178; x_135 = x_150; x_136 = x_151; x_137 = x_152; @@ -6487,8 +6487,8 @@ if (lean_obj_tag(x_185) == 0) { lean_dec_ref(x_185); x_132 = x_174; -x_133 = x_178; -x_134 = x_175; +x_133 = x_175; +x_134 = x_178; x_135 = x_150; x_136 = x_151; x_137 = x_152; @@ -6590,24 +6590,24 @@ return x_199; block_32: { size_t x_28; lean_object* x_29; -x_28 = lean_array_size(x_19); +x_28 = lean_array_size(x_20); lean_inc(x_26); lean_inc_ref(x_25); lean_inc(x_24); lean_inc_ref(x_23); -lean_inc(x_20); -x_29 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_wfRecursion_spec__18___redArg(x_20, x_18, x_19, x_28, x_12, x_18, x_23, x_24, x_25, x_26); +lean_inc(x_19); +x_29 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_wfRecursion_spec__18___redArg(x_19, x_18, x_20, x_28, x_12, x_18, x_23, x_24, x_25, x_26); if (lean_obj_tag(x_29) == 0) { lean_object* x_30; lean_dec_ref(x_29); lean_inc_ref(x_25); -x_30 = l_Lean_enableRealizationsForConst(x_20, x_25, x_26); +x_30 = l_Lean_enableRealizationsForConst(x_19, x_25, x_26); if (lean_obj_tag(x_30) == 0) { lean_object* x_31; lean_dec_ref(x_30); -x_31 = l_Lean_Elab_Mutual_addPreDefAttributes(x_19, x_21, x_22, x_23, x_24, x_25, x_26); +x_31 = l_Lean_Elab_Mutual_addPreDefAttributes(x_20, x_21, x_22, x_23, x_24, x_25, x_26); return x_31; } else @@ -6618,7 +6618,7 @@ lean_dec(x_24); lean_dec_ref(x_23); lean_dec(x_22); lean_dec_ref(x_21); -lean_dec_ref(x_19); +lean_dec_ref(x_20); return x_30; } } @@ -6630,8 +6630,8 @@ lean_dec(x_24); lean_dec_ref(x_23); lean_dec(x_22); lean_dec_ref(x_21); -lean_dec(x_20); -lean_dec_ref(x_19); +lean_dec_ref(x_20); +lean_dec(x_19); return x_29; } } diff --git a/stage0/stdlib/Lean/Elab/Tactic/Do/VCGen.c b/stage0/stdlib/Lean/Elab/Tactic/Do/VCGen.c index 708229ebd7dd..e54a9933270f 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Do/VCGen.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Do/VCGen.c @@ -28,156 +28,158 @@ LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_m LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13_spec__26___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__9___redArg___closed__10; lean_object* l_Lean_Meta_mkEqHEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13_spec__26___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13_spec__26(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__2; static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__15___closed__20; static lean_object* l___private_Init_Data_Array_Basic_0__Array_foldrMUnsafe_fold___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__10___closed__6; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldrMUnsafe_fold___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__10(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_inferArgumentTypesN(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__9___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__51___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_logWarningAt___at___00Lean_Elab_Tactic_Do_elabMVCGen_spec__5___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__14(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_collectFreshMVars___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__12___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__17(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__0; lean_object* l_Lean_MessageData_paren(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__6(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57_spec__65(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___closed__8; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__5___closed__11; static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__5___redArg___closed__4; LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__0; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg(lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_Do_ProofMode_MGoal_assumption(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__5_spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__43___lam__0(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___closed__12; LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabInvariants___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlMAux___at___00Lean_PersistentArray_foldlM___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__1_spec__4___boxed(lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__0; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Range_Basic_0__Std_Legacy_Range_forIn_x27_loop___at___00Lean_Elab_Tactic_Do_VCGen_genVCs_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntroForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__11___lam__0___closed__0; static lean_object* l_Lean_addTrace___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__5___lam__1___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_VCGen_genVCs___lam__0___closed__0; static lean_object* l_Lean_Elab_Tactic_Do_VCGen_genVCs___lam__0___closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_sortMVarIdArrayByIndex___at___00Lean_Elab_Tactic_sortMVarIdsByIndex___at___00Lean_Elab_Tactic_collectFreshMVars___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__12_spec__24_spec__31(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_collectFreshMVars___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__12___lam__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__6; lean_object* l_Lean_Elab_Tactic_Do_reduceProjBeta_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_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_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15___lam__4(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntro___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__3___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Core_instMonadCoreM___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__42___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__7; LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__2___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__1_spec__2___closed__0; uint64_t l_Lean_Meta_Context_configKey(lean_object*); +static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__10; static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__28_spec__33_spec__38___redArg___closed__4; -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56_spec__64___redArg(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__9___redArg___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntroForall___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_right(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_Do_burnOne___redArg(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13_spec__26___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__22(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__17___redArg___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__61___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mIntroForall___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__4_spec__5_spec__9___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__42___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__42___lam__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapFinIdxM_map___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__23___redArg___closed__1; LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00Lean_Elab_Tactic_Do_elabMVCGen_spec__1___redArg(lean_object*, lean_object*); lean_object* l_Lean_TSyntax_getId(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__26(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_VCGen_genVCs___lam__0___closed__5; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Elab_Tactic_Do_elabMVCGen_spec__0___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntroForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__11___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__1___closed__0; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___lam__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__16___lam__0(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeMGoal___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__21(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__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_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56_spec__64(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_String_Slice_toNat_x3f(lean_object*); +static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLetDeclImp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__15___closed__8; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static double l_Lean_addTrace___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__5___lam__1___closed__0; -static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__4; static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__14___closed__5; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__0; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___lam__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_Do_ProofMode_MGoal_toExpr(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__61___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntro___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__3___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__5_spec__7___boxed(lean_object**); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___lam__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___00Lean_Elab_Tactic_Do_ProofMode_mIntro___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__3_spec__3___redArg(lean_object*); lean_object* l_Lean_Elab_Tactic_Do_mkSpecContext(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_sortMVarIdArrayByIndex___at___00Lean_Elab_Tactic_sortMVarIdsByIndex___at___00Lean_Elab_Tactic_collectFreshMVars___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__12_spec__24_spec__31___redArg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__2___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_reverse___redArg(lean_object*); +static lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__7; LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_logErrorAt___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__1_spec__1___redArg(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_QSort_Basic_0__Array_qsort_sort___at___00Lean_Elab_Tactic_sortMVarIdArrayByIndex___at___00Lean_Elab_Tactic_sortMVarIdsByIndex___at___00Lean_Elab_Tactic_collectFreshMVars___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__12_spec__24_spec__31_spec__47___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_Do_addSubGoalAsVC(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62___lam__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__9___redArg___closed__13; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__56___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__19___redArg___lam__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13_spec__26___lam__1___closed__0; lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mStartMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__56(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_whnf(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__57___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__5___closed__12; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__63___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_elabVCs_evalAlts_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkApp7(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_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__50___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MetavarContext_getDecl(lean_object*, lean_object*); uint8_t lean_usize_dec_le(size_t, size_t); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57_spec__65___redArg(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__19___redArg___lam__4___closed__5; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mStart(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_indentD(lean_object*); double lean_float_div(double, double); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___lam__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeMGoal___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__21___lam__2___closed__0; -LEAN_EXPORT lean_object* l_Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_getGoals___redArg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__57___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_tryGoal_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at___00Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__6_spec__8_spec__12_spec__34___redArg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Elab_Tactic_Do_elabMVCGen_spec__0___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__54(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Exception_isInterrupt(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_tryGoal___lam__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkApp8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__1; lean_object* l_Lean_mkAppB(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeMGoal___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__21___closed__1; lean_object* l_Lean_Elab_Tactic_Do_knownJP_x3f___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_ProofMode_MGoal_withNewProg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___closed__0; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__10(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Elab_Tactic_Do_elabMVCGen_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -190,7 +192,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Ela lean_object* l_Lean_Meta_isExprDefEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_elabMVCGenHint(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_FileMap_toPosition(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__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_Elab_Tactic_Do_findSpec(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__9___redArg___closed__22; static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__15___closed__1; @@ -205,53 +206,53 @@ uint8_t l_Lean_Expr_isLet(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__15(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_mapTR_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__8___closed__0; uint8_t l_Lean_Expr_isAppOfArity(lean_object*, lean_object*, lean_object*); +lean_object* l_Array_mask___redArg(lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_Do_ProofMode_SPred_mkAnd(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp(lean_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_Do_VCGen_genVCs___lam__0(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___closed__1; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_tryGoal___lam__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__0___closed__1; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__49___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___closed__4; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__19_spec__24(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__19___redArg___boxed(lean_object*, 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_instBEqMVarId_beq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__28_spec__33_spec__38___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isApp(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___lam__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__8; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__63(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___lam__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__43___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntro___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__3___lam__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_Tactic_tacticElabAttribute; -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_elabVCs_evalAlts___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__1; static lean_object* l___private_Init_Data_Array_Basic_0__Array_foldrMUnsafe_fold___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__10___closed__7; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__7; lean_object* l_Lean_Expr_sort___override(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___lam__4(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_trySynthInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_qpartition___redArg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofList(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkSort(lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_Tactic_Do_elabMVCGen___lam__4(uint8_t, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Tactic_Do_elabMVCGen_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__15___closed__4; lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___lam__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mIntroForall___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__4_spec__5_spec__9(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__3___redArg(size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__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_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__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_Basic_0__Array_foldrMUnsafe_fold___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__25___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__16___closed__7; lean_object* l_Lean_Elab_Tactic_filterOldMVars___redArg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_mul(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__32___lam__1___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_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__9_spec__13___redArg___boxed(lean_object**); @@ -261,52 +262,53 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tact LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeMGoal___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__21___lam__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___closed__6; LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__9_spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabInvariants___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__3___closed__0; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__56___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); +static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__7; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_elabVCs_applyPreTac___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_KVMap_find(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_instReprTSyntax_repr___redArg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__61(lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instFunctorOfMonad___redArg___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntroForall___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__4___closed__4; static lean_object* l_Lean_Elab_Tactic_Do_elabInvariants___closed__0; lean_object* l_Lean_Syntax_getArgs(lean_object*); LEAN_EXPORT lean_object* l_Lean_logInfoAt___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___closed__7; static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__4___redArg___closed__0; static lean_object* l_Lean_Elab_Tactic_Do_elabInvariants___closed__7; -static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__6; lean_object* l___private_Lean_Log_0__Lean_MessageData_appendDescriptionWidgetIfNamed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeMGoal___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__21___lam__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__9___redArg___closed__19; static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__1___closed__4; static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___closed__1; +LEAN_EXPORT lean_object* l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__4_spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Tactic_Do_elabMVCGen_spec__2(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGenHint___lam__0___closed__3; LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__3___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___closed__5; lean_object* l_Lean_replaceRef(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62___lam__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_LocalDecl_isNondep(lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabInvariants___closed__12; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__27___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___lam__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__16___closed__10; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mIntroForall___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__4_spec__5_spec__9___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_sortMVarIdsByIndex___at___00Lean_Elab_Tactic_collectFreshMVars___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__12_spec__24___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__9; +static lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__5; lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__26___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__28_spec__33___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__9___redArg___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -316,44 +318,48 @@ lean_object* l_Lean_Elab_Tactic_Do_ProofMode_addLocalVarInfo(lean_object*, lean_ static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__9___closed__0; lean_object* l_Lean_Syntax_getTailPos_x3f(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at___00Lean_Elab_Tactic_Do_VCGen_genVCs_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__55___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabInvariants___closed__1; -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__3; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37(size_t, size_t, lean_object*); uint8_t l_Lean_instBEqMessageSeverity_beq(uint8_t, uint8_t); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__16___closed__11; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_sortMVarIdsByIndex___at___00Lean_Elab_Tactic_collectFreshMVars___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__12_spec__24(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__9___redArg___closed__20; -LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__48(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_getTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_tryGoal_spec__1___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__4; +static lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__1___closed__0; lean_object* l_Lean_Expr_fvarId_x21(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13_spec__28___lam__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__28_spec__33_spec__38___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_getKind(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__8___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___lam__5(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__9; +LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_float_decLt(double, double); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__2; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_VCGen_genVCs___lam__0___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__14___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__14___closed__4; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__7(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_collectFreshMVars___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__12___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeMGoal___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__21___lam__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__58___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__9___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntro___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__3___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__7___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_MessageData_hasSyntheticSorry(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15___lam__1(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___lam__5(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -369,22 +375,20 @@ LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_m LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntroForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__11___lam__0(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAuxAux(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isAppOf(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__50___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__51___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13_spec__29___redArg(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeMGoal___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__21___lam__0___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_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__5; LEAN_EXPORT lean_object* l_Lean_addTrace___at___00Lean_Elab_Tactic_Do_elabMVCGen_spec__4___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___closed__0; LEAN_EXPORT lean_object* l_Lean_Meta_lambdaBoundedTelescope___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__2(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__19___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__9_spec__13___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instApplicativeOfMonad___redArg___lam__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_tryGoal_spec__0___redArg___lam__0___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_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__5; static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___closed__10; +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___boxed__const__1; static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___closed__3; uint8_t l_Lean_Name_quickLt(lean_object*, lean_object*); lean_object* lean_io_get_num_heartbeats(); @@ -400,79 +404,99 @@ LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Uns static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__15___closed__16; extern lean_object* l_Lean_trace_profiler_useHeartbeats; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__51___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__61___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___lam__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13___redArg___lam__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_filterMapM___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__13___closed__0; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__43___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_instInhabitedPersistentArrayNode_default(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__12_spec__16_spec__18_spec__23(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_zipWithMAux___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntro___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__3___lam__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkApp4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13___redArg___closed__6; static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___regBuiltin_Lean_Elab_Tactic_Do_elabMVCGen__1___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntroForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__11___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___lam__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56_spec__64___boxed(lean_object**); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___lam__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__2; static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___closed__8; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__56___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_Do_ifOutOfFuel___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__5; static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__15___closed__19; static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__3___closed__1; uint8_t lean_string_dec_eq(lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_Do_getNumJoinParams(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__15___closed__13; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__0(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__6_spec__8_spec__12___redArg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___regBuiltin_Lean_Elab_Tactic_Do_elabMVCGen__1(); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__5_spec__7___redArg___boxed(lean_object**); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_mapTR_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__8___closed__3; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__11___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_throwErrorAt___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__7_spec__10___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_VCGen_genVCs___lam__0___closed__7; lean_object* l_Lean_Exception_toMessageData(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeMGoal___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__21___lam__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static size_t l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__12_spec__16_spec__18___redArg___closed__0; +LEAN_EXPORT uint8_t l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2(uint8_t, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__3___closed__0; +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__3___closed__0; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__21___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeMGoal___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__21___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__50___redArg___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_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mIntroForall___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__4_spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___lam__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___closed__5; -LEAN_EXPORT lean_object* l_Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_mapTR_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__8___closed__2; lean_object* l_Array_mkArray0(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__11(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntroForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__11___lam__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_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__51___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__5___closed__5; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__4___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_utf8_byte_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___closed__9; +static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__3; lean_object* l_Lean_Expr_appArg_x21(lean_object*); -LEAN_EXPORT lean_object* l_Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___lam__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instMonad___redArg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__10___redArg(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntro___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__3___lam__6(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__50___lam__0(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_VCGen_genVCs___lam__0___closed__4; static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___closed__0; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntro___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__3___lam__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__4; static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___regBuiltin_Lean_Elab_Tactic_Do_elabMVCGen__1___closed__4; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__19___boxed(lean_object**); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__9___redArg___closed__23; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___closed__4; -LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15___boxed__const__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___lam__9(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_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__14___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -482,11 +506,14 @@ LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldrMUnsafe LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13_spec__28___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_hasMVar(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___closed__0; +static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__11; LEAN_EXPORT lean_object* l_Lean_throwError___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__0___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_foldl___at___00Array_appendList_spec__0___redArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__10; static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___lam__7___closed__2; lean_object* l_Lean_Meta_instantiateMVarsIfMVarApp___redArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__15___closed__17; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__15___boxed(lean_object**); @@ -496,13 +523,11 @@ lean_object* l_Nat_reprFast(lean_object*); lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_withUserNamesImpl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isTypeCorrect(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___lam__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__8___closed__3; LEAN_EXPORT lean_object* l_Lean_addTrace___at___00Lean_Elab_Tactic_Do_elabMVCGen_spec__4___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___boxed(lean_object**); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mIntroForall___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__4_spec__5_spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_tryGoal_spec__0___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -510,7 +535,6 @@ LEAN_EXPORT uint8_t l_Lean_Option_get___at___00Lean_withTraceNode___at___00__pri lean_object* l_Lean_Elab_Tactic_Do_ProofMode_Hyp_toExpr(lean_object*); extern lean_object* l_Lean_Elab_Tactic_Do_mvcgen_warning; lean_object* l_Lean_mkAppRev(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGenHint___regBuiltin_Lean_Elab_Tactic_Do_elabMVCGenHint__1___closed__3; @@ -523,20 +547,23 @@ LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaq lean_object* lean_expr_abstract(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___closed__9; lean_object* l_Lean_Elab_Tactic_getMainGoal___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__55___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___lam__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___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_addMessageContextFull___at___00Lean_throwError___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__0_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabInvariants___closed__4; LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__6_spec__8(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__5___closed__13; lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_mkProj_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntroForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__11___lam__0___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___lam__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_logErrorAt___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__6; +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__0; static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__16___closed__9; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___lam__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -551,6 +578,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic lean_object* l_Lean_Meta_Simp_Result_mkEqMPR(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at___00Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__6_spec__8_spec__12_spec__34___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__12___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_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13_spec__29(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__5___closed__7; @@ -560,18 +588,16 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLoca LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__12_spec__16_spec__18___redArg(lean_object*, size_t, size_t, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__28_spec__33_spec__38___redArg___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__32___lam__1(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntroForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__11___lam__0___closed__5; uint8_t lean_expr_eqv(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__17___redArg(lean_object*, lean_object*); lean_object* l_Lean_Meta_Match_Extension_getMatcherInfo_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__19___redArg(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__16___closed__0; -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGenHint___lam__0___closed__0; LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__4_spec__5___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13(lean_object*, lean_object*, lean_object*, uint8_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_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___boxed(lean_object**); static lean_object* l_Lean_Elab_Tactic_Do_elabInvariants___closed__16; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__19___redArg___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -586,14 +612,12 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tact static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___closed__6; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__9___redArg(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___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__7___boxed(lean_object**); -static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__9; -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___boxed(lean_object**); lean_object* l_Lean_MessageData_ofSyntax(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___lam__7___boxed(lean_object**); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___lam__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_MVarId_assign___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_elabVCs_evalAlts(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__5___redArg___closed__2; lean_object* l_instMonadEST(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__28_spec__33_spec__38_spec__42___redArg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -602,23 +626,26 @@ LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at___00Lean_Elab_Tactic_Do_mSpec lean_object* l_Array_empty(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__15___closed__3; -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__14(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15___closed__2; LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__23___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_Do_elimLets(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__4___redArg___closed__1; -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__4___redArg___boxed(lean_object**); static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGenHint___regBuiltin_Lean_Elab_Tactic_Do_elabMVCGenHint__1___closed__0; +static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___lam__6(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Array_Basic_0__Array_foldrMUnsafe_fold___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__10___closed__0; +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_tryGoal___lam__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_FVarId_getValue_x3f___redArg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___closed__3; lean_object* l_Lean_Meta_mkHEqRefl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___lam__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__51___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00Lean_Elab_Tactic_Do_elabMVCGen_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_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__4___boxed(lean_object**); static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___closed__9; @@ -626,14 +653,19 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___0 lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_resTy(lean_object*); LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__23___redArg___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_mkFreshId___at___00Lean_Elab_Tactic_Do_ProofMode_mIntro___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__3_spec__3(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_abstractM(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntro___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__3___closed__0; static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__9___redArg___closed__3; lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withMVarContextImp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at___00Lean_Elab_Tactic_Do_elabMVCGen_spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGenHint___regBuiltin_Lean_Elab_Tactic_Do_elabMVCGenHint__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__43(lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___lam__9___boxed(lean_object**); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__16___boxed(lean_object**); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__43___lam__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__21___lam__0(lean_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_Do_ProofMode_mIntroForall___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__4___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__28_spec__33_spec__38___redArg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -641,25 +673,27 @@ static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_e LEAN_EXPORT lean_object* l_String_dropPrefix_x3f___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__8___redArg(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___lam__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___closed__12; -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGenHint___lam__0___closed__1; -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__45(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofFormat(lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___lam__7___closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_check(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal___lam__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabVCs___closed__2; static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__3___closed__0; +static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__28_spec__33_spec__38_spec__42___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabInvariants___closed__11; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__49(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___lam__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabVCs___closed__1; static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___closed__2; @@ -672,55 +706,46 @@ LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at___00__private_Lean_Elab_Tac lean_object* l_Lean_Meta_getLevel(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntroForall___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__4___lam__1___closed__0; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___lam__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGenHint___lam__0___closed__4; -static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__1; lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntroForall___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__4___lam__0(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__1(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_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__19___redArg___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_replaceMainGoal___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57_spec__65___boxed(lean_object**); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__9___redArg___closed__18; lean_object* l_ReaderT_instApplicativeOfMonad___redArg___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntroForall___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__4___closed__1; double lean_float_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__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_throwError___at___00Lean_throwErrorAt___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__7_spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___closed__0; lean_object* l_instInhabitedOfMonad___redArg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__48___redArg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__19___redArg___lam__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15___closed__8; -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___closed__3; lean_object* lean_st_ref_get(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__49___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__12_spec__16_spec__18_spec__24___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__14___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_tryGoal___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___closed__6; -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__42(lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_pop(lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkAppM(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__5(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__5___redArg___boxed(lean_object**); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__16(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__9___redArg___closed__11; -static lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__0; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___lam__6___boxed(lean_object**); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__49___lam__0(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__3___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Range_Basic_0__Std_Legacy_Range_forIn_x27_loop___at___00Lean_Elab_Tactic_Do_VCGen_genVCs_spec__1___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__50___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkAndN(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_QSort_Basic_0__Array_qsort_sort___at___00Lean_Elab_Tactic_sortMVarIdArrayByIndex___at___00Lean_Elab_Tactic_sortMVarIdsByIndex___at___00Lean_Elab_Tactic_collectFreshMVars___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__12_spec__24_spec__31_spec__47___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getOptional_x3f(lean_object*); lean_object* lean_st_mk_ref(lean_object*); lean_object* l_Lean_Meta_reduceMatcher_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -728,8 +753,9 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tact LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__19___redArg___lam__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__0___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__3___closed__1; LEAN_EXPORT lean_object* l_Lean_addTrace___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*); lean_object* l_Lean_Meta_mkEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -737,15 +763,17 @@ static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_e LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__12_spec__16_spec__18___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabInvariants___closed__19; LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__17___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51(lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__5; static lean_object* l_Lean_Elab_Tactic_Do_elabInvariants___closed__20; LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__6_spec__8___redArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__8___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_tryGoal_spec__0(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_MessageData_hasTag(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Range_Basic_0__Std_Legacy_Range_forIn_x27_loop___at___00Lean_Elab_Tactic_Do_VCGen_genVCs_spec__1___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGenHint___lam__0___closed__2; static lean_object* l_List_mapTR_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__8___closed__1; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); @@ -753,13 +781,14 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tact static lean_object* l_Lean_Elab_Tactic_Do_elabInvariants___closed__17; lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntroForall___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__4___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__12_spec__16_spec__18_spec__23___redArg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__1_spec__2_spec__6(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15___closed__6; +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__9___boxed(lean_object**); uint8_t l_Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__28_spec__33_spec__38___redArg___closed__0; static lean_object* l_Lean_Elab_Tactic_Do_VCGen_genVCs___lam__0___closed__10; @@ -767,15 +796,14 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_sortMVarIdArrayByIndex___at___00Lean lean_object* l_Lean_mkConst(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addTrace___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__5___lam__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__42___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__6_spec__8___redArg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__10_spec__15(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15___closed__1; -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44(lean_object*, lean_object*, lean_object*, 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_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__19___redArg___lam__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13_spec__28_spec__36(size_t, size_t, lean_object*); +static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___closed__1; LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___00Lean_Elab_Tactic_Do_ProofMode_mIntro___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__3_spec__3___redArg___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___regBuiltin_Lean_Elab_Tactic_Do_elabMVCGen__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__32___lam__0(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___lam__9___closed__0; @@ -785,60 +813,55 @@ static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___lam__7___closed__0; static lean_object* l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13___redArg___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntroForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at___00Lean_Elab_Tactic_Do_VCGen_genVCs_spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__6___boxed(lean_object**); +static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__3; lean_object* l_Lean_Expr_constName_x21(lean_object*); static lean_object* l___private_Init_Data_Array_Basic_0__Array_foldrMUnsafe_fold___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__10___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__3(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_Basic_0__Array_foldlMUnsafe_fold___at___00Array_filterMapM___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__13_spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___closed__2; extern lean_object* l_Lean_instInhabitedExpr; lean_object* l_Lean_KeyedDeclsAttribute_addBuiltin___redArg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_elabVCs___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__11___redArg(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__0; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__46___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_VCGen_genVCs___lam__0___closed__8; LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at___00Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__6_spec__8_spec__12_spec__34(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__11; -LEAN_EXPORT lean_object* l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__54___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__17(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___lam__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___lam__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_String_dropPrefix_x3f___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__8___redArg___closed__0; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_tryGoal___lam__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47(uint8_t, 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_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldrMUnsafe_fold___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__6_spec__8___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentArray_push___redArg(lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__5___redArg___closed__3; lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_lambdaTelescopeImp(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__46(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGenHint___lam__0___closed__5; -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__0___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__50(lean_object*, lean_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_Do_elabMVCGenHint___regBuiltin_Lean_Elab_Tactic_Do_elabMVCGenHint__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__2; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___closed__7; static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___lam__4___closed__0; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Range_Basic_0__Std_Legacy_Range_forIn_x27_loop___at___00Lean_Elab_Tactic_Do_VCGen_genVCs_spec__1___redArg___closed__0; -static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__5; LEAN_EXPORT lean_object* l_Lean_Option_get___at___00Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13_spec__30(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__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___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__47(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_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__16___closed__18; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___lam__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_letE___override(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); static lean_object* l___private_Init_Data_Array_Basic_0__Array_foldrMUnsafe_fold___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__10___closed__2; lean_object* l_Lean_Meta_mapErrorImp___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -846,6 +869,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__0(lean_object*, lean_object* l_Lean_Meta_mkLetFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__5___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__11(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal___closed__1; static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___closed__9; @@ -853,7 +877,6 @@ lean_object* l_Lean_Elab_Tactic_pruneSolvedGoals(lean_object*, lean_object*, lea uint8_t l_Lean_Elab_Tactic_Do_isJP(lean_object*); lean_object* l_Lean_MVarId_setTag___redArg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__14(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addMessageContextFull___at___00Lean_throwError___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__0_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_Do_ProofMode_TypeList_length(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_VCGen_genVCs___lam__0___closed__1; @@ -865,8 +888,8 @@ static lean_object* l_Lean_Elab_Tactic_Do_elabInvariants___closed__13; static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___regBuiltin_Lean_Elab_Tactic_Do_elabMVCGen__1___closed__0; static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___closed__11; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13_spec__28___lam__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_warningAsError; -static lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___closed__0; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__28_spec__33_spec__38_spec__42___redArg___lam__0(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__7___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabInvariants___closed__15; @@ -874,15 +897,16 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___0 static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__16___closed__3; static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__5___closed__14; static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__19___redArg___lam__4___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__4; lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___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_MVarId_isAssigned___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__0(lean_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_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_QSort_Basic_0__Array_qsort_sort___at___00Lean_Elab_Tactic_sortMVarIdArrayByIndex___at___00Lean_Elab_Tactic_sortMVarIdsByIndex___at___00Lean_Elab_Tactic_collectFreshMVars___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__12_spec__24_spec__31_spec__47___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__3___closed__1; -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___lam__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_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__7; static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntro___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__3___closed__1; lean_object* l_Lean_Elab_Tactic_Do_withJP___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__9(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -893,12 +917,12 @@ static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___closed__1; extern lean_object* l_Lean_trace_profiler_threshold; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___boxed(lean_object**); lean_object* l_Lean_mkNot(lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_VCGen_genVCs___lam__0___closed__6; -static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__10; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__28_spec__33___boxed(lean_object*, lean_object*, lean_object*, 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_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__16___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__46___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_extract___redArg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at___00Lean_Elab_Tactic_Do_VCGen_genVCs_spec__2___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_collectFreshMVars___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__12___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -909,13 +933,11 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00__private_Lean_Elab_Tacti LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___redArg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15___closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__11(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__50___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Subarray_toArray___redArg(lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeMGoal___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__21___lam__2___closed__1; static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntroForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__11___lam__0___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___lam__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__1; +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__0; LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__9_spec__13___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_Option_get___at___00Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13_spec__27___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___00Lean_Elab_Tactic_Do_ProofMode_mIntro___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__3_spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -923,21 +945,28 @@ static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Ela lean_object* l_Lean_Meta_Match_forallAltVarsTelescope___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_appFn_x21(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_syntax_ident(lean_object*); +LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mIntroForall___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__4_spec__5_spec__9___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___00Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__12_spec__16_spec__18_spec__23_spec__32___redArg(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__0(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___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__32(lean_object*, uint8_t, uint8_t, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15___lam__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_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_lambdaBoundedTelescope___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__15___redArg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__6; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__28_spec__33_spec__38_spec__42___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__12_spec__16___redArg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13___redArg___lam__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_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57_spec__65___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__14(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__3; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__2(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__5(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_tryGoal_spec__1___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_logWarningAt___at___00Lean_Elab_Tactic_Do_elabMVCGen_spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntroForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__11___lam__0___closed__4; @@ -945,10 +974,9 @@ static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntroForall___at___00__priv LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__9___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at___00Lean_Elab_Tactic_Do_elabMVCGen_spec__3(lean_object*, lean_object*); +static lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntro___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__3___closed__3; -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59(lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_borrowed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__2(uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13_spec__29___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -958,20 +986,17 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tact LEAN_EXPORT lean_object* l___private_Init_Data_Range_Basic_0__Std_Legacy_Range_forIn_x27_loop___at___00Lean_Elab_Tactic_Do_VCGen_genVCs_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeMGoal___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__21___lam__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___lam__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__10; lean_object* lean_usize_to_nat(size_t); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__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___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__19___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___lam__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___lam__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Core_mkFreshUserName(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56_spec__64___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_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___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__3___boxed(lean_object**); static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__5___closed__8; lean_object* l_String_Slice_pos_x21(lean_object*, lean_object*); static lean_object* l___private_Init_Data_Array_Basic_0__Array_foldrMUnsafe_fold___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__10___closed__4; -static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___closed__0; static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__19___redArg___lam__4___closed__6; -static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__12_spec__16_spec__18_spec__24___redArg(size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__19(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -983,47 +1008,42 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tact static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___closed__8; lean_object* l_Lean_Meta_mkAndIntroN(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofExpr(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_Do_SpecAttr_SpecProof_instantiate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__3; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__21(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__57(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__5; lean_object* l_Lean_FVarId_getUserName___redArg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntroForall___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__4___lam__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__28_spec__33_spec__38___redArg___closed__5; -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_tryGoal___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_LocalDecl_isImplementationDetail(lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at___00Lean_Elab_Tactic_Do_elabMVCGen_spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___closed__12; lean_object* l_Lean_mkLambda(lean_object*, uint8_t, lean_object*, lean_object*); -static lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__55___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeMGoal___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__21___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppNumArgs(lean_object*); +static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__6; lean_object* l_Lean_LocalDecl_userName(lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__8(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__28_spec__33_spec__38___redArg___closed__2; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg(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_mkStr6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__8(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Context_config(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__55___redArg___lam__0(lean_object*, lean_object*, lean_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_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_Do_ProofMode_getFreshHypName(lean_object*, lean_object*, lean_object*); lean_object* lean_name_append_index_after(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntro___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__3___lam__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_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__3___closed__1; lean_object* l_Lean_Elab_Tactic_Do_ProofMode_SPred_mkPure(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__45___lam__0___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_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__28_spec__33___redArg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__5___closed__6; @@ -1044,45 +1064,35 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tact static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_patchVCAltIntoCaseTactic___closed__0; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__8(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_filterMapM___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__13___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__4; -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__8___closed__6; LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlMAux___at___00Lean_PersistentArray_foldlM___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__1_spec__4(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15___lam__2(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___lam__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48(uint8_t, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__9___redArg___closed__0; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_elabVCs_applyPreTac_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__15___closed__5; -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___closed__2; lean_object* lean_array_fget(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__0; -LEAN_EXPORT lean_object* l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__2; lean_object* l_Lean_Expr_constLevels_x21(lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__17___closed__0; -static lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__0___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__15___closed__21; static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___closed__2; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__50(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Subarray_empty(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__32___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_tryGoal(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__1; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__7___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mIntroForall___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__4_spec__5___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at___00Lean_Elab_Tactic_Do_VCGen_genVCs_spec__2___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__1; lean_object* l_Lean_Expr_appFnCleanup___redArg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__1(uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntroForall___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__4___closed__0; +LEAN_EXPORT lean_object* l_Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__15___boxed__const__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntro___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__1_spec__3(lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60(lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___closed__10; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal(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_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__19___redArg___lam__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1092,52 +1102,59 @@ static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_ lean_object* l_Lean_Meta_instMonadMetaM___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__16___closed__12; static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__61___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__12___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_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_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_Do_elabMVCGenHint___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__19___redArg___lam__4___closed__7; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__2___boxed(lean_object**); static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___closed__4; -LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntroForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__11___lam__0___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___00Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__12_spec__16_spec__18_spec__23_spec__32(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__15___closed__2; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_instMonadMetaM___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntroForall___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__4___closed__3; -static lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__4; static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__5___closed__4; -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___boxed__const__1; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__61___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15___lam__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkApp3(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___closed__13; -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__6(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__9; static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeMGoal___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__21___closed__0; +LEAN_EXPORT lean_object* l_Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_elabMVCGenHint___regBuiltin_Lean_Elab_Tactic_Do_elabMVCGenHint__1(); lean_object* l_Lean_PersistentHashMap_mkCollisionNode___redArg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__5___closed__2; +static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntro___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__3___lam__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_decLevel(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2(uint8_t, uint8_t, lean_object*); lean_object* l_Lean_Elab_Tactic_Do_emitVC(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__4; LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__23(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabInvariants___closed__21; lean_object* l_Lean_Elab_Tactic_withMainContext___redArg(lean_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_Do_elabInvariants___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___closed__0; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); uint64_t l_Lean_instHashableMVarId_hash(lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__49(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_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13_spec__26___lam__1___closed__2; lean_object* l_Lean_Meta_mkFreshExprMVar(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__0(lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13___redArg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__18(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__12_spec__16_spec__18_spec__24(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Option_get___at___00Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13_spec__30___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__28_spec__33_spec__38_spec__42___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__9___redArg___closed__21; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_tryGoal___lam__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1145,12 +1162,10 @@ static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___regBuiltin_Lean_Elab_Tact static lean_object* l_Lean_Elab_Tactic_Do_elabInvariants___closed__2; static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__15___closed__11; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_collectFreshMVars___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__12___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__0; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntro___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__3___lam__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__12_spec__16_spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_TermElabM_run___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___closed__5; -LEAN_EXPORT lean_object* l_Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_isEqvAux___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__9___redArg(lean_object*, lean_object*, lean_object*); @@ -1159,25 +1174,27 @@ static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_ LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__9(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__5___closed__0; static lean_object* l_Lean_MVarId_isAssigned___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__6___closed__0; -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__45___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__16___boxed(lean_object**); static lean_object* l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__12_spec__16_spec__18___redArg___closed__2; lean_object* l_Lean_PersistentArray_toArray___redArg(lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabVCs___closed__0; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mIntroForall___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__4_spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkEqRefl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__9___redArg___closed__7; static double l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13___redArg___closed__2; lean_object* l_Lean_indentExpr(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___lam__7(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__49___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isHEq(lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_getMVarsNoDelayed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__63___lam__1___closed__0; lean_object* l_Lean_Meta_isExprDefEqGuarded(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__58(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__4___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__6(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__28(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1185,6 +1202,7 @@ LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at___00__private_Lean_Elab_Tac LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__11___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__1; static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___lam__4___closed__1; lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___closed__6; @@ -1195,24 +1213,24 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___0 LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___lam__7(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); lean_object* l_Lean_Meta_mkForallFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__49___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___closed__2; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__27___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___closed__8; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__6(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__28_spec__32___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mIntroForall___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__4_spec__5_spec__9___redArg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn_x27(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__14___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15___lam__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___lam__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___closed__3; LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__0___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_logErrorAt___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__1_spec__1(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_MatcherApp_toExpr(lean_object*); -static lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__3; static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntroForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__11___lam__0___closed__6; -static lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___closed__1; lean_object* l_Lean_LocalDecl_type(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__19___redArg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13_spec__28_spec__36___boxed(lean_object*, lean_object*, lean_object*); @@ -1224,31 +1242,27 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tacti static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__19___redArg___lam__4___closed__3; static lean_object* l___private_Init_Data_Array_Basic_0__Array_foldrMUnsafe_fold___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__10___closed__5; static lean_object* l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__28_spec__33___redArg___closed__0; -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__9___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__2___boxed(lean_object**); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__1_spec__2_spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___closed__1; static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__16___closed__14; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_tryGoal___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__2___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_reverse___redArg(lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at___00Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__6_spec__8_spec__12_spec__34___redArg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___lam__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__5___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__2___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_get_match_equations_for(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntroForall___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__16___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_tryGoal_spec__0___redArg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__4___boxed(lean_object**); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__16___closed__4; lean_object* l_Lean_Elab_Tactic_Do_ProofMode_addHypInfo(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); @@ -1256,13 +1270,13 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLoca lean_object* l_Array_append___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___lam__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13___redArg___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_tryGoal_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabInvariants___closed__18; lean_object* l_ReaderT_instApplicativeOfMonad___redArg___lam__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__46___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__27___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__9___redArg___closed__17; +static lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__1; LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__3___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__12_spec__16_spec__18_spec__24___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13_spec__26___lam__1___closed__1; @@ -1273,25 +1287,22 @@ uint64_t lean_uint64_shift_left(uint64_t, uint64_t); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_tryGoal_spec__1___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__8___closed__1; -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_logInfoAt___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__19___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__43___lam__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_mkArrow(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__12___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___lam__7___closed__1; -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___closed__4; static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_patchVCAltIntoCaseTactic___closed__2; lean_object* l_Lean_toMessageList(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__2; static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__1___closed__2; uint8_t lean_string_memcmp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__9___redArg___closed__8; lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_throwTypeMismatchError___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__12(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___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__61___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_LocalDecl_binderInfo(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__1_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__16___lam__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1305,47 +1316,40 @@ LEAN_EXPORT lean_object* l___private_Init_Data_Array_QSort_Basic_0__Array_qsort_ LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00Lean_Elab_Tactic_Do_elabMVCGen_spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__3(uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__12___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__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_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__61(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_erase_macro_scopes(lean_object*); LEAN_EXPORT lean_object* l_String_dropPrefix_x3f___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__8(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__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_filterMapM___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__13(lean_object*, lean_object*, lean_object*); -static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntro___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__3___lam__0(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__8___closed__0; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___lam__9(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_tryGoal_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___closed__0; +lean_object* lean_expr_instantiate_rev(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__6; static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__9___redArg___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___lam__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Simp_mkCongrArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_MatcherApp_altNumParams(lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__5___closed__9; -static lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62___lam__1___closed__0; -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_isEqvAux___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___lam__0(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_sub(size_t, size_t); +static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__2; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__28_spec__32(size_t, size_t, lean_object*); lean_object* lean_array_mk(lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_lambdaBoundedTelescope___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__15___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___closed__0; static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___closed__7; -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__14___closed__3; -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__14___closed__1; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_elabVCs_applyPreTac_spec__0(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__10(lean_object*, lean_object*); lean_object* l_Lean_instantiateMVarsCore(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__11___redArg___lam__0(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___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___redArg(lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___boxed(lean_object**); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13_spec__26___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__11___closed__2; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Array_filterMapM___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__13_spec__18(lean_object*, size_t, size_t, lean_object*); @@ -1354,50 +1358,56 @@ static lean_object* l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_ LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__5_spec__7___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__3___redArg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__16___closed__17; +static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__3; LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__23___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__16___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__56___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_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_Do_VCGen_genVCs___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__50___redArg(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_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__46___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__7; size_t lean_usize_add(size_t, size_t); lean_object* l_mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__3(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15___closed__7; -static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__12; LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_throwErrorAt___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__7_spec__10___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__8; +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15___lam__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__58___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__13(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_forallTelescopeReducingAux(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_elabVCs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__16___closed__8; -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_throwErrorAt___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__7_spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkMVar(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52(lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_elabMVCGenHint___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___closed__8; static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___closed__6; -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_betaRev(lean_object*, lean_object*, uint8_t, uint8_t); static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___closed__2; lean_object* lean_array_uget(lean_object*, size_t); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___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___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__2(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_array_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__19___redArg___lam__4___closed__2; -static lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__5; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__16(lean_object*, uint8_t, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__47___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___closed__0; LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__28_spec__33_spec__38(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_trace_profiler; +LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__49___redArg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__63___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Range_Basic_0__Std_Legacy_Range_forIn_x27_loop___at___00Lean_Elab_Tactic_Do_VCGen_genVCs_spec__0___redArg___closed__0; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__9_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___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkFVar(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__42___lam__0(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__0___closed__0; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal___lam__0(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___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__4(lean_object*, lean_object*); @@ -1408,7 +1418,6 @@ LEAN_EXPORT lean_object* l___private_Init_Data_Range_Basic_0__Std_Legacy_Range_f LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_elabInvariants(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__8___closed__7; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___closed__13; static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15___closed__4; LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__15___redArg(lean_object*, lean_object*); @@ -1419,43 +1428,38 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_Elab_Tactic_Do_ProofMod lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_local_ctx_num_indices(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___boxed(lean_object**); static lean_object* l_Lean_Elab_Tactic_Do_VCGen_genVCs___lam__0___closed__3; +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__7; LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__28_spec__33_spec__38___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___closed__7; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__0___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_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___closed__0; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__28_spec__33(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___closed__11; lean_object* l_Lean_Meta_Simp_simp___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___closed__4; -static lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__0___closed__0; -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_Do_withLetDeclShared___redArg(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntroForall___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__4___lam__1___closed__1; static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___closed__0; static lean_object* l_Array_mapFinIdxM_map___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__23___redArg___closed__0; lean_object* l_Lean_Elab_Tactic_evalTactic(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__15___closed__7; -static lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59___lam__1___closed__0; +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_setKind(lean_object*, lean_object*); lean_object* l_Lean_Meta_Match_MatcherInfo_getNumDiscrEqs(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__48___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldrMUnsafe_fold___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__10___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__5___redArg___closed__1; static lean_object* l_Lean_Elab_Tactic_Do_elabVCs___closed__3; LEAN_EXPORT lean_object* l_Lean_throwError___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__0___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntro___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__3___lam__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__55___redArg(lean_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_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___lam__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__1(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_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__6___redArg(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_collectFreshMVars___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__12___lam__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___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__45___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___closed__5; lean_object* l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1466,48 +1470,58 @@ lean_object* l_Lean_Expr_headBeta(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntro___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__3___lam__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___lam__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__5___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__8; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__14(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__14___closed__0; lean_object* l_Pi_instInhabited___redArg(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_value_x3f(lean_object*, uint8_t); lean_object* lean_array_get_size(lean_object*); +static lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__0; lean_object* l_Lean_LocalDecl_toExpr(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___lam__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_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__2; lean_object* l_Lean_Elab_Tactic_setGoals___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__1_spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__5___closed__1; LEAN_EXPORT lean_object* l_Lean_logErrorAt___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__12; lean_object* l_Lean_Elab_Tactic_Do_ProofMode_pushForallContextIntoHyps(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Range_Basic_0__Std_Legacy_Range_forIn_x27_loop___at___00Lean_Elab_Tactic_Do_VCGen_genVCs_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__3; lean_object* l_Lean_mkAtom(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal___lam__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___closed__10; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeMGoal___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__21___lam__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_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Range_Basic_0__Std_Legacy_Range_forIn_x27_loop___at___00Lean_Elab_Tactic_Do_VCGen_genVCs_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_tryGoal___lam__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__0; static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___closed__3; static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal___closed__0; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__3___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62___lam__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_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__28_spec__33_spec__38_spec__42(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntroForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__9___redArg___closed__2; +static lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__5; uint8_t lean_usize_dec_lt(size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__55(lean_object*, lean_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_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___boxed(lean_object**); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__6___boxed(lean_object**); +LEAN_EXPORT lean_object* l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__55___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_simpDiscrs_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___lam__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkLambdaFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___lam__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__46(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__14___closed__0; static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__9___redArg___closed__5; uint64_t l_Lean_Meta_TransparencyMode_toUInt64(uint8_t); @@ -1524,107 +1538,114 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_Elab_Tactic_Do_ProofMod LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__17(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__10_spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__2; LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_run___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_elabMVCGenHint___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15___closed__5; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__6___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___lam__4___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60___lam__0___boxed(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_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__50___lam__0___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_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isConst(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__15___closed__14; uint8_t l_Lean_Exception_isRuntime(lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__16___closed__13; -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabInvariants___closed__10; lean_object* l_Lean_Expr_beta(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkApp5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22___boxed(lean_object**); -static lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static size_t l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__12_spec__16_spec__18___redArg___closed__1; +uint8_t l_Lean_Expr_isFVar(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15___lam__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13(lean_object*, lean_object*, lean_object*, uint8_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_patchVCAltIntoCaseTactic(lean_object*); lean_object* l_Lean_Meta_instantiateForall(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabInvariants___closed__8; LEAN_EXPORT lean_object* l_String_dropPrefix_x3f___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__8___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__3___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__0___closed__0; -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__16___closed__0; lean_object* l_Lean_Expr_mvarId_x21(lean_object*); -LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_tryGoal_spec__1___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__48___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__1___closed__3; static lean_object* l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13___redArg___closed__3; -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__2(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t l___private_Lean_Meta_Basic_0__Lean_Meta_Config_toKey(lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__15___closed__0; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__63___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_consumeMData(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_lambdaBoundedTelescope___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__15(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___boxed(lean_object**); +static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__6; static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__5___closed__3; LEAN_EXPORT lean_object* l_Lean_PersistentArray_foldlM___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mIntro___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__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___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_Do_suggestInvariant___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__58___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__5___redArg___closed__0; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__6___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabMVCGenHint___regBuiltin_Lean_Elab_Tactic_Do_elabMVCGenHint__1___closed__2; lean_object* l_Array_zip___redArg(lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__12_spec__16_spec__18(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__10(lean_object*, lean_object*); uint8_t l_Lean_isCasesOnRecursor(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_isEqvAux___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__55(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__1_spec__2(lean_object*, size_t, size_t, lean_object*); lean_object* lean_expr_instantiate1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___regBuiltin_Lean_Elab_Tactic_Do_elabMVCGen__1___boxed(lean_object*); static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__28_spec__33_spec__38___redArg___closed__1; LEAN_EXPORT uint8_t l___private_Init_Data_Array_QSort_Basic_0__Array_qsort_sort___at___00Lean_Elab_Tactic_sortMVarIdArrayByIndex___at___00Lean_Elab_Tactic_sortMVarIdsByIndex___at___00Lean_Elab_Tactic_collectFreshMVars___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__12_spec__24_spec__31_spec__47___redArg___lam__0(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_Do_ProofMode_transferHypNames(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__63___lam__1___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__6_spec__8_spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mIntroForall___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__4_spec__5___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59___lam__1___closed__1; LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__4_spec__5___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__0; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_elabVCs_applyPreTac(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__19___redArg___lam__4___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onFail___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_Elab_Tactic_Do_ProofMode_mIntro___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__3___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__6_spec__8_spec__12(lean_object*, lean_object*, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__3___boxed(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_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__19_spec__24___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59___lam__0(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__15___closed__12; lean_object* l_Lean_Elab_Tactic_Do_ProofMode_emptyHyp(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15___lam__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__16(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__8; +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13___redArg___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__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_MessageLog_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__57___lam__0(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_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg(lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_setKind___redArg(lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13_spec__26___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___lam__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_patchVCAltIntoCaseTactic___closed__1; lean_object* l_Lean_Meta_whnfR(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_Do_getSplitInfo_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__16___closed__15; LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1636,11 +1657,14 @@ static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_l LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___lam__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); double lean_float_sub(double, double); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13___redArg___lam__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___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__51(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__9___closed__2; lean_object* l_Lean_Expr_getNumHeadLambdas(lean_object*); -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_elabInvariants___closed__14; LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__17___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__56___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__15___redArg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* _init_l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_ProofMode_MGoal_withNewProg___closed__0() { @@ -8036,16 +8060,16 @@ x_69 = lean_box(x_2); x_70 = lean_box(x_67); x_71 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__11___boxed), 13, 5); lean_closure_set(x_71, 0, x_3); -lean_closure_set(x_71, 1, x_65); +lean_closure_set(x_71, 1, x_62); lean_closure_set(x_71, 2, x_68); lean_closure_set(x_71, 3, x_69); lean_closure_set(x_71, 4, x_70); lean_inc(x_60); -lean_inc_ref(x_63); -lean_inc(x_64); lean_inc_ref(x_59); -lean_inc_ref(x_62); -x_72 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_71, x_62, x_58, x_59, x_64, x_63, x_60); +lean_inc(x_57); +lean_inc_ref(x_63); +lean_inc_ref(x_64); +x_72 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_71, x_64, x_58, x_63, x_57, x_59, x_60); if (lean_obj_tag(x_72) == 0) { lean_object* x_73; @@ -8054,11 +8078,11 @@ lean_inc(x_73); lean_dec_ref(x_72); x_22 = x_61; x_23 = x_73; -x_24 = x_62; +x_24 = x_64; x_25 = x_58; -x_26 = x_59; -x_27 = x_64; -x_28 = x_63; +x_26 = x_63; +x_27 = x_57; +x_28 = x_59; x_29 = x_60; x_30 = lean_box(0); goto block_56; @@ -8066,12 +8090,12 @@ goto block_56; else { uint8_t x_74; -lean_dec(x_64); +lean_dec_ref(x_64); lean_dec_ref(x_63); -lean_dec_ref(x_62); lean_dec_ref(x_61); lean_dec(x_60); lean_dec_ref(x_59); +lean_dec(x_57); lean_dec_ref(x_13); x_74 = !lean_is_exclusive(x_72); if (x_74 == 0) @@ -8106,15 +8130,15 @@ if (x_93 == 0) lean_dec_ref(x_90); lean_dec(x_6); lean_dec(x_4); -x_57 = lean_box(0); +x_57 = x_80; x_58 = x_83; -x_59 = x_78; +x_59 = x_81; x_60 = x_79; x_61 = x_85; -x_62 = x_82; -x_63 = x_80; -x_64 = x_81; -x_65 = x_84; +x_62 = x_84; +x_63 = x_82; +x_64 = x_78; +x_65 = lean_box(0); x_66 = x_93; goto block_77; } @@ -8127,15 +8151,15 @@ if (x_94 == 0) { lean_dec(x_6); lean_dec(x_4); -x_57 = lean_box(0); +x_57 = x_80; x_58 = x_83; -x_59 = x_78; +x_59 = x_81; x_60 = x_79; x_61 = x_85; -x_62 = x_82; -x_63 = x_80; -x_64 = x_81; -x_65 = x_84; +x_62 = x_84; +x_63 = x_82; +x_64 = x_78; +x_65 = lean_box(0); x_66 = x_94; goto block_77; } @@ -8149,11 +8173,11 @@ x_97 = l_Lean_mkConst(x_96, x_4); x_98 = l_Lean_mkLambda(x_6, x_95, x_97, x_84); x_22 = x_85; x_23 = x_98; -x_24 = x_82; +x_24 = x_78; x_25 = x_83; -x_26 = x_78; -x_27 = x_81; -x_28 = x_80; +x_26 = x_82; +x_27 = x_80; +x_28 = x_81; x_29 = x_79; x_30 = lean_box(0); goto block_56; @@ -8221,11 +8245,11 @@ if (x_121 == 0) { lean_dec_ref(x_10); lean_dec(x_1); -x_78 = x_102; +x_78 = x_100; x_79 = x_105; -x_80 = x_104; -x_81 = x_103; -x_82 = x_100; +x_80 = x_103; +x_81 = x_104; +x_82 = x_102; x_83 = x_101; x_84 = x_119; x_85 = x_116; @@ -8263,11 +8287,11 @@ lean_inc(x_128); x_129 = lean_ctor_get(x_127, 1); lean_inc(x_129); lean_dec(x_127); -x_78 = x_102; +x_78 = x_100; x_79 = x_105; -x_80 = x_104; -x_81 = x_103; -x_82 = x_100; +x_80 = x_103; +x_81 = x_104; +x_82 = x_102; x_83 = x_101; x_84 = x_128; x_85 = x_129; @@ -9034,8 +9058,8 @@ lean_inc_ref(x_70); lean_inc(x_69); lean_inc_ref(x_68); lean_inc_ref(x_66); -lean_inc(x_65); -x_75 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__3(x_65, x_66, x_67, x_68, x_69, x_70, x_71); +lean_inc(x_64); +x_75 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__3(x_64, x_66, x_67, x_68, x_69, x_70, x_71); if (lean_obj_tag(x_75) == 0) { lean_object* x_76; lean_object* x_77; uint8_t x_78; @@ -9049,8 +9073,8 @@ x_78 = lean_unbox(x_76); lean_dec(x_76); if (x_78 == 0) { -lean_dec(x_65); -lean_dec_ref(x_64); +lean_dec_ref(x_65); +lean_dec(x_64); x_35 = x_77; x_36 = x_66; x_37 = x_67; @@ -9069,7 +9093,7 @@ lean_inc_ref(x_70); lean_inc(x_69); lean_inc_ref(x_68); lean_inc_ref(x_66); -x_79 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_64, x_66, x_67, x_68, x_69, x_70, x_71); +x_79 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_65, x_66, x_67, x_68, x_69, x_70, x_71); if (lean_obj_tag(x_79) == 0) { lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; @@ -9086,7 +9110,7 @@ lean_inc_ref(x_70); lean_inc(x_69); lean_inc_ref(x_68); lean_inc_ref(x_66); -x_84 = l_Lean_addTrace___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__5(x_65, x_83, x_66, x_67, x_68, x_69, x_70, x_71); +x_84 = l_Lean_addTrace___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__5(x_64, x_83, x_66, x_67, x_68, x_69, x_70, x_71); if (lean_obj_tag(x_84) == 0) { lean_dec_ref(x_84); @@ -9143,7 +9167,7 @@ lean_dec(x_69); lean_dec_ref(x_68); lean_dec(x_67); lean_dec_ref(x_66); -lean_dec(x_65); +lean_dec(x_64); lean_dec_ref(x_34); lean_dec_ref(x_30); lean_dec_ref(x_14); @@ -9165,8 +9189,8 @@ lean_dec(x_69); lean_dec_ref(x_68); lean_dec(x_67); lean_dec_ref(x_66); -lean_dec(x_65); -lean_dec_ref(x_64); +lean_dec_ref(x_65); +lean_dec(x_64); lean_dec_ref(x_63); lean_dec_ref(x_34); lean_dec_ref(x_30); @@ -9201,8 +9225,8 @@ lean_dec(x_69); lean_dec_ref(x_68); lean_dec(x_67); lean_dec_ref(x_66); -lean_dec(x_65); -lean_dec_ref(x_64); +lean_dec_ref(x_65); +lean_dec(x_64); lean_dec_ref(x_63); lean_dec_ref(x_34); lean_dec_ref(x_30); @@ -9254,7 +9278,7 @@ lean_inc(x_104); lean_inc_ref(x_103); lean_inc(x_102); lean_inc_ref(x_101); -x_112 = l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__11___redArg(x_110, x_108, x_99, x_111, x_111, x_101, x_102, x_103, x_104, x_105, x_106); +x_112 = l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__11___redArg(x_110, x_108, x_100, x_111, x_111, x_101, x_102, x_103, x_104, x_105, x_106); if (lean_obj_tag(x_112) == 0) { lean_object* x_113; uint8_t x_114; @@ -9305,8 +9329,8 @@ lean_inc_ref(x_105); lean_inc(x_104); lean_inc_ref(x_103); lean_inc_ref(x_101); -lean_inc(x_100); -x_123 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__3(x_100, x_101, x_102, x_103, x_104, x_105, x_106); +lean_inc(x_99); +x_123 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__3(x_99, x_101, x_102, x_103, x_104, x_105, x_106); if (lean_obj_tag(x_123) == 0) { lean_object* x_124; lean_object* x_125; lean_object* x_126; uint8_t x_127; @@ -9327,8 +9351,8 @@ lean_dec(x_115); lean_dec_ref(x_27); x_62 = x_125; x_63 = x_126; -x_64 = x_97; -x_65 = x_100; +x_64 = x_99; +x_65 = x_97; x_66 = x_101; x_67 = x_102; x_68 = x_103; @@ -9387,15 +9411,15 @@ lean_inc_ref(x_105); lean_inc(x_104); lean_inc_ref(x_103); lean_inc_ref(x_101); -lean_inc(x_100); -x_143 = l_Lean_addTrace___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__5(x_100, x_142, x_101, x_102, x_103, x_104, x_105, x_106); +lean_inc(x_99); +x_143 = l_Lean_addTrace___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__5(x_99, x_142, x_101, x_102, x_103, x_104, x_105, x_106); if (lean_obj_tag(x_143) == 0) { lean_dec_ref(x_143); x_62 = x_125; x_63 = x_126; -x_64 = x_97; -x_65 = x_100; +x_64 = x_99; +x_65 = x_97; x_66 = x_101; x_67 = x_102; x_68 = x_103; @@ -9416,7 +9440,7 @@ lean_dec(x_104); lean_dec_ref(x_103); lean_dec(x_102); lean_dec_ref(x_101); -lean_dec(x_100); +lean_dec(x_99); lean_dec_ref(x_97); lean_dec_ref(x_34); lean_dec_ref(x_30); @@ -9454,7 +9478,7 @@ lean_dec(x_104); lean_dec_ref(x_103); lean_dec(x_102); lean_dec_ref(x_101); -lean_dec(x_100); +lean_dec(x_99); lean_dec_ref(x_97); lean_dec_ref(x_34); lean_dec_ref(x_30); @@ -9481,7 +9505,7 @@ lean_dec(x_104); lean_dec_ref(x_103); lean_dec(x_102); lean_dec_ref(x_101); -lean_dec(x_100); +lean_dec(x_99); lean_dec_ref(x_97); lean_dec_ref(x_34); lean_dec_ref(x_30); @@ -9514,7 +9538,7 @@ else lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_dec(x_116); lean_dec(x_115); -lean_dec(x_100); +lean_dec(x_99); lean_dec_ref(x_97); lean_dec_ref(x_34); lean_dec_ref(x_30); @@ -9550,7 +9574,7 @@ lean_dec(x_104); lean_dec_ref(x_103); lean_dec(x_102); lean_dec_ref(x_101); -lean_dec(x_100); +lean_dec(x_99); lean_dec_ref(x_97); lean_dec_ref(x_34); lean_dec_ref(x_30); @@ -9576,7 +9600,7 @@ lean_dec(x_104); lean_dec_ref(x_103); lean_dec(x_102); lean_dec_ref(x_101); -lean_dec(x_100); +lean_dec(x_99); lean_dec_ref(x_97); lean_dec_ref(x_34); lean_dec_ref(x_30); @@ -9650,8 +9674,8 @@ lean_inc_ref(x_105); lean_inc(x_104); lean_inc_ref(x_103); lean_inc_ref(x_101); -lean_inc(x_100); -x_166 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__3(x_100, x_101, x_102, x_103, x_104, x_105, x_106); +lean_inc(x_99); +x_166 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__3(x_99, x_101, x_102, x_103, x_104, x_105, x_106); if (lean_obj_tag(x_166) == 0) { lean_object* x_167; lean_object* x_168; lean_object* x_169; uint8_t x_170; @@ -9671,8 +9695,8 @@ lean_dec(x_158); lean_dec_ref(x_27); x_62 = x_168; x_63 = x_169; -x_64 = x_97; -x_65 = x_100; +x_64 = x_99; +x_65 = x_97; x_66 = x_101; x_67 = x_102; x_68 = x_103; @@ -9731,15 +9755,15 @@ lean_inc_ref(x_105); lean_inc(x_104); lean_inc_ref(x_103); lean_inc_ref(x_101); -lean_inc(x_100); -x_187 = l_Lean_addTrace___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__5(x_100, x_186, x_101, x_102, x_103, x_104, x_105, x_106); +lean_inc(x_99); +x_187 = l_Lean_addTrace___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__5(x_99, x_186, x_101, x_102, x_103, x_104, x_105, x_106); if (lean_obj_tag(x_187) == 0) { lean_dec_ref(x_187); x_62 = x_168; x_63 = x_169; -x_64 = x_97; -x_65 = x_100; +x_64 = x_99; +x_65 = x_97; x_66 = x_101; x_67 = x_102; x_68 = x_103; @@ -9760,7 +9784,7 @@ lean_dec(x_104); lean_dec_ref(x_103); lean_dec(x_102); lean_dec_ref(x_101); -lean_dec(x_100); +lean_dec(x_99); lean_dec_ref(x_97); lean_dec_ref(x_34); lean_dec_ref(x_30); @@ -9798,7 +9822,7 @@ lean_dec(x_104); lean_dec_ref(x_103); lean_dec(x_102); lean_dec_ref(x_101); -lean_dec(x_100); +lean_dec(x_99); lean_dec_ref(x_97); lean_dec_ref(x_34); lean_dec_ref(x_30); @@ -9824,7 +9848,7 @@ lean_dec(x_104); lean_dec_ref(x_103); lean_dec(x_102); lean_dec_ref(x_101); -lean_dec(x_100); +lean_dec(x_99); lean_dec_ref(x_97); lean_dec_ref(x_34); lean_dec_ref(x_30); @@ -9858,7 +9882,7 @@ else lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_dec(x_159); lean_dec(x_158); -lean_dec(x_100); +lean_dec(x_99); lean_dec_ref(x_97); lean_dec_ref(x_34); lean_dec_ref(x_30); @@ -9893,7 +9917,7 @@ lean_dec(x_104); lean_dec_ref(x_103); lean_dec(x_102); lean_dec_ref(x_101); -lean_dec(x_100); +lean_dec(x_99); lean_dec_ref(x_97); lean_dec_ref(x_34); lean_dec_ref(x_30); @@ -9918,7 +9942,7 @@ lean_dec(x_104); lean_dec_ref(x_103); lean_dec(x_102); lean_dec_ref(x_101); -lean_dec(x_100); +lean_dec(x_99); lean_dec_ref(x_97); lean_dec_ref(x_34); lean_dec_ref(x_30); @@ -9957,7 +9981,7 @@ lean_dec(x_104); lean_dec_ref(x_103); lean_dec(x_102); lean_dec_ref(x_101); -lean_dec(x_100); +lean_dec(x_99); lean_dec_ref(x_97); lean_dec_ref(x_34); lean_dec_ref(x_30); @@ -10001,7 +10025,7 @@ lean_inc(x_104); lean_inc_ref(x_103); lean_inc(x_102); lean_inc_ref(x_101); -x_209 = l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__11___redArg(x_206, x_207, x_99, x_208, x_208, x_101, x_102, x_103, x_104, x_105, x_106); +x_209 = l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__11___redArg(x_206, x_207, x_100, x_208, x_208, x_101, x_102, x_103, x_104, x_105, x_106); if (lean_obj_tag(x_209) == 0) { lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; @@ -10058,8 +10082,8 @@ lean_inc_ref(x_105); lean_inc(x_104); lean_inc_ref(x_103); lean_inc_ref(x_101); -lean_inc(x_100); -x_220 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__3(x_100, x_101, x_102, x_103, x_104, x_105, x_106); +lean_inc(x_99); +x_220 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__3(x_99, x_101, x_102, x_103, x_104, x_105, x_106); if (lean_obj_tag(x_220) == 0) { lean_object* x_221; lean_object* x_222; lean_object* x_223; uint8_t x_224; @@ -10080,8 +10104,8 @@ lean_dec(x_211); lean_dec_ref(x_27); x_62 = x_222; x_63 = x_223; -x_64 = x_97; -x_65 = x_100; +x_64 = x_99; +x_65 = x_97; x_66 = x_101; x_67 = x_102; x_68 = x_103; @@ -10145,15 +10169,15 @@ lean_inc_ref(x_105); lean_inc(x_104); lean_inc_ref(x_103); lean_inc_ref(x_101); -lean_inc(x_100); -x_241 = l_Lean_addTrace___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__5(x_100, x_240, x_101, x_102, x_103, x_104, x_105, x_106); +lean_inc(x_99); +x_241 = l_Lean_addTrace___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__5(x_99, x_240, x_101, x_102, x_103, x_104, x_105, x_106); if (lean_obj_tag(x_241) == 0) { lean_dec_ref(x_241); x_62 = x_222; x_63 = x_223; -x_64 = x_97; -x_65 = x_100; +x_64 = x_99; +x_65 = x_97; x_66 = x_101; x_67 = x_102; x_68 = x_103; @@ -10174,7 +10198,7 @@ lean_dec(x_104); lean_dec_ref(x_103); lean_dec(x_102); lean_dec_ref(x_101); -lean_dec(x_100); +lean_dec(x_99); lean_dec_ref(x_97); lean_dec_ref(x_34); lean_dec_ref(x_30); @@ -10213,7 +10237,7 @@ lean_dec(x_104); lean_dec_ref(x_103); lean_dec(x_102); lean_dec_ref(x_101); -lean_dec(x_100); +lean_dec(x_99); lean_dec_ref(x_97); lean_dec_ref(x_34); lean_dec_ref(x_30); @@ -10240,7 +10264,7 @@ lean_dec(x_104); lean_dec_ref(x_103); lean_dec(x_102); lean_dec_ref(x_101); -lean_dec(x_100); +lean_dec(x_99); lean_dec_ref(x_97); lean_dec_ref(x_34); lean_dec_ref(x_30); @@ -10274,7 +10298,7 @@ else lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_dec(x_212); lean_dec(x_211); -lean_dec(x_100); +lean_dec(x_99); lean_dec_ref(x_97); lean_dec_ref(x_34); lean_dec_ref(x_30); @@ -10315,7 +10339,7 @@ lean_dec(x_104); lean_dec_ref(x_103); lean_dec(x_102); lean_dec_ref(x_101); -lean_dec(x_100); +lean_dec(x_99); lean_dec_ref(x_97); lean_dec_ref(x_34); lean_dec_ref(x_30); @@ -10341,7 +10365,7 @@ lean_dec(x_104); lean_dec_ref(x_103); lean_dec(x_102); lean_dec_ref(x_101); -lean_dec(x_100); +lean_dec(x_99); lean_dec_ref(x_97); lean_dec_ref(x_34); lean_dec_ref(x_30); @@ -10379,7 +10403,7 @@ lean_dec(x_104); lean_dec_ref(x_103); lean_dec(x_102); lean_dec_ref(x_101); -lean_dec(x_100); +lean_dec(x_99); lean_dec_ref(x_97); lean_dec_ref(x_34); lean_dec_ref(x_30); @@ -10418,8 +10442,8 @@ lean_dec(x_104); lean_dec_ref(x_103); lean_dec(x_102); lean_dec_ref(x_101); -lean_dec(x_100); -lean_dec_ref(x_99); +lean_dec_ref(x_100); +lean_dec(x_99); lean_dec_ref(x_97); lean_dec_ref(x_34); lean_dec_ref(x_30); @@ -10489,8 +10513,8 @@ lean_dec(x_267); if (x_276 == 0) { lean_dec_ref(x_272); -x_99 = x_275; -x_100 = x_265; +x_99 = x_265; +x_100 = x_275; x_101 = x_15; x_102 = x_16; x_103 = x_17; @@ -10532,8 +10556,8 @@ x_289 = l_Lean_addTrace___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Ela if (lean_obj_tag(x_289) == 0) { lean_dec_ref(x_289); -x_99 = x_275; -x_100 = x_265; +x_99 = x_265; +x_100 = x_275; x_101 = x_15; x_102 = x_16; x_103 = x_17; @@ -11404,12 +11428,12 @@ goto block_147; block_95: { lean_object* x_41; -lean_inc(x_39); +lean_inc(x_37); +lean_inc_ref(x_34); +lean_inc(x_38); lean_inc_ref(x_33); -lean_inc(x_34); lean_inc_ref(x_35); -lean_inc_ref(x_36); -x_41 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_31, x_36, x_38, x_35, x_34, x_33, x_39); +x_41 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_31, x_35, x_39, x_33, x_38, x_34, x_37); if (lean_obj_tag(x_41) == 0) { lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; @@ -11420,7 +11444,7 @@ lean_inc_ref(x_40); x_43 = l_Lean_Elab_Tactic_Do_ProofMode_pushForallContextIntoHyps(x_40, x_42); lean_inc_ref(x_40); lean_inc(x_14); -x_44 = l_Lean_Elab_Tactic_Do_ProofMode_SPred_mkAnd(x_14, x_40, x_43, x_37); +x_44 = l_Lean_Elab_Tactic_Do_ProofMode_SPred_mkAnd(x_14, x_40, x_43, x_36); x_45 = !lean_is_exclusive(x_44); if (x_45 == 0) { @@ -11443,21 +11467,21 @@ lean_ctor_set(x_52, 0, x_14); lean_ctor_set(x_52, 1, x_40); lean_ctor_set(x_52, 2, x_46); lean_ctor_set(x_52, 3, x_51); -lean_inc(x_39); +lean_inc(x_37); +lean_inc_ref(x_34); +lean_inc(x_38); lean_inc_ref(x_33); -lean_inc(x_34); +lean_inc(x_39); lean_inc_ref(x_35); -lean_inc(x_38); -lean_inc_ref(x_36); -x_53 = lean_apply_8(x_4, x_52, x_36, x_38, x_35, x_34, x_33, x_39, lean_box(0)); +x_53 = lean_apply_8(x_4, x_52, x_35, x_39, x_33, x_38, x_34, x_37, lean_box(0)); if (lean_obj_tag(x_53) == 0) { lean_object* x_54; lean_object* x_55; x_54 = lean_ctor_get(x_53, 0); lean_inc(x_54); lean_dec_ref(x_53); -x_55 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_29, x_36, x_38, x_35, x_34, x_33, x_39); -lean_dec(x_38); +x_55 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_32, x_35, x_39, x_33, x_38, x_34, x_37); +lean_dec(x_39); if (lean_obj_tag(x_55) == 0) { uint8_t x_56; @@ -11476,7 +11500,7 @@ x_61 = l_Lean_mkAppN(x_46, x_28); x_62 = l_Lean_mkAppN(x_47, x_28); x_63 = l_Lean_mkAppN(x_54, x_28); lean_dec_ref(x_28); -x_64 = l_Lean_mkApp8(x_60, x_15, x_57, x_16, x_61, x_17, x_32, x_62, x_63); +x_64 = l_Lean_mkApp8(x_60, x_15, x_57, x_16, x_61, x_17, x_30, x_62, x_63); lean_ctor_set(x_55, 0, x_64); return x_55; } @@ -11496,7 +11520,7 @@ x_69 = l_Lean_mkAppN(x_46, x_28); x_70 = l_Lean_mkAppN(x_47, x_28); x_71 = l_Lean_mkAppN(x_54, x_28); lean_dec_ref(x_28); -x_72 = l_Lean_mkApp8(x_68, x_15, x_65, x_16, x_69, x_17, x_32, x_70, x_71); +x_72 = l_Lean_mkApp8(x_68, x_15, x_65, x_16, x_69, x_17, x_30, x_70, x_71); x_73 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_73, 0, x_72); return x_73; @@ -11508,7 +11532,7 @@ lean_dec(x_54); lean_free_object(x_44); lean_dec(x_47); lean_dec(x_46); -lean_dec_ref(x_32); +lean_dec_ref(x_30); lean_dec_ref(x_28); lean_dec_ref(x_17); lean_dec_ref(x_16); @@ -11524,12 +11548,12 @@ lean_dec(x_47); lean_dec(x_46); lean_dec(x_39); lean_dec(x_38); -lean_dec_ref(x_36); +lean_dec(x_37); lean_dec_ref(x_35); -lean_dec(x_34); +lean_dec_ref(x_34); lean_dec_ref(x_33); lean_dec_ref(x_32); -lean_dec_ref(x_29); +lean_dec_ref(x_30); lean_dec_ref(x_28); lean_dec_ref(x_17); lean_dec_ref(x_16); @@ -11562,21 +11586,21 @@ lean_ctor_set(x_80, 0, x_14); lean_ctor_set(x_80, 1, x_40); lean_ctor_set(x_80, 2, x_74); lean_ctor_set(x_80, 3, x_79); -lean_inc(x_39); +lean_inc(x_37); +lean_inc_ref(x_34); +lean_inc(x_38); lean_inc_ref(x_33); -lean_inc(x_34); +lean_inc(x_39); lean_inc_ref(x_35); -lean_inc(x_38); -lean_inc_ref(x_36); -x_81 = lean_apply_8(x_4, x_80, x_36, x_38, x_35, x_34, x_33, x_39, lean_box(0)); +x_81 = lean_apply_8(x_4, x_80, x_35, x_39, x_33, x_38, x_34, x_37, lean_box(0)); if (lean_obj_tag(x_81) == 0) { lean_object* x_82; lean_object* x_83; x_82 = lean_ctor_get(x_81, 0); lean_inc(x_82); lean_dec_ref(x_81); -x_83 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_29, x_36, x_38, x_35, x_34, x_33, x_39); -lean_dec(x_38); +x_83 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_32, x_35, x_39, x_33, x_38, x_34, x_37); +lean_dec(x_39); if (lean_obj_tag(x_83) == 0) { lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; @@ -11599,7 +11623,7 @@ x_90 = l_Lean_mkAppN(x_74, x_28); x_91 = l_Lean_mkAppN(x_75, x_28); x_92 = l_Lean_mkAppN(x_82, x_28); lean_dec_ref(x_28); -x_93 = l_Lean_mkApp8(x_89, x_15, x_84, x_16, x_90, x_17, x_32, x_91, x_92); +x_93 = l_Lean_mkApp8(x_89, x_15, x_84, x_16, x_90, x_17, x_30, x_91, x_92); if (lean_is_scalar(x_85)) { x_94 = lean_alloc_ctor(0, 1, 0); } else { @@ -11613,7 +11637,7 @@ else lean_dec(x_82); lean_dec(x_75); lean_dec(x_74); -lean_dec_ref(x_32); +lean_dec_ref(x_30); lean_dec_ref(x_28); lean_dec_ref(x_17); lean_dec_ref(x_16); @@ -11628,12 +11652,12 @@ lean_dec(x_75); lean_dec(x_74); lean_dec(x_39); lean_dec(x_38); -lean_dec_ref(x_36); +lean_dec(x_37); lean_dec_ref(x_35); -lean_dec(x_34); +lean_dec_ref(x_34); lean_dec_ref(x_33); lean_dec_ref(x_32); -lean_dec_ref(x_29); +lean_dec_ref(x_30); lean_dec_ref(x_28); lean_dec_ref(x_17); lean_dec_ref(x_16); @@ -11648,13 +11672,13 @@ else lean_dec_ref(x_40); lean_dec(x_39); lean_dec(x_38); -lean_dec_ref(x_37); +lean_dec(x_37); lean_dec_ref(x_36); lean_dec_ref(x_35); -lean_dec(x_34); +lean_dec_ref(x_34); lean_dec_ref(x_33); lean_dec_ref(x_32); -lean_dec_ref(x_29); +lean_dec_ref(x_30); lean_dec_ref(x_28); lean_dec_ref(x_25); lean_dec_ref(x_22); @@ -11769,17 +11793,17 @@ if (x_132 == 0) { lean_dec(x_111); lean_inc_ref(x_15); -x_29 = x_131; -x_30 = lean_box(0); +x_29 = lean_box(0); +x_30 = x_129; x_31 = x_130; -x_32 = x_129; -x_33 = x_102; -x_34 = x_101; -x_35 = x_100; -x_36 = x_98; -x_37 = x_120; -x_38 = x_99; -x_39 = x_103; +x_32 = x_131; +x_33 = x_100; +x_34 = x_102; +x_35 = x_98; +x_36 = x_120; +x_37 = x_103; +x_38 = x_101; +x_39 = x_99; x_40 = x_15; goto block_95; } @@ -11791,17 +11815,17 @@ lean_inc_ref(x_15); lean_inc(x_14); x_134 = l___private_Init_Data_Array_Basic_0__Array_foldrMUnsafe_fold___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__25(x_14, x_111, x_133, x_106, x_15); lean_dec(x_111); -x_29 = x_131; -x_30 = lean_box(0); +x_29 = lean_box(0); +x_30 = x_129; x_31 = x_130; -x_32 = x_129; -x_33 = x_102; -x_34 = x_101; -x_35 = x_100; -x_36 = x_98; -x_37 = x_120; -x_38 = x_99; -x_39 = x_103; +x_32 = x_131; +x_33 = x_100; +x_34 = x_102; +x_35 = x_98; +x_36 = x_120; +x_37 = x_103; +x_38 = x_101; +x_39 = x_99; x_40 = x_134; goto block_95; } @@ -14221,7 +14245,7 @@ x_11 = l_Lean_Elab_Tactic_Do_ProofMode_mIntroForallN___at___00__private_Lean_Ela return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_10; @@ -14229,16 +14253,7 @@ x_10 = l_Lean_Meta_getLevel(x_1, x_5, x_6, x_7, x_8); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_11; -x_11 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_11, 0, x_1); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { lean_object* x_9; @@ -14247,7 +14262,7 @@ lean_ctor_set(x_9, 0, x_1); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { lean_object* x_9; @@ -14256,6 +14271,14 @@ lean_ctor_set(x_9, 0, x_1); return x_9; } } +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Expr_abstractM(x_1, x_2, x_6, x_7, x_8, x_9); +return x_11; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { @@ -14852,7 +14875,7 @@ x_11 = l_Lean_MVarId_withContext___at___00__private_Lean_Elab_Tactic_Do_VCGen_0_ return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_10; @@ -14860,7 +14883,7 @@ x_10 = l_Lean_Meta_isTypeCorrect(x_1, x_5, x_6, x_7, x_8); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_10; @@ -14868,7 +14891,7 @@ x_10 = lean_get_match_equations_for(x_1, x_5, x_6, x_7, x_8); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_10; @@ -14876,7 +14899,7 @@ x_10 = l_Lean_Meta_check(x_1, x_5, x_6, x_7, x_8); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_10; @@ -14884,7 +14907,7 @@ x_10 = l_Lean_Meta_getLevel(x_1, x_5, x_6, x_7, x_8); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__14(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__14(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { lean_object* x_14; @@ -14892,7 +14915,7 @@ x_14 = l_Lean_Meta_mkLambdaFVars(x_1, x_2, x_3, x_4, x_3, x_4, x_5, x_9, x_10, x return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_11; @@ -14900,7 +14923,7 @@ x_11 = l_Lean_Meta_inferArgumentTypesN(x_1, x_2, x_6, x_7, x_8, x_9); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__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_10; @@ -14908,7 +14931,7 @@ x_10 = l_Lean_Meta_check(x_1, x_5, x_6, x_7, x_8); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_10; @@ -14916,7 +14939,7 @@ x_10 = l_Lean_Meta_isTypeCorrect(x_1, x_5, x_6, x_7, x_8); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_11; @@ -14924,7 +14947,7 @@ x_11 = l_Lean_Meta_inferArgumentTypesN(x_1, x_2, x_6, x_7, x_8, x_9); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_10; @@ -14932,7 +14955,7 @@ x_10 = l_Lean_Meta_isTypeCorrect(x_1, x_5, x_6, x_7, x_8); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_11; @@ -14940,7 +14963,7 @@ x_11 = l_Lean_Meta_inferArgumentTypesN(x_1, x_2, x_6, x_7, x_8, x_9); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_10; @@ -14948,7 +14971,7 @@ x_10 = l_Lean_Meta_check(x_1, x_5, x_6, x_7, x_8); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { lean_object* x_9; @@ -14956,7 +14979,7 @@ x_9 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7 return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__50___redArg___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__51___redArg___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { lean_object* x_9; @@ -14964,19 +14987,19 @@ x_9 = lean_apply_7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, lean_box(0)); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__50___redArg___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__51___redArg___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; -x_9 = l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__50___redArg___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_9 = l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__51___redArg___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__50___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__51___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_11; lean_object* x_12; -x_11 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__50___redArg___lam__0___boxed), 8, 3); +x_11 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__51___redArg___lam__0___boxed), 8, 3); lean_closure_set(x_11, 0, x_3); lean_closure_set(x_11, 1, x_4); lean_closure_set(x_11, 2, x_5); @@ -15006,11 +15029,11 @@ return x_15; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__50(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__51(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_12; -x_12 = l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__50___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_12 = l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__51___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_12; } } @@ -15121,7 +15144,7 @@ x_13 = l_Lean_Meta_lambdaBoundedTelescope___at___00__private_Lean_Elab_Tactic_Do return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__48___redArg(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__49___redArg(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_11; uint8_t x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; @@ -15158,11 +15181,11 @@ return x_18; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__48(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__49(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_12; -x_12 = l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__48___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_12 = l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__49___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_12; } } @@ -15208,7 +15231,7 @@ x_15 = l_Lean_Meta_withLetDecl___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__L return x_15; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59___lam__0(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__0(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { lean_object* x_14; @@ -15216,7 +15239,7 @@ x_14 = l_Lean_Meta_mkLambdaFVars(x_1, x_2, x_3, x_4, x_3, x_4, x_5, x_9, x_10, x return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__2(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__2(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { lean_object* x_14; @@ -15224,7 +15247,7 @@ x_14 = l_Lean_Meta_mkForallFVars(x_1, x_2, x_3, x_4, x_4, x_5, x_9, x_10, x_11, return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__42___lam__0(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__43___lam__0(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { lean_object* x_14; @@ -15301,7 +15324,7 @@ x_7 = l_Lean_isTracingEnabledFor___at___00Lean_Elab_Tactic_Do_mSpec___at___00__p return x_7; } } -LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_10; @@ -15828,7 +15851,7 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___closed__0() { +static lean_object* _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___closed__0() { _start: { lean_object* x_1; @@ -16358,8 +16381,8 @@ return x_10; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; double x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; double x_35; uint8_t x_36; lean_object* x_37; lean_object* x_38; lean_object* x_43; double x_44; lean_object* x_45; lean_object* x_46; double x_47; uint8_t x_48; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; double x_72; lean_object* x_73; double x_74; lean_object* x_75; uint8_t x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_84; lean_object* x_85; double x_86; double x_87; lean_object* x_88; uint8_t x_89; lean_object* x_95; lean_object* x_96; -x_95 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___closed__0; +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_31; double x_32; double x_33; lean_object* x_34; uint8_t x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_43; double x_44; double x_45; lean_object* x_46; uint8_t x_47; lean_object* x_48; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; double x_72; lean_object* x_73; lean_object* x_74; double x_75; uint8_t x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; double x_84; lean_object* x_85; uint8_t x_86; double x_87; lean_object* x_88; lean_object* x_89; lean_object* x_95; lean_object* x_96; +x_95 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___closed__0; lean_inc(x_11); lean_inc_ref(x_10); lean_inc(x_9); @@ -16381,7 +16404,7 @@ lean_inc(x_1); x_98 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__3(x_1, x_6, x_7, x_8, x_9, x_10, x_11); if (lean_obj_tag(x_98) == 0) { -lean_object* x_99; lean_object* x_100; double x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; double x_105; uint8_t x_106; uint8_t x_107; lean_object* x_115; lean_object* x_116; double x_117; lean_object* x_118; lean_object* x_119; double x_120; uint8_t x_121; double x_122; lean_object* x_126; lean_object* x_127; uint8_t x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_154; double x_155; lean_object* x_156; lean_object* x_157; double x_158; lean_object* x_159; uint8_t x_160; uint8_t x_161; lean_object* x_169; lean_object* x_170; lean_object* x_171; double x_172; double x_173; lean_object* x_174; uint8_t x_175; double x_176; lean_object* x_180; lean_object* x_181; uint8_t x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_205; lean_object* x_206; uint8_t x_241; +lean_object* x_99; lean_object* x_100; lean_object* x_101; double x_102; double x_103; uint8_t x_104; lean_object* x_105; lean_object* x_106; uint8_t x_107; lean_object* x_115; lean_object* x_116; double x_117; double x_118; lean_object* x_119; uint8_t x_120; lean_object* x_121; double x_122; lean_object* x_126; lean_object* x_127; lean_object* x_128; uint8_t x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_154; double x_155; lean_object* x_156; double x_157; uint8_t x_158; lean_object* x_159; lean_object* x_160; uint8_t x_161; lean_object* x_169; double x_170; lean_object* x_171; uint8_t x_172; double x_173; lean_object* x_174; lean_object* x_175; double x_176; lean_object* x_180; lean_object* x_181; uint8_t x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_205; lean_object* x_206; uint8_t x_241; x_99 = lean_ctor_get(x_98, 0); lean_inc(x_99); lean_dec_ref(x_98); @@ -16425,7 +16448,7 @@ if (x_108 == 0) if (x_107 == 0) { lean_object* x_109; -lean_dec_ref(x_104); +lean_dec_ref(x_105); lean_dec_ref(x_5); lean_dec_ref(x_2); lean_dec(x_1); @@ -16435,13 +16458,13 @@ if (lean_obj_tag(x_109) == 0) { lean_object* x_110; lean_dec_ref(x_109); -x_110 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13_spec__29___redArg(x_102); +x_110 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13_spec__29___redArg(x_101); return x_110; } else { uint8_t x_111; -lean_dec_ref(x_102); +lean_dec_ref(x_101); x_111 = !lean_is_exclusive(x_109); if (x_111 == 0) { @@ -16462,39 +16485,39 @@ return x_113; else { lean_dec_ref(x_100); -x_43 = x_102; -x_44 = x_101; -x_45 = lean_box(0); -x_46 = x_104; -x_47 = x_105; -x_48 = x_106; +x_43 = x_101; +x_44 = x_102; +x_45 = x_103; +x_46 = x_105; +x_47 = x_104; +x_48 = lean_box(0); goto block_53; } } else { lean_dec_ref(x_100); -x_43 = x_102; -x_44 = x_101; -x_45 = lean_box(0); -x_46 = x_104; -x_47 = x_105; -x_48 = x_106; +x_43 = x_101; +x_44 = x_102; +x_45 = x_103; +x_46 = x_105; +x_47 = x_104; +x_48 = lean_box(0); goto block_53; } } block_125: { double x_123; uint8_t x_124; -x_123 = lean_float_sub(x_120, x_117); +x_123 = lean_float_sub(x_117, x_118); x_124 = lean_float_decLt(x_122, x_123); x_100 = x_115; -x_101 = x_117; -x_102 = x_116; -x_103 = lean_box(0); -x_104 = x_119; -x_105 = x_120; -x_106 = x_121; +x_101 = x_116; +x_102 = x_117; +x_103 = x_118; +x_104 = x_120; +x_105 = x_119; +x_106 = lean_box(0); x_107 = x_124; goto block_114; } @@ -16506,14 +16529,14 @@ lean_inc_ref(x_10); lean_inc(x_9); lean_inc_ref(x_8); lean_inc_ref(x_6); -x_133 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_130, x_6, x_7, x_8, x_9, x_10, x_11); +x_133 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_128, x_6, x_7, x_8, x_9, x_10, x_11); if (lean_obj_tag(x_133) == 0) { lean_object* x_134; double x_135; double x_136; double x_137; double x_138; double x_139; lean_object* x_140; uint8_t x_141; x_134 = lean_ctor_get(x_133, 0); lean_inc(x_134); lean_dec_ref(x_133); -x_135 = lean_float_of_nat(x_127); +x_135 = lean_float_of_nat(x_126); x_136 = l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13___redArg___closed__2; x_137 = lean_float_div(x_135, x_136); x_138 = lean_float_of_nat(x_134); @@ -16523,19 +16546,19 @@ x_141 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00__private_Lean_El if (x_141 == 0) { lean_dec(x_97); -x_100 = x_126; -x_101 = x_137; -x_102 = x_131; -x_103 = lean_box(0); -x_104 = x_129; -x_105 = x_139; -x_106 = x_141; +x_100 = x_127; +x_101 = x_131; +x_102 = x_139; +x_103 = x_137; +x_104 = x_141; +x_105 = x_130; +x_106 = lean_box(0); x_107 = x_141; goto block_114; } else { -if (x_128 == 0) +if (x_129 == 0) { lean_object* x_142; lean_object* x_143; double x_144; double x_145; double x_146; x_142 = l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13___redArg___closed__4; @@ -16544,13 +16567,13 @@ lean_dec(x_97); x_144 = lean_float_of_nat(x_143); x_145 = l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13___redArg___closed__5; x_146 = lean_float_div(x_144, x_145); -x_115 = x_126; +x_115 = x_127; x_116 = x_131; -x_117 = x_137; -x_118 = lean_box(0); -x_119 = x_129; -x_120 = x_139; -x_121 = x_141; +x_117 = x_139; +x_118 = x_137; +x_119 = x_130; +x_120 = x_141; +x_121 = lean_box(0); x_122 = x_146; goto block_125; } @@ -16561,13 +16584,13 @@ x_147 = l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lea x_148 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13_spec__30(x_97, x_147); lean_dec(x_97); x_149 = lean_float_of_nat(x_148); -x_115 = x_126; +x_115 = x_127; x_116 = x_131; -x_117 = x_137; -x_118 = lean_box(0); -x_119 = x_129; -x_120 = x_139; -x_121 = x_141; +x_117 = x_139; +x_118 = x_137; +x_119 = x_130; +x_120 = x_141; +x_121 = lean_box(0); x_122 = x_149; goto block_125; } @@ -16577,9 +16600,9 @@ else { uint8_t x_150; lean_dec_ref(x_131); -lean_dec_ref(x_129); -lean_dec(x_127); -lean_dec_ref(x_126); +lean_dec_ref(x_130); +lean_dec_ref(x_127); +lean_dec(x_126); lean_dec(x_99); lean_dec(x_97); lean_dec(x_11); @@ -16655,39 +16678,39 @@ return x_167; else { lean_dec_ref(x_154); -x_84 = lean_box(0); +x_84 = x_155; x_85 = x_156; -x_86 = x_155; -x_87 = x_158; +x_86 = x_158; +x_87 = x_157; x_88 = x_159; -x_89 = x_160; +x_89 = lean_box(0); goto block_94; } } else { lean_dec_ref(x_154); -x_84 = lean_box(0); +x_84 = x_155; x_85 = x_156; -x_86 = x_155; -x_87 = x_158; +x_86 = x_158; +x_87 = x_157; x_88 = x_159; -x_89 = x_160; +x_89 = lean_box(0); goto block_94; } } block_179: { double x_177; uint8_t x_178; -x_177 = lean_float_sub(x_173, x_172); +x_177 = lean_float_sub(x_170, x_173); x_178 = lean_float_decLt(x_176, x_177); x_154 = x_169; -x_155 = x_172; +x_155 = x_170; x_156 = x_171; -x_157 = lean_box(0); -x_158 = x_173; +x_157 = x_173; +x_158 = x_172; x_159 = x_174; -x_160 = x_175; +x_160 = lean_box(0); x_161 = x_178; goto block_168; } @@ -16706,20 +16729,20 @@ lean_object* x_188; double x_189; double x_190; lean_object* x_191; uint8_t x_19 x_188 = lean_ctor_get(x_187, 0); lean_inc(x_188); lean_dec_ref(x_187); -x_189 = lean_float_of_nat(x_180); +x_189 = lean_float_of_nat(x_181); x_190 = lean_float_of_nat(x_188); x_191 = l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13___redArg___closed__3; x_192 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13_spec__27(x_97, x_191); if (x_192 == 0) { lean_dec(x_97); -x_154 = x_181; -x_155 = x_189; +x_154 = x_180; +x_155 = x_190; x_156 = x_185; -x_157 = lean_box(0); -x_158 = x_190; +x_157 = x_189; +x_158 = x_192; x_159 = x_183; -x_160 = x_192; +x_160 = lean_box(0); x_161 = x_192; goto block_168; } @@ -16734,13 +16757,13 @@ lean_dec(x_97); x_195 = lean_float_of_nat(x_194); x_196 = l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13___redArg___closed__5; x_197 = lean_float_div(x_195, x_196); -x_169 = x_181; -x_170 = lean_box(0); +x_169 = x_180; +x_170 = x_190; x_171 = x_185; -x_172 = x_189; -x_173 = x_190; +x_172 = x_192; +x_173 = x_189; x_174 = x_183; -x_175 = x_192; +x_175 = lean_box(0); x_176 = x_197; goto block_179; } @@ -16751,13 +16774,13 @@ x_198 = l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lea x_199 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13_spec__30(x_97, x_198); lean_dec(x_97); x_200 = lean_float_of_nat(x_199); -x_169 = x_181; -x_170 = lean_box(0); +x_169 = x_180; +x_170 = x_190; x_171 = x_185; -x_172 = x_189; -x_173 = x_190; +x_172 = x_192; +x_173 = x_189; x_174 = x_183; -x_175 = x_192; +x_175 = lean_box(0); x_176 = x_200; goto block_179; } @@ -16768,8 +16791,8 @@ else uint8_t x_201; lean_dec_ref(x_185); lean_dec_ref(x_183); -lean_dec_ref(x_181); -lean_dec(x_180); +lean_dec(x_181); +lean_dec_ref(x_180); lean_dec(x_99); lean_dec(x_97); lean_dec(x_11); @@ -16849,11 +16872,11 @@ x_215 = !lean_is_exclusive(x_214); if (x_215 == 0) { lean_ctor_set_tag(x_214, 1); -x_126 = x_213; -x_127 = x_212; -x_128 = x_210; -x_129 = x_208; -x_130 = x_205; +x_126 = x_212; +x_127 = x_213; +x_128 = x_205; +x_129 = x_210; +x_130 = x_208; x_131 = x_214; x_132 = lean_box(0); goto block_153; @@ -16866,11 +16889,11 @@ lean_inc(x_216); lean_dec(x_214); x_217 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_217, 0, x_216); -x_126 = x_213; -x_127 = x_212; -x_128 = x_210; -x_129 = x_208; -x_130 = x_205; +x_126 = x_212; +x_127 = x_213; +x_128 = x_205; +x_129 = x_210; +x_130 = x_208; x_131 = x_217; x_132 = lean_box(0); goto block_153; @@ -16883,11 +16906,11 @@ x_218 = !lean_is_exclusive(x_214); if (x_218 == 0) { lean_ctor_set_tag(x_214, 0); -x_126 = x_213; -x_127 = x_212; -x_128 = x_210; -x_129 = x_208; -x_130 = x_205; +x_126 = x_212; +x_127 = x_213; +x_128 = x_205; +x_129 = x_210; +x_130 = x_208; x_131 = x_214; x_132 = lean_box(0); goto block_153; @@ -16900,11 +16923,11 @@ lean_inc(x_219); lean_dec(x_214); x_220 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_220, 0, x_219); -x_126 = x_213; -x_127 = x_212; -x_128 = x_210; -x_129 = x_208; -x_130 = x_205; +x_126 = x_212; +x_127 = x_213; +x_128 = x_205; +x_129 = x_210; +x_130 = x_208; x_131 = x_220; x_132 = lean_box(0); goto block_153; @@ -16979,8 +17002,8 @@ x_228 = !lean_is_exclusive(x_227); if (x_228 == 0) { lean_ctor_set_tag(x_227, 1); -x_180 = x_225; -x_181 = x_226; +x_180 = x_226; +x_181 = x_225; x_182 = x_210; x_183 = x_208; x_184 = x_206; @@ -16996,8 +17019,8 @@ lean_inc(x_229); lean_dec(x_227); x_230 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_230, 0, x_229); -x_180 = x_225; -x_181 = x_226; +x_180 = x_226; +x_181 = x_225; x_182 = x_210; x_183 = x_208; x_184 = x_206; @@ -17013,8 +17036,8 @@ x_231 = !lean_is_exclusive(x_227); if (x_231 == 0) { lean_ctor_set_tag(x_227, 0); -x_180 = x_225; -x_181 = x_226; +x_180 = x_226; +x_181 = x_225; x_182 = x_210; x_183 = x_208; x_184 = x_206; @@ -17030,8 +17053,8 @@ lean_inc(x_232); lean_dec(x_227); x_233 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_233, 0, x_232); -x_180 = x_225; -x_181 = x_226; +x_180 = x_226; +x_181 = x_225; x_182 = x_210; x_183 = x_208; x_184 = x_206; @@ -17206,7 +17229,7 @@ return x_29; } block_42: { -if (x_36 == 0) +if (x_35 == 0) { double x_39; lean_object* x_40; x_39 = l_Lean_addTrace___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__5___lam__1___closed__0; @@ -17216,9 +17239,9 @@ lean_ctor_set(x_40, 1, x_5); lean_ctor_set_float(x_40, sizeof(void*)*2, x_39); lean_ctor_set_float(x_40, sizeof(void*)*2 + 8, x_39); lean_ctor_set_uint8(x_40, sizeof(void*)*2 + 16, x_4); -x_13 = x_32; -x_14 = x_33; -x_15 = x_34; +x_13 = x_31; +x_14 = x_34; +x_15 = x_36; x_16 = x_37; x_17 = x_40; x_18 = x_6; @@ -17236,12 +17259,12 @@ lean_object* x_41; x_41 = lean_alloc_ctor(0, 2, 17); lean_ctor_set(x_41, 0, x_1); lean_ctor_set(x_41, 1, x_5); -lean_ctor_set_float(x_41, sizeof(void*)*2, x_31); -lean_ctor_set_float(x_41, sizeof(void*)*2 + 8, x_35); +lean_ctor_set_float(x_41, sizeof(void*)*2, x_33); +lean_ctor_set_float(x_41, sizeof(void*)*2 + 8, x_32); lean_ctor_set_uint8(x_41, sizeof(void*)*2 + 16, x_4); -x_13 = x_32; -x_14 = x_33; -x_15 = x_34; +x_13 = x_31; +x_14 = x_34; +x_15 = x_36; x_16 = x_37; x_17 = x_41; x_18 = x_6; @@ -17273,12 +17296,12 @@ x_51 = lean_ctor_get(x_50, 0); lean_inc(x_51); lean_dec_ref(x_50); lean_inc(x_49); -x_31 = x_44; -x_32 = x_43; -x_33 = x_49; -x_34 = x_46; +x_31 = x_43; +x_32 = x_44; +x_33 = x_45; +x_34 = x_49; x_35 = x_47; -x_36 = x_48; +x_36 = x_46; x_37 = x_51; x_38 = lean_box(0); goto block_42; @@ -17289,12 +17312,12 @@ lean_object* x_52; lean_dec_ref(x_50); x_52 = l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13___redArg___closed__1; lean_inc(x_49); -x_31 = x_44; -x_32 = x_43; -x_33 = x_49; -x_34 = x_46; +x_31 = x_43; +x_32 = x_44; +x_33 = x_45; +x_34 = x_49; x_35 = x_47; -x_36 = x_48; +x_36 = x_46; x_37 = x_52; x_38 = lean_box(0); goto block_42; @@ -17303,7 +17326,7 @@ goto block_42; block_71: { lean_object* x_66; -x_66 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13_spec__28(x_56, x_58, x_57, x_55, x_59, x_60, x_61, x_62, x_63, x_64); +x_66 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13_spec__28(x_57, x_58, x_55, x_56, x_59, x_60, x_61, x_62, x_63, x_64); lean_dec(x_60); if (lean_obj_tag(x_66) == 0) { @@ -17346,8 +17369,8 @@ lean_ctor_set_float(x_81, sizeof(void*)*2, x_80); lean_ctor_set_float(x_81, sizeof(void*)*2 + 8, x_80); lean_ctor_set_uint8(x_81, sizeof(void*)*2 + 16, x_4); x_54 = x_73; -x_55 = x_78; -x_56 = x_75; +x_55 = x_74; +x_56 = x_78; x_57 = x_77; x_58 = x_81; x_59 = x_6; @@ -17365,12 +17388,12 @@ lean_object* x_82; x_82 = lean_alloc_ctor(0, 2, 17); lean_ctor_set(x_82, 0, x_1); lean_ctor_set(x_82, 1, x_5); -lean_ctor_set_float(x_82, sizeof(void*)*2, x_72); -lean_ctor_set_float(x_82, sizeof(void*)*2 + 8, x_74); +lean_ctor_set_float(x_82, sizeof(void*)*2, x_75); +lean_ctor_set_float(x_82, sizeof(void*)*2 + 8, x_72); lean_ctor_set_uint8(x_82, sizeof(void*)*2 + 16, x_4); x_54 = x_73; -x_55 = x_78; -x_56 = x_75; +x_55 = x_74; +x_56 = x_78; x_57 = x_77; x_58 = x_82; x_59 = x_6; @@ -17402,12 +17425,12 @@ x_92 = lean_ctor_get(x_91, 0); lean_inc(x_92); lean_dec_ref(x_91); lean_inc(x_90); -x_72 = x_86; +x_72 = x_84; x_73 = x_85; -x_74 = x_87; -x_75 = x_88; -x_76 = x_89; -x_77 = x_90; +x_74 = x_90; +x_75 = x_87; +x_76 = x_86; +x_77 = x_88; x_78 = x_92; x_79 = lean_box(0); goto block_83; @@ -17418,12 +17441,12 @@ lean_object* x_93; lean_dec_ref(x_91); x_93 = l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13___redArg___closed__1; lean_inc(x_90); -x_72 = x_86; +x_72 = x_84; x_73 = x_85; -x_74 = x_87; -x_75 = x_88; -x_76 = x_89; -x_77 = x_90; +x_74 = x_90; +x_75 = x_87; +x_76 = x_86; +x_77 = x_88; x_78 = x_93; x_79 = lean_box(0); goto block_83; @@ -17583,7 +17606,7 @@ x_11 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs return x_11; } } -static lean_object* _init_l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__0___closed__1() { +static lean_object* _init_l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__0___closed__1() { _start: { lean_object* x_1; @@ -17591,7 +17614,7 @@ x_1 = lean_mk_string_unchecked("apply", 5, 5); return x_1; } } -static lean_object* _init_l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__0___closed__0() { +static lean_object* _init_l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__0___closed__0() { _start: { lean_object* x_1; @@ -17603,8 +17626,8 @@ static lean_object* _init_l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_El _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__0___closed__1; -x_2 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__0___closed__0; +x_1 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__0___closed__1; +x_2 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__0___closed__0; x_3 = l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___closed__3; x_4 = l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___closed__2; x_5 = l_Lean_Name_mkStr4(x_4, x_3, x_2, x_1); @@ -17988,7 +18011,7 @@ lean_dec(x_2); return x_10; } } -static lean_object* _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__0() { +static lean_object* _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__0() { _start: { lean_object* x_1; @@ -17996,16 +18019,16 @@ x_1 = lean_mk_string_unchecked("Tried to assign already assigned metavariable `" return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__1() { +static lean_object* _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__1() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__0; +x_1 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__0; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__2() { +static lean_object* _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__2() { _start: { lean_object* x_1; @@ -18013,16 +18036,16 @@ x_1 = lean_mk_string_unchecked("`. MVar: ", 9, 9); return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__3() { +static lean_object* _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__2; +x_1 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__2; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__4() { +static lean_object* _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__4() { _start: { lean_object* x_1; @@ -18030,16 +18053,16 @@ x_1 = lean_mk_string_unchecked("\nAssignment: ", 13, 13); return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__5() { +static lean_object* _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__4; +x_1 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__4; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__6() { +static lean_object* _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__6() { _start: { lean_object* x_1; @@ -18047,11 +18070,11 @@ x_1 = lean_mk_string_unchecked("\nNew assignment: ", 17, 17); return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__7() { +static lean_object* _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__6; +x_1 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__6; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } @@ -18175,12 +18198,12 @@ lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean x_39 = lean_ctor_get(x_38, 0); lean_inc(x_39); lean_dec_ref(x_38); -x_40 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__1; +x_40 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__1; x_41 = l_Lean_MessageData_ofName(x_39); lean_ctor_set_tag(x_4, 7); lean_ctor_set(x_4, 1, x_41); lean_ctor_set(x_4, 0, x_40); -x_42 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__3; +x_42 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__3; x_43 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_43, 0, x_4); lean_ctor_set(x_43, 1, x_42); @@ -18189,7 +18212,7 @@ lean_ctor_set(x_21, 0, x_15); x_44 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_44, 0, x_43); lean_ctor_set(x_44, 1, x_21); -x_45 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__5; +x_45 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__5; x_46 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_46, 0, x_44); lean_ctor_set(x_46, 1, x_45); @@ -18198,7 +18221,7 @@ x_48 = l_Lean_MessageData_ofExpr(x_47); x_49 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_49, 0, x_46); lean_ctor_set(x_49, 1, x_48); -x_50 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__7; +x_50 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__7; x_51 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_51, 0, x_49); lean_ctor_set(x_51, 1, x_50); @@ -18281,7 +18304,7 @@ if (lean_obj_tag(x_31) == 0) { lean_object* x_32; lean_dec_ref(x_31); -x_32 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg(x_1, x_2, x_16, x_2, x_6, x_7, x_8, x_9, x_10, x_11); +x_32 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg(x_1, x_2, x_16, x_2, x_6, x_7, x_8, x_9, x_10, x_11); return x_32; } else @@ -18356,12 +18379,12 @@ lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean x_77 = lean_ctor_get(x_76, 0); lean_inc(x_77); lean_dec_ref(x_76); -x_78 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__1; +x_78 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__1; x_79 = l_Lean_MessageData_ofName(x_77); lean_ctor_set_tag(x_4, 7); lean_ctor_set(x_4, 1, x_79); lean_ctor_set(x_4, 0, x_78); -x_80 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__3; +x_80 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__3; x_81 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_81, 0, x_4); lean_ctor_set(x_81, 1, x_80); @@ -18371,7 +18394,7 @@ lean_ctor_set(x_82, 0, x_15); x_83 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_83, 0, x_81); lean_ctor_set(x_83, 1, x_82); -x_84 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__5; +x_84 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__5; x_85 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_85, 0, x_83); lean_ctor_set(x_85, 1, x_84); @@ -18380,7 +18403,7 @@ x_87 = l_Lean_MessageData_ofExpr(x_86); x_88 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_88, 0, x_85); lean_ctor_set(x_88, 1, x_87); -x_89 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__7; +x_89 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__7; x_90 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_90, 0, x_88); lean_ctor_set(x_90, 1, x_89); @@ -18463,7 +18486,7 @@ if (lean_obj_tag(x_69) == 0) { lean_object* x_70; lean_dec_ref(x_69); -x_70 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg(x_1, x_2, x_16, x_2, x_6, x_7, x_8, x_9, x_10, x_11); +x_70 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg(x_1, x_2, x_16, x_2, x_6, x_7, x_8, x_9, x_10, x_11); return x_70; } else @@ -18494,7 +18517,7 @@ if (lean_obj_tag(x_100) == 0) { lean_object* x_101; lean_dec_ref(x_100); -x_101 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg(x_1, x_2, x_16, x_2, x_6, x_7, x_8, x_9, x_10, x_11); +x_101 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg(x_1, x_2, x_16, x_2, x_6, x_7, x_8, x_9, x_10, x_11); return x_101; } else @@ -18546,7 +18569,7 @@ else lean_object* x_105; lean_free_object(x_4); lean_dec(x_15); -x_105 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg(x_1, x_2, x_16, x_2, x_6, x_7, x_8, x_9, x_10, x_11); +x_105 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg(x_1, x_2, x_16, x_2, x_6, x_7, x_8, x_9, x_10, x_11); return x_105; } } @@ -18684,12 +18707,12 @@ lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; x_133 = lean_ctor_get(x_132, 0); lean_inc(x_133); lean_dec_ref(x_132); -x_134 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__1; +x_134 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__1; x_135 = l_Lean_MessageData_ofName(x_133); x_136 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_136, 0, x_134); lean_ctor_set(x_136, 1, x_135); -x_137 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__3; +x_137 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__3; x_138 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_138, 0, x_136); lean_ctor_set(x_138, 1, x_137); @@ -18703,7 +18726,7 @@ lean_ctor_set(x_139, 0, x_109); x_140 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_140, 0, x_138); lean_ctor_set(x_140, 1, x_139); -x_141 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__5; +x_141 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__5; x_142 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_142, 0, x_140); lean_ctor_set(x_142, 1, x_141); @@ -18712,7 +18735,7 @@ x_144 = l_Lean_MessageData_ofExpr(x_143); x_145 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_145, 0, x_142); lean_ctor_set(x_145, 1, x_144); -x_146 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__7; +x_146 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__7; x_147 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_147, 0, x_145); lean_ctor_set(x_147, 1, x_146); @@ -18795,7 +18818,7 @@ if (lean_obj_tag(x_125) == 0) { lean_object* x_126; lean_dec_ref(x_125); -x_126 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg(x_1, x_2, x_110, x_2, x_6, x_7, x_8, x_9, x_10, x_11); +x_126 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg(x_1, x_2, x_110, x_2, x_6, x_7, x_8, x_9, x_10, x_11); return x_126; } else @@ -18824,7 +18847,7 @@ if (lean_obj_tag(x_157) == 0) { lean_object* x_158; lean_dec_ref(x_157); -x_158 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg(x_1, x_2, x_110, x_2, x_6, x_7, x_8, x_9, x_10, x_11); +x_158 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg(x_1, x_2, x_110, x_2, x_6, x_7, x_8, x_9, x_10, x_11); return x_158; } else @@ -18875,7 +18898,7 @@ else { lean_object* x_162; lean_dec(x_109); -x_162 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg(x_1, x_2, x_110, x_2, x_6, x_7, x_8, x_9, x_10, x_11); +x_162 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg(x_1, x_2, x_110, x_2, x_6, x_7, x_8, x_9, x_10, x_11); return x_162; } } @@ -21168,7 +21191,7 @@ x_3 = l_Lean_Name_mkStr2(x_2, x_1); return x_3; } } -static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__8() { +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__8() { _start: { lean_object* x_1; lean_object* x_2; @@ -21688,7 +21711,7 @@ x_93 = lean_ctor_get(x_92, 0); lean_inc(x_93); lean_dec_ref(x_92); x_94 = l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___closed__10; -x_95 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__8; +x_95 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__8; x_96 = lean_array_push(x_95, x_81); x_97 = lean_array_push(x_96, x_93); x_98 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___lam__11___boxed), 10, 2); @@ -22704,7 +22727,7 @@ x_66 = lean_ctor_get(x_65, 0); lean_inc(x_66); lean_dec_ref(x_65); x_67 = l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___closed__10; -x_68 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__8; +x_68 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__8; x_69 = lean_array_push(x_68, x_54); x_70 = lean_array_push(x_69, x_66); x_71 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__2___boxed), 10, 2); @@ -23177,7 +23200,7 @@ x_166 = lean_ctor_get(x_165, 0); lean_inc(x_166); lean_dec_ref(x_165); x_167 = l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___closed__10; -x_168 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__8; +x_168 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__8; x_169 = lean_array_push(x_168, x_154); x_170 = lean_array_push(x_169, x_166); x_171 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite___lam__2___boxed), 10, 2); @@ -23392,7 +23415,7 @@ static lean_object* _init_l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_El { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__8___closed__4; -x_2 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__0___closed__0; +x_2 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__0___closed__0; x_3 = l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___closed__3; x_4 = l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___closed__2; x_5 = l_Lean_Name_mkStr4(x_4, x_3, x_2, x_1); @@ -25569,7 +25592,7 @@ x_76 = l_Lean_Meta_Context_config(x_12); x_77 = !lean_is_exclusive(x_76); if (x_77 == 0) { -uint8_t x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; uint8_t x_85; uint8_t x_86; uint8_t x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; uint8_t x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; uint8_t x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; uint8_t x_130; lean_object* x_131; uint8_t x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; uint8_t x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; uint8_t x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; uint8_t x_158; lean_object* x_159; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; uint8_t x_382; lean_object* x_383; lean_object* x_391; uint64_t x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; +uint8_t x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; uint8_t x_85; uint8_t x_86; uint8_t x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; uint8_t x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; uint8_t x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; uint8_t x_130; uint8_t x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; uint8_t x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; uint8_t x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; uint8_t x_157; lean_object* x_158; lean_object* x_159; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; uint8_t x_382; lean_object* x_383; lean_object* x_391; uint64_t x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; x_78 = lean_ctor_get_uint8(x_12, sizeof(void*)*7); x_79 = lean_ctor_get(x_12, 1); x_80 = lean_ctor_get(x_12, 2); @@ -25716,10 +25739,10 @@ return x_402; { lean_object* x_112; lean_inc(x_95); -lean_inc_ref(x_106); +lean_inc_ref(x_109); lean_inc_ref(x_74); lean_inc_ref(x_72); -x_112 = l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20(x_72, x_74, x_106, x_95, x_111, x_105, x_104, x_109, x_110, x_102, x_108); +x_112 = l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20(x_72, x_74, x_109, x_95, x_111, x_110, x_103, x_101, x_106, x_108, x_102); if (lean_obj_tag(x_112) == 0) { lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; @@ -25738,12 +25761,12 @@ lean_ctor_set(x_116, 0, x_64); lean_ctor_set(x_116, 1, x_115); lean_inc_ref(x_116); x_117 = l_Lean_mkConst(x_114, x_116); -lean_inc_ref(x_106); +lean_inc_ref(x_109); lean_inc_ref(x_5); lean_inc_ref(x_72); lean_inc_ref(x_74); lean_inc_ref(x_117); -x_118 = l_Lean_mkApp4(x_117, x_74, x_72, x_5, x_106); +x_118 = l_Lean_mkApp4(x_117, x_74, x_72, x_5, x_109); lean_inc(x_95); lean_inc_ref(x_5); lean_inc_ref(x_72); @@ -25752,7 +25775,7 @@ x_119 = l_Lean_mkApp4(x_117, x_74, x_72, x_5, x_95); x_120 = l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__8___closed__5; lean_inc_ref(x_116); x_121 = l_Lean_mkConst(x_120, x_116); -x_122 = l_Lean_mkApp6(x_121, x_74, x_72, x_5, x_106, x_95, x_113); +x_122 = l_Lean_mkApp6(x_121, x_74, x_72, x_5, x_109, x_95, x_113); x_123 = lean_box(x_99); x_124 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__5___boxed), 11, 10); lean_closure_set(x_124, 0, x_60); @@ -25765,7 +25788,7 @@ lean_closure_set(x_124, 6, x_119); lean_closure_set(x_124, 7, x_122); lean_closure_set(x_124, 8, x_65); lean_closure_set(x_124, 9, x_100); -x_17 = x_103; +x_17 = x_104; x_18 = x_107; x_19 = x_124; x_20 = lean_box(0); @@ -25773,9 +25796,9 @@ goto block_24; } else { +lean_dec_ref(x_109); lean_dec_ref(x_107); -lean_dec_ref(x_106); -lean_dec_ref(x_103); +lean_dec_ref(x_104); lean_dec_ref(x_100); lean_dec_ref(x_98); lean_dec(x_95); @@ -25790,7 +25813,7 @@ return x_112; } block_143: { -if (x_132 == 0) +if (x_131 == 0) { if (x_63 == 0) { @@ -25800,7 +25823,7 @@ lean_dec(x_137); lean_dec_ref(x_136); lean_dec(x_135); lean_dec_ref(x_134); -lean_dec_ref(x_129); +lean_dec_ref(x_132); lean_dec_ref(x_127); lean_dec_ref(x_98); lean_dec(x_95); @@ -25811,31 +25834,31 @@ lean_dec(x_64); lean_dec(x_29); lean_dec(x_6); lean_dec_ref(x_5); -x_17 = x_128; +x_17 = x_129; x_18 = x_133; -x_19 = x_131; +x_19 = x_128; x_20 = lean_box(0); goto block_24; } else { -lean_dec_ref(x_131); +lean_dec_ref(x_128); if (x_130 == 0) { if (x_63 == 0) { x_99 = x_126; x_100 = x_127; -x_101 = lean_box(0); -x_102 = x_138; -x_103 = x_128; -x_104 = x_135; -x_105 = x_134; -x_106 = x_129; +x_101 = x_136; +x_102 = x_139; +x_103 = x_135; +x_104 = x_129; +x_105 = lean_box(0); +x_106 = x_137; x_107 = x_133; -x_108 = x_139; -x_109 = x_136; -x_110 = x_137; +x_108 = x_138; +x_109 = x_132; +x_110 = x_134; x_111 = x_6; goto block_125; } @@ -25846,16 +25869,16 @@ x_141 = l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePos x_142 = l_Lean_Name_append(x_6, x_141); x_99 = x_126; x_100 = x_127; -x_101 = lean_box(0); -x_102 = x_138; -x_103 = x_128; -x_104 = x_135; -x_105 = x_134; -x_106 = x_129; +x_101 = x_136; +x_102 = x_139; +x_103 = x_135; +x_104 = x_129; +x_105 = lean_box(0); +x_106 = x_137; x_107 = x_133; -x_108 = x_139; -x_109 = x_136; -x_110 = x_137; +x_108 = x_138; +x_109 = x_132; +x_110 = x_134; x_111 = x_142; goto block_125; } @@ -25864,16 +25887,16 @@ else { x_99 = x_126; x_100 = x_127; -x_101 = lean_box(0); -x_102 = x_138; -x_103 = x_128; -x_104 = x_135; -x_105 = x_134; -x_106 = x_129; +x_101 = x_136; +x_102 = x_139; +x_103 = x_135; +x_104 = x_129; +x_105 = lean_box(0); +x_106 = x_137; x_107 = x_133; -x_108 = x_139; -x_109 = x_136; -x_110 = x_137; +x_108 = x_138; +x_109 = x_132; +x_110 = x_134; x_111 = x_6; goto block_125; } @@ -25887,7 +25910,7 @@ lean_dec(x_137); lean_dec_ref(x_136); lean_dec(x_135); lean_dec_ref(x_134); -lean_dec_ref(x_129); +lean_dec_ref(x_132); lean_dec_ref(x_127); lean_dec_ref(x_98); lean_dec(x_95); @@ -25898,9 +25921,9 @@ lean_dec(x_64); lean_dec(x_29); lean_dec(x_6); lean_dec_ref(x_5); -x_17 = x_128; +x_17 = x_129; x_18 = x_133; -x_19 = x_131; +x_19 = x_128; x_20 = lean_box(0); goto block_24; } @@ -25908,7 +25931,7 @@ goto block_24; block_168: { lean_object* x_160; lean_object* x_161; -lean_inc_ref(x_153); +lean_inc_ref(x_148); lean_inc_ref(x_66); lean_inc_ref(x_65); lean_inc(x_64); @@ -25920,13 +25943,13 @@ if (lean_is_scalar(x_68)) { lean_ctor_set(x_160, 0, x_64); lean_ctor_set(x_160, 1, x_65); lean_ctor_set(x_160, 2, x_66); -lean_ctor_set(x_160, 3, x_153); -lean_inc(x_148); -lean_inc_ref(x_149); -lean_inc(x_157); -lean_inc_ref(x_154); -lean_inc_ref(x_156); -x_161 = l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeMGoal___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__21(x_160, x_159, x_8, x_156, x_147, x_154, x_157, x_149, x_148); +lean_ctor_set(x_160, 3, x_148); +lean_inc(x_153); +lean_inc_ref(x_152); +lean_inc(x_149); +lean_inc_ref(x_147); +lean_inc_ref(x_146); +x_161 = l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeMGoal___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__21(x_160, x_159, x_8, x_146, x_154, x_147, x_149, x_152, x_153); if (lean_obj_tag(x_161) == 0) { lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; @@ -25950,37 +25973,37 @@ x_167 = lean_alloc_closure((void*)(l_Lean_mkApp6), 7, 6); lean_closure_set(x_167, 0, x_166); lean_closure_set(x_167, 1, x_65); lean_closure_set(x_167, 2, x_66); -lean_closure_set(x_167, 3, x_153); +lean_closure_set(x_167, 3, x_148); lean_closure_set(x_167, 4, x_67); lean_closure_set(x_167, 5, x_162); x_126 = x_144; x_127 = x_145; -x_128 = x_146; -x_129 = x_152; -x_130 = x_150; -x_131 = x_151; +x_128 = x_155; +x_129 = x_156; +x_130 = x_157; +x_131 = x_150; x_132 = x_158; x_133 = x_167; -x_134 = x_156; -x_135 = x_147; -x_136 = x_154; -x_137 = x_157; -x_138 = x_149; -x_139 = x_148; +x_134 = x_146; +x_135 = x_154; +x_136 = x_147; +x_137 = x_149; +x_138 = x_152; +x_139 = x_153; x_140 = lean_box(0); goto block_143; } else { -lean_dec(x_157); +lean_dec_ref(x_158); lean_dec_ref(x_156); -lean_dec_ref(x_154); -lean_dec_ref(x_153); +lean_dec_ref(x_155); +lean_dec(x_154); +lean_dec(x_153); lean_dec_ref(x_152); -lean_dec_ref(x_151); -lean_dec_ref(x_149); -lean_dec(x_148); -lean_dec(x_147); +lean_dec(x_149); +lean_dec_ref(x_148); +lean_dec_ref(x_147); lean_dec_ref(x_146); lean_dec_ref(x_145); lean_dec_ref(x_98); @@ -26165,11 +26188,11 @@ lean_dec(x_216); lean_inc_ref(x_7); x_126 = x_206; x_127 = x_207; -x_128 = x_217; -x_129 = x_182; +x_128 = x_7; +x_129 = x_217; x_130 = x_219; -x_131 = x_7; -x_132 = x_220; +x_131 = x_220; +x_132 = x_182; x_133 = x_7; x_134 = x_171; x_135 = x_172; @@ -26189,27 +26212,27 @@ if (x_221 == 0) if (x_63 == 0) { uint8_t x_222; uint8_t x_223; -x_222 = lean_unbox(x_213); -lean_dec(x_213); -x_223 = lean_unbox(x_216); +x_222 = lean_unbox(x_216); lean_dec(x_216); +x_223 = lean_unbox(x_213); +lean_dec(x_213); lean_inc(x_6); lean_inc_ref(x_207); x_144 = x_206; x_145 = x_207; -x_146 = x_217; -x_147 = x_172; -x_148 = x_176; -x_149 = x_175; +x_146 = x_171; +x_147 = x_173; +x_148 = x_207; +x_149 = x_174; x_150 = x_222; -x_151 = x_7; -x_152 = x_182; -x_153 = x_207; -x_154 = x_173; -x_155 = lean_box(0); -x_156 = x_171; -x_157 = x_174; -x_158 = x_223; +x_151 = lean_box(0); +x_152 = x_175; +x_153 = x_176; +x_154 = x_172; +x_155 = x_7; +x_156 = x_217; +x_157 = x_223; +x_158 = x_182; x_159 = x_6; goto block_168; } @@ -26219,26 +26242,26 @@ lean_object* x_224; lean_object* x_225; uint8_t x_226; uint8_t x_227; x_224 = l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__8___closed__8; lean_inc(x_6); x_225 = l_Lean_Name_append(x_6, x_224); -x_226 = lean_unbox(x_213); -lean_dec(x_213); -x_227 = lean_unbox(x_216); +x_226 = lean_unbox(x_216); lean_dec(x_216); +x_227 = lean_unbox(x_213); +lean_dec(x_213); lean_inc_ref(x_207); x_144 = x_206; x_145 = x_207; -x_146 = x_217; -x_147 = x_172; -x_148 = x_176; -x_149 = x_175; +x_146 = x_171; +x_147 = x_173; +x_148 = x_207; +x_149 = x_174; x_150 = x_226; -x_151 = x_7; -x_152 = x_182; -x_153 = x_207; -x_154 = x_173; -x_155 = lean_box(0); -x_156 = x_171; -x_157 = x_174; -x_158 = x_227; +x_151 = lean_box(0); +x_152 = x_175; +x_153 = x_176; +x_154 = x_172; +x_155 = x_7; +x_156 = x_217; +x_157 = x_227; +x_158 = x_182; x_159 = x_225; goto block_168; } @@ -26246,27 +26269,27 @@ goto block_168; else { uint8_t x_228; uint8_t x_229; -x_228 = lean_unbox(x_213); -lean_dec(x_213); -x_229 = lean_unbox(x_216); +x_228 = lean_unbox(x_216); lean_dec(x_216); +x_229 = lean_unbox(x_213); +lean_dec(x_213); lean_inc(x_6); lean_inc_ref(x_207); x_144 = x_206; x_145 = x_207; -x_146 = x_217; -x_147 = x_172; -x_148 = x_176; -x_149 = x_175; +x_146 = x_171; +x_147 = x_173; +x_148 = x_207; +x_149 = x_174; x_150 = x_228; -x_151 = x_7; -x_152 = x_182; -x_153 = x_207; -x_154 = x_173; -x_155 = lean_box(0); -x_156 = x_171; -x_157 = x_174; -x_158 = x_229; +x_151 = lean_box(0); +x_152 = x_175; +x_153 = x_176; +x_154 = x_172; +x_155 = x_7; +x_156 = x_217; +x_157 = x_229; +x_158 = x_182; x_159 = x_6; goto block_168; } @@ -26286,11 +26309,11 @@ lean_dec(x_216); lean_inc_ref(x_7); x_126 = x_206; x_127 = x_207; -x_128 = x_217; -x_129 = x_182; +x_128 = x_7; +x_129 = x_217; x_130 = x_230; -x_131 = x_7; -x_132 = x_231; +x_131 = x_231; +x_132 = x_182; x_133 = x_7; x_134 = x_171; x_135 = x_172; @@ -26504,11 +26527,11 @@ lean_dec(x_264); lean_inc_ref(x_7); x_126 = x_253; x_127 = x_254; -x_128 = x_265; -x_129 = x_182; +x_128 = x_7; +x_129 = x_265; x_130 = x_267; -x_131 = x_7; -x_132 = x_268; +x_131 = x_268; +x_132 = x_182; x_133 = x_7; x_134 = x_171; x_135 = x_172; @@ -26528,27 +26551,27 @@ if (x_269 == 0) if (x_63 == 0) { uint8_t x_270; uint8_t x_271; -x_270 = lean_unbox(x_261); -lean_dec(x_261); -x_271 = lean_unbox(x_264); +x_270 = lean_unbox(x_264); lean_dec(x_264); +x_271 = lean_unbox(x_261); +lean_dec(x_261); lean_inc(x_6); lean_inc_ref(x_254); x_144 = x_253; x_145 = x_254; -x_146 = x_265; -x_147 = x_172; -x_148 = x_176; -x_149 = x_175; +x_146 = x_171; +x_147 = x_173; +x_148 = x_254; +x_149 = x_174; x_150 = x_270; -x_151 = x_7; -x_152 = x_182; -x_153 = x_254; -x_154 = x_173; -x_155 = lean_box(0); -x_156 = x_171; -x_157 = x_174; -x_158 = x_271; +x_151 = lean_box(0); +x_152 = x_175; +x_153 = x_176; +x_154 = x_172; +x_155 = x_7; +x_156 = x_265; +x_157 = x_271; +x_158 = x_182; x_159 = x_6; goto block_168; } @@ -26558,26 +26581,26 @@ lean_object* x_272; lean_object* x_273; uint8_t x_274; uint8_t x_275; x_272 = l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__8___closed__8; lean_inc(x_6); x_273 = l_Lean_Name_append(x_6, x_272); -x_274 = lean_unbox(x_261); -lean_dec(x_261); -x_275 = lean_unbox(x_264); +x_274 = lean_unbox(x_264); lean_dec(x_264); +x_275 = lean_unbox(x_261); +lean_dec(x_261); lean_inc_ref(x_254); x_144 = x_253; x_145 = x_254; -x_146 = x_265; -x_147 = x_172; -x_148 = x_176; -x_149 = x_175; +x_146 = x_171; +x_147 = x_173; +x_148 = x_254; +x_149 = x_174; x_150 = x_274; -x_151 = x_7; -x_152 = x_182; -x_153 = x_254; -x_154 = x_173; -x_155 = lean_box(0); -x_156 = x_171; -x_157 = x_174; -x_158 = x_275; +x_151 = lean_box(0); +x_152 = x_175; +x_153 = x_176; +x_154 = x_172; +x_155 = x_7; +x_156 = x_265; +x_157 = x_275; +x_158 = x_182; x_159 = x_273; goto block_168; } @@ -26585,27 +26608,27 @@ goto block_168; else { uint8_t x_276; uint8_t x_277; -x_276 = lean_unbox(x_261); -lean_dec(x_261); -x_277 = lean_unbox(x_264); +x_276 = lean_unbox(x_264); lean_dec(x_264); +x_277 = lean_unbox(x_261); +lean_dec(x_261); lean_inc(x_6); lean_inc_ref(x_254); x_144 = x_253; x_145 = x_254; -x_146 = x_265; -x_147 = x_172; -x_148 = x_176; -x_149 = x_175; +x_146 = x_171; +x_147 = x_173; +x_148 = x_254; +x_149 = x_174; x_150 = x_276; -x_151 = x_7; -x_152 = x_182; -x_153 = x_254; -x_154 = x_173; -x_155 = lean_box(0); -x_156 = x_171; -x_157 = x_174; -x_158 = x_277; +x_151 = lean_box(0); +x_152 = x_175; +x_153 = x_176; +x_154 = x_172; +x_155 = x_7; +x_156 = x_265; +x_157 = x_277; +x_158 = x_182; x_159 = x_6; goto block_168; } @@ -26625,11 +26648,11 @@ lean_dec(x_264); lean_inc_ref(x_7); x_126 = x_253; x_127 = x_254; -x_128 = x_265; -x_129 = x_182; +x_128 = x_7; +x_129 = x_265; x_130 = x_278; -x_131 = x_7; -x_132 = x_279; +x_131 = x_279; +x_132 = x_182; x_133 = x_7; x_134 = x_171; x_135 = x_172; @@ -26934,11 +26957,11 @@ lean_dec(x_352); lean_inc_ref(x_7); x_126 = x_341; x_127 = x_342; -x_128 = x_353; -x_129 = x_182; +x_128 = x_7; +x_129 = x_353; x_130 = x_355; -x_131 = x_7; -x_132 = x_356; +x_131 = x_356; +x_132 = x_182; x_133 = x_7; x_134 = x_171; x_135 = x_172; @@ -26958,27 +26981,27 @@ if (x_357 == 0) if (x_63 == 0) { uint8_t x_358; uint8_t x_359; -x_358 = lean_unbox(x_349); -lean_dec(x_349); -x_359 = lean_unbox(x_352); +x_358 = lean_unbox(x_352); lean_dec(x_352); +x_359 = lean_unbox(x_349); +lean_dec(x_349); lean_inc(x_6); lean_inc_ref(x_342); x_144 = x_341; x_145 = x_342; -x_146 = x_353; -x_147 = x_172; -x_148 = x_176; -x_149 = x_175; +x_146 = x_171; +x_147 = x_173; +x_148 = x_342; +x_149 = x_174; x_150 = x_358; -x_151 = x_7; -x_152 = x_182; -x_153 = x_342; -x_154 = x_173; -x_155 = lean_box(0); -x_156 = x_171; -x_157 = x_174; -x_158 = x_359; +x_151 = lean_box(0); +x_152 = x_175; +x_153 = x_176; +x_154 = x_172; +x_155 = x_7; +x_156 = x_353; +x_157 = x_359; +x_158 = x_182; x_159 = x_6; goto block_168; } @@ -26988,26 +27011,26 @@ lean_object* x_360; lean_object* x_361; uint8_t x_362; uint8_t x_363; x_360 = l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__8___closed__8; lean_inc(x_6); x_361 = l_Lean_Name_append(x_6, x_360); -x_362 = lean_unbox(x_349); -lean_dec(x_349); -x_363 = lean_unbox(x_352); +x_362 = lean_unbox(x_352); lean_dec(x_352); +x_363 = lean_unbox(x_349); +lean_dec(x_349); lean_inc_ref(x_342); x_144 = x_341; x_145 = x_342; -x_146 = x_353; -x_147 = x_172; -x_148 = x_176; -x_149 = x_175; +x_146 = x_171; +x_147 = x_173; +x_148 = x_342; +x_149 = x_174; x_150 = x_362; -x_151 = x_7; -x_152 = x_182; -x_153 = x_342; -x_154 = x_173; -x_155 = lean_box(0); -x_156 = x_171; -x_157 = x_174; -x_158 = x_363; +x_151 = lean_box(0); +x_152 = x_175; +x_153 = x_176; +x_154 = x_172; +x_155 = x_7; +x_156 = x_353; +x_157 = x_363; +x_158 = x_182; x_159 = x_361; goto block_168; } @@ -27015,27 +27038,27 @@ goto block_168; else { uint8_t x_364; uint8_t x_365; -x_364 = lean_unbox(x_349); -lean_dec(x_349); -x_365 = lean_unbox(x_352); +x_364 = lean_unbox(x_352); lean_dec(x_352); +x_365 = lean_unbox(x_349); +lean_dec(x_349); lean_inc(x_6); lean_inc_ref(x_342); x_144 = x_341; x_145 = x_342; -x_146 = x_353; -x_147 = x_172; -x_148 = x_176; -x_149 = x_175; +x_146 = x_171; +x_147 = x_173; +x_148 = x_342; +x_149 = x_174; x_150 = x_364; -x_151 = x_7; -x_152 = x_182; -x_153 = x_342; -x_154 = x_173; -x_155 = lean_box(0); -x_156 = x_171; -x_157 = x_174; -x_158 = x_365; +x_151 = lean_box(0); +x_152 = x_175; +x_153 = x_176; +x_154 = x_172; +x_155 = x_7; +x_156 = x_353; +x_157 = x_365; +x_158 = x_182; x_159 = x_6; goto block_168; } @@ -27055,11 +27078,11 @@ lean_dec(x_352); lean_inc_ref(x_7); x_126 = x_341; x_127 = x_342; -x_128 = x_353; -x_129 = x_182; +x_128 = x_7; +x_129 = x_353; x_130 = x_366; -x_131 = x_7; -x_132 = x_367; +x_131 = x_367; +x_132 = x_182; x_133 = x_7; x_134 = x_171; x_135 = x_172; @@ -27353,7 +27376,7 @@ goto block_377; } else { -uint8_t x_403; uint8_t x_404; uint8_t x_405; uint8_t x_406; uint8_t x_407; uint8_t x_408; uint8_t x_409; uint8_t x_410; uint8_t x_411; uint8_t x_412; uint8_t x_413; uint8_t x_414; uint8_t x_415; uint8_t x_416; uint8_t x_417; uint8_t x_418; uint8_t x_419; uint8_t x_420; uint8_t x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; uint8_t x_428; uint8_t x_429; uint8_t x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; uint8_t x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; uint8_t x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; uint8_t x_473; lean_object* x_474; uint8_t x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; uint8_t x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; uint8_t x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; uint8_t x_501; lean_object* x_502; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; uint8_t x_624; lean_object* x_625; lean_object* x_633; lean_object* x_634; uint64_t x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; +uint8_t x_403; uint8_t x_404; uint8_t x_405; uint8_t x_406; uint8_t x_407; uint8_t x_408; uint8_t x_409; uint8_t x_410; uint8_t x_411; uint8_t x_412; uint8_t x_413; uint8_t x_414; uint8_t x_415; uint8_t x_416; uint8_t x_417; uint8_t x_418; uint8_t x_419; uint8_t x_420; uint8_t x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; uint8_t x_428; uint8_t x_429; uint8_t x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; uint8_t x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; uint8_t x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; uint8_t x_473; uint8_t x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; uint8_t x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; uint8_t x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; uint8_t x_500; lean_object* x_501; lean_object* x_502; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; uint8_t x_624; lean_object* x_625; lean_object* x_633; lean_object* x_634; uint64_t x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; x_403 = lean_ctor_get_uint8(x_76, 0); x_404 = lean_ctor_get_uint8(x_76, 1); x_405 = lean_ctor_get_uint8(x_76, 2); @@ -27539,10 +27562,10 @@ return x_645; { lean_object* x_455; lean_inc(x_438); -lean_inc_ref(x_449); +lean_inc_ref(x_452); lean_inc_ref(x_74); lean_inc_ref(x_72); -x_455 = l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20(x_72, x_74, x_449, x_438, x_454, x_448, x_447, x_452, x_453, x_445, x_451); +x_455 = l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20(x_72, x_74, x_452, x_438, x_454, x_453, x_446, x_444, x_449, x_451, x_445); if (lean_obj_tag(x_455) == 0) { lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; @@ -27561,12 +27584,12 @@ lean_ctor_set(x_459, 0, x_64); lean_ctor_set(x_459, 1, x_458); lean_inc_ref(x_459); x_460 = l_Lean_mkConst(x_457, x_459); -lean_inc_ref(x_449); +lean_inc_ref(x_452); lean_inc_ref(x_5); lean_inc_ref(x_72); lean_inc_ref(x_74); lean_inc_ref(x_460); -x_461 = l_Lean_mkApp4(x_460, x_74, x_72, x_5, x_449); +x_461 = l_Lean_mkApp4(x_460, x_74, x_72, x_5, x_452); lean_inc(x_438); lean_inc_ref(x_5); lean_inc_ref(x_72); @@ -27575,7 +27598,7 @@ x_462 = l_Lean_mkApp4(x_460, x_74, x_72, x_5, x_438); x_463 = l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__8___closed__5; lean_inc_ref(x_459); x_464 = l_Lean_mkConst(x_463, x_459); -x_465 = l_Lean_mkApp6(x_464, x_74, x_72, x_5, x_449, x_438, x_456); +x_465 = l_Lean_mkApp6(x_464, x_74, x_72, x_5, x_452, x_438, x_456); x_466 = lean_box(x_442); x_467 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__5___boxed), 11, 10); lean_closure_set(x_467, 0, x_60); @@ -27588,7 +27611,7 @@ lean_closure_set(x_467, 6, x_462); lean_closure_set(x_467, 7, x_465); lean_closure_set(x_467, 8, x_65); lean_closure_set(x_467, 9, x_443); -x_17 = x_446; +x_17 = x_447; x_18 = x_450; x_19 = x_467; x_20 = lean_box(0); @@ -27596,9 +27619,9 @@ goto block_24; } else { +lean_dec_ref(x_452); lean_dec_ref(x_450); -lean_dec_ref(x_449); -lean_dec_ref(x_446); +lean_dec_ref(x_447); lean_dec_ref(x_443); lean_dec_ref(x_441); lean_dec(x_438); @@ -27613,7 +27636,7 @@ return x_455; } block_486: { -if (x_475 == 0) +if (x_474 == 0) { if (x_63 == 0) { @@ -27623,7 +27646,7 @@ lean_dec(x_480); lean_dec_ref(x_479); lean_dec(x_478); lean_dec_ref(x_477); -lean_dec_ref(x_472); +lean_dec_ref(x_475); lean_dec_ref(x_470); lean_dec_ref(x_441); lean_dec(x_438); @@ -27634,31 +27657,31 @@ lean_dec(x_64); lean_dec(x_29); lean_dec(x_6); lean_dec_ref(x_5); -x_17 = x_471; +x_17 = x_472; x_18 = x_476; -x_19 = x_474; +x_19 = x_471; x_20 = lean_box(0); goto block_24; } else { -lean_dec_ref(x_474); +lean_dec_ref(x_471); if (x_473 == 0) { if (x_63 == 0) { x_442 = x_469; x_443 = x_470; -x_444 = lean_box(0); -x_445 = x_481; -x_446 = x_471; -x_447 = x_478; -x_448 = x_477; -x_449 = x_472; +x_444 = x_479; +x_445 = x_482; +x_446 = x_478; +x_447 = x_472; +x_448 = lean_box(0); +x_449 = x_480; x_450 = x_476; -x_451 = x_482; -x_452 = x_479; -x_453 = x_480; +x_451 = x_481; +x_452 = x_475; +x_453 = x_477; x_454 = x_6; goto block_468; } @@ -27669,16 +27692,16 @@ x_484 = l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePos x_485 = l_Lean_Name_append(x_6, x_484); x_442 = x_469; x_443 = x_470; -x_444 = lean_box(0); -x_445 = x_481; -x_446 = x_471; -x_447 = x_478; -x_448 = x_477; -x_449 = x_472; +x_444 = x_479; +x_445 = x_482; +x_446 = x_478; +x_447 = x_472; +x_448 = lean_box(0); +x_449 = x_480; x_450 = x_476; -x_451 = x_482; -x_452 = x_479; -x_453 = x_480; +x_451 = x_481; +x_452 = x_475; +x_453 = x_477; x_454 = x_485; goto block_468; } @@ -27687,16 +27710,16 @@ else { x_442 = x_469; x_443 = x_470; -x_444 = lean_box(0); -x_445 = x_481; -x_446 = x_471; -x_447 = x_478; -x_448 = x_477; -x_449 = x_472; +x_444 = x_479; +x_445 = x_482; +x_446 = x_478; +x_447 = x_472; +x_448 = lean_box(0); +x_449 = x_480; x_450 = x_476; -x_451 = x_482; -x_452 = x_479; -x_453 = x_480; +x_451 = x_481; +x_452 = x_475; +x_453 = x_477; x_454 = x_6; goto block_468; } @@ -27710,7 +27733,7 @@ lean_dec(x_480); lean_dec_ref(x_479); lean_dec(x_478); lean_dec_ref(x_477); -lean_dec_ref(x_472); +lean_dec_ref(x_475); lean_dec_ref(x_470); lean_dec_ref(x_441); lean_dec(x_438); @@ -27721,9 +27744,9 @@ lean_dec(x_64); lean_dec(x_29); lean_dec(x_6); lean_dec_ref(x_5); -x_17 = x_471; +x_17 = x_472; x_18 = x_476; -x_19 = x_474; +x_19 = x_471; x_20 = lean_box(0); goto block_24; } @@ -27731,7 +27754,7 @@ goto block_24; block_511: { lean_object* x_503; lean_object* x_504; -lean_inc_ref(x_496); +lean_inc_ref(x_491); lean_inc_ref(x_66); lean_inc_ref(x_65); lean_inc(x_64); @@ -27743,13 +27766,13 @@ if (lean_is_scalar(x_68)) { lean_ctor_set(x_503, 0, x_64); lean_ctor_set(x_503, 1, x_65); lean_ctor_set(x_503, 2, x_66); -lean_ctor_set(x_503, 3, x_496); -lean_inc(x_491); -lean_inc_ref(x_492); -lean_inc(x_500); -lean_inc_ref(x_497); -lean_inc_ref(x_499); -x_504 = l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeMGoal___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__21(x_503, x_502, x_8, x_499, x_490, x_497, x_500, x_492, x_491); +lean_ctor_set(x_503, 3, x_491); +lean_inc(x_496); +lean_inc_ref(x_495); +lean_inc(x_492); +lean_inc_ref(x_490); +lean_inc_ref(x_489); +x_504 = l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeMGoal___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__21(x_503, x_502, x_8, x_489, x_497, x_490, x_492, x_495, x_496); 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; @@ -27773,37 +27796,37 @@ x_510 = lean_alloc_closure((void*)(l_Lean_mkApp6), 7, 6); lean_closure_set(x_510, 0, x_509); lean_closure_set(x_510, 1, x_65); lean_closure_set(x_510, 2, x_66); -lean_closure_set(x_510, 3, x_496); +lean_closure_set(x_510, 3, x_491); lean_closure_set(x_510, 4, x_67); lean_closure_set(x_510, 5, x_505); x_469 = x_487; x_470 = x_488; -x_471 = x_489; -x_472 = x_495; -x_473 = x_493; -x_474 = x_494; +x_471 = x_498; +x_472 = x_499; +x_473 = x_500; +x_474 = x_493; x_475 = x_501; x_476 = x_510; -x_477 = x_499; -x_478 = x_490; -x_479 = x_497; -x_480 = x_500; -x_481 = x_492; -x_482 = x_491; +x_477 = x_489; +x_478 = x_497; +x_479 = x_490; +x_480 = x_492; +x_481 = x_495; +x_482 = x_496; x_483 = lean_box(0); goto block_486; } else { -lean_dec(x_500); +lean_dec_ref(x_501); lean_dec_ref(x_499); -lean_dec_ref(x_497); -lean_dec_ref(x_496); +lean_dec_ref(x_498); +lean_dec(x_497); +lean_dec(x_496); lean_dec_ref(x_495); -lean_dec_ref(x_494); -lean_dec_ref(x_492); -lean_dec(x_491); -lean_dec(x_490); +lean_dec(x_492); +lean_dec_ref(x_491); +lean_dec_ref(x_490); lean_dec_ref(x_489); lean_dec_ref(x_488); lean_dec_ref(x_441); @@ -28068,11 +28091,11 @@ lean_dec(x_594); lean_inc_ref(x_7); x_469 = x_583; x_470 = x_584; -x_471 = x_595; -x_472 = x_525; +x_471 = x_7; +x_472 = x_595; x_473 = x_597; -x_474 = x_7; -x_475 = x_598; +x_474 = x_598; +x_475 = x_525; x_476 = x_7; x_477 = x_514; x_478 = x_515; @@ -28092,27 +28115,27 @@ if (x_599 == 0) if (x_63 == 0) { uint8_t x_600; uint8_t x_601; -x_600 = lean_unbox(x_591); -lean_dec(x_591); -x_601 = lean_unbox(x_594); +x_600 = lean_unbox(x_594); lean_dec(x_594); +x_601 = lean_unbox(x_591); +lean_dec(x_591); lean_inc(x_6); lean_inc_ref(x_584); x_487 = x_583; x_488 = x_584; -x_489 = x_595; -x_490 = x_515; -x_491 = x_519; -x_492 = x_518; +x_489 = x_514; +x_490 = x_516; +x_491 = x_584; +x_492 = x_517; x_493 = x_600; -x_494 = x_7; -x_495 = x_525; -x_496 = x_584; -x_497 = x_516; -x_498 = lean_box(0); -x_499 = x_514; -x_500 = x_517; -x_501 = x_601; +x_494 = lean_box(0); +x_495 = x_518; +x_496 = x_519; +x_497 = x_515; +x_498 = x_7; +x_499 = x_595; +x_500 = x_601; +x_501 = x_525; x_502 = x_6; goto block_511; } @@ -28122,26 +28145,26 @@ lean_object* x_602; lean_object* x_603; uint8_t x_604; uint8_t x_605; x_602 = l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__8___closed__8; lean_inc(x_6); x_603 = l_Lean_Name_append(x_6, x_602); -x_604 = lean_unbox(x_591); -lean_dec(x_591); -x_605 = lean_unbox(x_594); +x_604 = lean_unbox(x_594); lean_dec(x_594); +x_605 = lean_unbox(x_591); +lean_dec(x_591); lean_inc_ref(x_584); x_487 = x_583; x_488 = x_584; -x_489 = x_595; -x_490 = x_515; -x_491 = x_519; -x_492 = x_518; +x_489 = x_514; +x_490 = x_516; +x_491 = x_584; +x_492 = x_517; x_493 = x_604; -x_494 = x_7; -x_495 = x_525; -x_496 = x_584; -x_497 = x_516; -x_498 = lean_box(0); -x_499 = x_514; -x_500 = x_517; -x_501 = x_605; +x_494 = lean_box(0); +x_495 = x_518; +x_496 = x_519; +x_497 = x_515; +x_498 = x_7; +x_499 = x_595; +x_500 = x_605; +x_501 = x_525; x_502 = x_603; goto block_511; } @@ -28149,27 +28172,27 @@ goto block_511; else { uint8_t x_606; uint8_t x_607; -x_606 = lean_unbox(x_591); -lean_dec(x_591); -x_607 = lean_unbox(x_594); +x_606 = lean_unbox(x_594); lean_dec(x_594); +x_607 = lean_unbox(x_591); +lean_dec(x_591); lean_inc(x_6); lean_inc_ref(x_584); x_487 = x_583; x_488 = x_584; -x_489 = x_595; -x_490 = x_515; -x_491 = x_519; -x_492 = x_518; +x_489 = x_514; +x_490 = x_516; +x_491 = x_584; +x_492 = x_517; x_493 = x_606; -x_494 = x_7; -x_495 = x_525; -x_496 = x_584; -x_497 = x_516; -x_498 = lean_box(0); -x_499 = x_514; -x_500 = x_517; -x_501 = x_607; +x_494 = lean_box(0); +x_495 = x_518; +x_496 = x_519; +x_497 = x_515; +x_498 = x_7; +x_499 = x_595; +x_500 = x_607; +x_501 = x_525; x_502 = x_6; goto block_511; } @@ -28189,11 +28212,11 @@ lean_dec(x_594); lean_inc_ref(x_7); x_469 = x_583; x_470 = x_584; -x_471 = x_595; -x_472 = x_525; +x_471 = x_7; +x_472 = x_595; x_473 = x_608; -x_474 = x_7; -x_475 = x_609; +x_474 = x_609; +x_475 = x_525; x_476 = x_7; x_477 = x_514; x_478 = x_515; @@ -29619,50 +29642,51 @@ lean_dec_ref(x_2); return x_11; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___lam__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___lam__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_14 = lean_array_get_borrowed(x_1, x_2, x_5); -x_15 = l_Array_append___redArg(x_3, x_6); -lean_inc(x_14); -x_16 = l_Lean_mkAppN(x_14, x_15); -lean_dec_ref(x_15); -x_17 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_17, 0, x_16); -return x_17; +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_15 = lean_array_get_borrowed(x_1, x_2, x_6); +x_16 = l_Array_append___redArg(x_3, x_7); +lean_inc(x_15); +x_17 = l_Lean_mkAppN(x_15, x_16); +lean_dec_ref(x_16); +x_18 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_18, 0, x_17); +return x_18; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___lam__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___lam__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_14; -x_14 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___lam__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_12); -lean_dec_ref(x_11); -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec(x_8); +lean_object* x_15; +x_15 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___lam__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_13); +lean_dec_ref(x_12); +lean_dec(x_11); +lean_dec_ref(x_10); +lean_dec(x_9); +lean_dec_ref(x_8); lean_dec_ref(x_7); -lean_dec_ref(x_6); -lean_dec(x_5); +lean_dec(x_6); +lean_dec_ref(x_5); lean_dec(x_4); lean_dec_ref(x_2); -return x_14; +return x_15; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_10 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_4); lean_dec_ref(x_3); lean_dec(x_2); return x_10; } } -static lean_object* _init_l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59___lam__1___closed__0() { +static lean_object* _init_l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__1___closed__0() { _start: { lean_object* x_1; @@ -29670,16 +29694,16 @@ x_1 = lean_mk_string_unchecked("isTrue", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59___lam__1___closed__1() { +static lean_object* _init_l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__1___closed__1() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59___lam__1___closed__0; +x_1 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__1___closed__0; x_2 = l_Lean_Name_mkStr1(x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62___lam__1___closed__0() { +static lean_object* _init_l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__63___lam__1___closed__0() { _start: { lean_object* x_1; @@ -29687,16 +29711,16 @@ x_1 = lean_mk_string_unchecked("isFalse", 7, 7); return x_1; } } -static lean_object* _init_l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62___lam__1___closed__1() { +static lean_object* _init_l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__63___lam__1___closed__1() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62___lam__1___closed__0; +x_1 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__63___lam__1___closed__0; x_2 = l_Lean_Name_mkStr1(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__0() { +static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__0() { _start: { lean_object* x_1; @@ -29704,23 +29728,23 @@ x_1 = lean_mk_string_unchecked("ite", 3, 3); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__1() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__0; +x_1 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__0; x_2 = l_Lean_Name_mkStr1(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { uint8_t x_14; uint8_t x_15; uint8_t x_16; lean_object* x_17; x_14 = lean_unbox(x_3); x_15 = lean_unbox(x_4); x_16 = lean_unbox(x_5); -x_17 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59___lam__0(x_1, x_2, x_14, x_15, x_16, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_17 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__0(x_1, x_2, x_14, x_15, x_16, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_12); lean_dec_ref(x_11); lean_dec(x_10); @@ -29732,202 +29756,204 @@ lean_dec_ref(x_1); return x_17; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_12 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59___lam__1___closed__1; -x_13 = lean_unsigned_to_nat(0u); -x_14 = l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__28_spec__33___redArg___closed__0; -lean_inc(x_10); -lean_inc_ref(x_9); -lean_inc(x_8); -lean_inc_ref(x_7); -lean_inc(x_3); -lean_inc_ref(x_2); -x_15 = lean_apply_10(x_1, x_12, x_13, x_14, x_2, x_3, x_7, x_8, x_9, x_10, lean_box(0)); -if (lean_obj_tag(x_15) == 0) +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__1___closed__1; +x_14 = lean_unsigned_to_nat(0u); +x_15 = l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__28_spec__33___redArg___closed__0; +lean_inc(x_11); +lean_inc_ref(x_10); +lean_inc(x_9); +lean_inc_ref(x_8); +lean_inc(x_4); +lean_inc_ref(x_3); +x_16 = lean_apply_11(x_1, x_13, x_2, x_14, x_15, x_3, x_4, x_8, x_9, x_10, x_11, lean_box(0)); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -lean_dec_ref(x_15); -x_17 = lean_mk_empty_array_with_capacity(x_4); -x_18 = lean_array_push(x_17, x_6); -x_19 = 0; -x_20 = 1; -x_21 = lean_box(x_19); -x_22 = lean_box(x_5); -x_23 = lean_box(x_20); -x_24 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59___lam__0___boxed), 13, 5); -lean_closure_set(x_24, 0, x_18); -lean_closure_set(x_24, 1, x_16); -lean_closure_set(x_24, 2, x_21); -lean_closure_set(x_24, 3, x_22); -lean_closure_set(x_24, 4, x_23); -x_25 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_24, x_2, x_3, x_7, x_8, x_9, x_10); -lean_dec(x_3); -return x_25; +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec_ref(x_16); +x_18 = lean_mk_empty_array_with_capacity(x_5); +x_19 = lean_array_push(x_18, x_7); +x_20 = 0; +x_21 = 1; +x_22 = lean_box(x_20); +x_23 = lean_box(x_6); +x_24 = lean_box(x_21); +x_25 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__0___boxed), 13, 5); +lean_closure_set(x_25, 0, x_19); +lean_closure_set(x_25, 1, x_17); +lean_closure_set(x_25, 2, x_22); +lean_closure_set(x_25, 3, x_23); +lean_closure_set(x_25, 4, x_24); +x_26 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_25, x_3, x_4, x_8, x_9, x_10, x_11); +lean_dec(x_4); +return x_26; } else { -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec(x_8); +lean_dec(x_11); +lean_dec_ref(x_10); +lean_dec(x_9); +lean_dec_ref(x_8); lean_dec_ref(x_7); -lean_dec_ref(x_6); -lean_dec(x_3); -lean_dec_ref(x_2); -return x_15; +lean_dec(x_4); +lean_dec_ref(x_3); +return x_16; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -uint8_t x_12; lean_object* x_13; -x_12 = lean_unbox(x_5); -x_13 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59___lam__1(x_1, x_2, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_4); -return x_13; +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_6); +x_14 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__1(x_1, x_2, x_3, x_4, x_5, x_13, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_5); +return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, uint8_t x_6, lean_object* x_7, uint8_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_box(x_3); -x_16 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59___lam__1___boxed), 11, 5); -lean_closure_set(x_16, 0, x_1); -lean_closure_set(x_16, 1, x_8); -lean_closure_set(x_16, 2, x_9); -lean_closure_set(x_16, 3, x_2); -lean_closure_set(x_16, 4, x_15); -x_17 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp(lean_box(0), x_4, x_5, x_6, x_16, x_7, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_17) == 0) +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_box(x_4); +x_17 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__1___boxed), 12, 6); +lean_closure_set(x_17, 0, x_1); +lean_closure_set(x_17, 1, x_2); +lean_closure_set(x_17, 2, x_9); +lean_closure_set(x_17, 3, x_10); +lean_closure_set(x_17, 4, x_3); +lean_closure_set(x_17, 5, x_16); +x_18 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp(lean_box(0), x_5, x_6, x_7, x_17, x_8, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_18) == 0) { -return x_17; +return x_18; } else { -uint8_t x_18; -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 0) +uint8_t x_19; +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) { -return x_17; +return x_18; } else { -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_17, 0); -lean_inc(x_19); -lean_dec(x_17); -x_20 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_20, 0, x_19); -return x_20; +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_18, 0); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_21, 0, x_20); +return x_21; } } } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__61___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62___lam__1___closed__1; -x_13 = l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__28_spec__33___redArg___closed__0; -lean_inc(x_10); -lean_inc_ref(x_9); -lean_inc(x_8); -lean_inc_ref(x_7); -lean_inc(x_4); -lean_inc_ref(x_3); -lean_inc(x_2); -x_14 = lean_apply_10(x_1, x_12, x_2, x_13, x_3, x_4, x_7, x_8, x_9, x_10, lean_box(0)); -if (lean_obj_tag(x_14) == 0) +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__63___lam__1___closed__1; +x_14 = l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__28_spec__33___redArg___closed__0; +lean_inc(x_11); +lean_inc_ref(x_10); +lean_inc(x_9); +lean_inc_ref(x_8); +lean_inc(x_5); +lean_inc_ref(x_4); +lean_inc(x_3); +x_15 = lean_apply_11(x_1, x_13, x_2, x_3, x_14, x_4, x_5, x_8, x_9, x_10, x_11, lean_box(0)); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; uint8_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -lean_dec_ref(x_14); -x_16 = lean_mk_empty_array_with_capacity(x_2); -lean_dec(x_2); -x_17 = lean_array_push(x_16, x_6); -x_18 = 0; -x_19 = 1; -x_20 = lean_box(x_18); -x_21 = lean_box(x_5); -x_22 = lean_box(x_19); -x_23 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59___lam__0___boxed), 13, 5); -lean_closure_set(x_23, 0, x_17); -lean_closure_set(x_23, 1, x_15); -lean_closure_set(x_23, 2, x_20); -lean_closure_set(x_23, 3, x_21); -lean_closure_set(x_23, 4, x_22); -x_24 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_23, x_3, x_4, x_7, x_8, x_9, x_10); -lean_dec(x_4); -return x_24; +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +lean_dec_ref(x_15); +x_17 = lean_mk_empty_array_with_capacity(x_3); +lean_dec(x_3); +x_18 = lean_array_push(x_17, x_7); +x_19 = 0; +x_20 = 1; +x_21 = lean_box(x_19); +x_22 = lean_box(x_6); +x_23 = lean_box(x_20); +x_24 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__0___boxed), 13, 5); +lean_closure_set(x_24, 0, x_18); +lean_closure_set(x_24, 1, x_16); +lean_closure_set(x_24, 2, x_21); +lean_closure_set(x_24, 3, x_22); +lean_closure_set(x_24, 4, x_23); +x_25 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_24, x_4, x_5, x_8, x_9, x_10, x_11); +lean_dec(x_5); +return x_25; } else { -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec(x_8); +lean_dec(x_11); +lean_dec_ref(x_10); +lean_dec(x_9); +lean_dec_ref(x_8); lean_dec_ref(x_7); -lean_dec_ref(x_6); -lean_dec(x_4); -lean_dec_ref(x_3); -lean_dec(x_2); -return x_14; +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec(x_3); +return x_15; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__61___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -uint8_t x_12; lean_object* x_13; -x_12 = lean_unbox(x_5); -x_13 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__1(x_1, x_2, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10); -return x_13; +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_6); +x_14 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__61___lam__1(x_1, x_2, x_3, x_4, x_5, x_13, x_7, x_8, x_9, x_10, x_11); +return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__61(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, uint8_t x_6, lean_object* x_7, uint8_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_box(x_3); -x_16 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__1___boxed), 11, 5); -lean_closure_set(x_16, 0, x_1); -lean_closure_set(x_16, 1, x_2); -lean_closure_set(x_16, 2, x_8); -lean_closure_set(x_16, 3, x_9); -lean_closure_set(x_16, 4, x_15); -x_17 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp(lean_box(0), x_4, x_5, x_6, x_16, x_7, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_17) == 0) +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_box(x_4); +x_17 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__61___lam__1___boxed), 12, 6); +lean_closure_set(x_17, 0, x_1); +lean_closure_set(x_17, 1, x_2); +lean_closure_set(x_17, 2, x_3); +lean_closure_set(x_17, 3, x_9); +lean_closure_set(x_17, 4, x_10); +lean_closure_set(x_17, 5, x_16); +x_18 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp(lean_box(0), x_5, x_6, x_7, x_17, x_8, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_18) == 0) { -return x_17; +return x_18; } else { -uint8_t x_18; -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 0) +uint8_t x_19; +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) { -return x_17; +return x_18; } else { -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_17, 0); -lean_inc(x_19); -lean_dec(x_17); -x_20 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_20, 0, x_19); -return x_20; +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_18, 0); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_21, 0, x_20); +return x_21; } } } } -static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__2() { +static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__2() { _start: { lean_object* x_1; @@ -29935,171 +29961,79 @@ x_1 = lean_mk_string_unchecked("dite", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__3() { +static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__2; +x_1 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__2; x_2 = l_Lean_Name_mkStr1(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__61___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_10 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59___lam__1___closed__1; -x_11 = lean_unsigned_to_nat(0u); -x_12 = l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___lam__4___closed__0; -x_13 = lean_array_push(x_12, x_4); -lean_inc(x_8); -lean_inc_ref(x_7); -lean_inc(x_6); -lean_inc_ref(x_5); -lean_inc(x_3); -lean_inc_ref(x_2); -lean_inc_ref(x_13); -x_14 = lean_apply_10(x_1, x_10, x_11, x_13, x_2, x_3, x_5, x_6, x_7, x_8, lean_box(0)); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; uint8_t x_16; uint8_t x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -lean_dec_ref(x_14); -x_16 = 0; -x_17 = 1; -x_18 = 1; -x_19 = lean_box(x_16); -x_20 = lean_box(x_17); -x_21 = lean_box(x_18); -x_22 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_ProofMode_mIntroForall___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__4___lam__0___boxed), 13, 5); -lean_closure_set(x_22, 0, x_13); -lean_closure_set(x_22, 1, x_15); -lean_closure_set(x_22, 2, x_19); -lean_closure_set(x_22, 3, x_20); -lean_closure_set(x_22, 4, x_21); -x_23 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_22, x_2, x_3, x_5, x_6, x_7, x_8); -lean_dec(x_3); -return x_23; -} -else -{ -lean_dec_ref(x_13); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_3); -lean_dec_ref(x_2); -return x_14; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__61___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; -x_10 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__61___lam__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__61(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_13; lean_object* x_14; -x_13 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__61___lam__1___boxed), 9, 3); -lean_closure_set(x_13, 0, x_1); -lean_closure_set(x_13, 1, x_6); -lean_closure_set(x_13, 2, x_7); -x_14 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp(lean_box(0), x_2, x_3, x_4, x_13, x_5, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_14) == 0) -{ -return x_14; -} -else -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) -{ -return x_14; -} -else -{ -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_14, 0); -lean_inc(x_16); -lean_dec(x_14); -x_17 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_17, 0, x_16); -return x_17; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_11 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62___lam__1___closed__1; -x_12 = lean_mk_empty_array_with_capacity(x_1); -x_13 = lean_array_push(x_12, x_5); +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_11 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__1___closed__1; +x_12 = lean_unsigned_to_nat(0u); +x_13 = l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___lam__4___closed__0; +x_14 = lean_array_push(x_13, x_5); lean_inc(x_9); lean_inc_ref(x_8); lean_inc(x_7); lean_inc_ref(x_6); lean_inc(x_4); lean_inc_ref(x_3); -lean_inc_ref(x_13); -x_14 = lean_apply_10(x_2, x_11, x_1, x_13, x_3, x_4, x_6, x_7, x_8, x_9, lean_box(0)); -if (lean_obj_tag(x_14) == 0) +lean_inc_ref(x_14); +x_15 = lean_apply_11(x_1, x_11, x_2, x_12, x_14, x_3, x_4, x_6, x_7, x_8, x_9, lean_box(0)); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_15; uint8_t x_16; uint8_t x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -lean_dec_ref(x_14); -x_16 = 0; -x_17 = 1; +lean_object* x_16; uint8_t x_17; uint8_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; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +lean_dec_ref(x_15); +x_17 = 0; x_18 = 1; -x_19 = lean_box(x_16); +x_19 = 1; x_20 = lean_box(x_17); x_21 = lean_box(x_18); -x_22 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_ProofMode_mIntroForall___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__4___lam__0___boxed), 13, 5); -lean_closure_set(x_22, 0, x_13); -lean_closure_set(x_22, 1, x_15); -lean_closure_set(x_22, 2, x_19); -lean_closure_set(x_22, 3, x_20); -lean_closure_set(x_22, 4, x_21); -x_23 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_22, x_3, x_4, x_6, x_7, x_8, x_9); +x_22 = lean_box(x_19); +x_23 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_ProofMode_mIntroForall___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__4___lam__0___boxed), 13, 5); +lean_closure_set(x_23, 0, x_14); +lean_closure_set(x_23, 1, x_16); +lean_closure_set(x_23, 2, x_20); +lean_closure_set(x_23, 3, x_21); +lean_closure_set(x_23, 4, x_22); +x_24 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_23, x_3, x_4, x_6, x_7, x_8, x_9); lean_dec(x_4); -return x_23; +return x_24; } else { -lean_dec_ref(x_13); +lean_dec_ref(x_14); lean_dec(x_9); lean_dec_ref(x_8); lean_dec(x_7); lean_dec_ref(x_6); lean_dec(x_4); lean_dec_ref(x_3); -return x_14; +return x_15; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62___lam__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62___lam__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { lean_object* x_14; lean_object* x_15; -x_14 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62___lam__1___boxed), 10, 4); +x_14 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62___lam__1___boxed), 10, 4); lean_closure_set(x_14, 0, x_1); lean_closure_set(x_14, 1, x_2); lean_closure_set(x_14, 2, x_7); @@ -30130,49 +30064,169 @@ return x_18; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__63___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_13 = l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___closed__13; -x_14 = lean_unsigned_to_nat(1u); -x_15 = lean_nat_add(x_2, x_14); -x_16 = lean_name_append_index_after(x_13, x_15); -x_17 = lean_apply_10(x_1, x_16, x_2, x_4, x_6, x_7, x_8, x_9, x_10, x_11, lean_box(0)); -return x_17; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__63___lam__1___closed__1; +x_13 = lean_mk_empty_array_with_capacity(x_1); +x_14 = lean_array_push(x_13, x_6); +lean_inc(x_10); +lean_inc_ref(x_9); +lean_inc(x_8); +lean_inc_ref(x_7); +lean_inc(x_5); +lean_inc_ref(x_4); +lean_inc_ref(x_14); +x_15 = lean_apply_11(x_2, x_12, x_3, x_1, x_14, x_4, x_5, x_7, x_8, x_9, x_10, lean_box(0)); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; uint8_t x_17; uint8_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; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +lean_dec_ref(x_15); +x_17 = 0; +x_18 = 1; +x_19 = 1; +x_20 = lean_box(x_17); +x_21 = lean_box(x_18); +x_22 = lean_box(x_19); +x_23 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_ProofMode_mIntroForall___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onGoal_spec__4___lam__0___boxed), 13, 5); +lean_closure_set(x_23, 0, x_14); +lean_closure_set(x_23, 1, x_16); +lean_closure_set(x_23, 2, x_20); +lean_closure_set(x_23, 3, x_21); +lean_closure_set(x_23, 4, x_22); +x_24 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_23, x_4, x_5, x_7, x_8, x_9, x_10); +lean_dec(x_5); +return x_24; } +else +{ +lean_dec_ref(x_14); +lean_dec(x_10); +lean_dec_ref(x_9); +lean_dec(x_8); +lean_dec_ref(x_7); +lean_dec(x_5); +lean_dec_ref(x_4); +return x_15; } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__63___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_13; -x_13 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec_ref(x_5); -lean_dec_ref(x_3); -return x_13; +lean_object* x_12; +x_12 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__63___lam__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__63(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__63___lam__1___boxed), 11, 5); +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_closure_set(x_15, 3, x_8); +lean_closure_set(x_15, 4, x_9); +x_16 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp(lean_box(0), x_4, x_5, x_6, x_15, x_7, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_16) == 0) +{ +return x_16; +} +else +{ +uint8_t x_17; +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +return x_16; +} +else +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +lean_dec(x_16); +x_19 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_19, 0, x_18); +return x_19; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; size_t x_9; size_t x_10; lean_object* x_11; lean_object* x_12; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = l_Lean_Expr_isFVar(x_5); +lean_dec(x_5); +x_9 = 1; +x_10 = lean_usize_add(x_2, x_9); +x_11 = lean_box(x_8); +x_12 = lean_array_uset(x_7, x_2, x_11); +x_2 = x_10; +x_3 = x_12; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; -x_11 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_9); lean_dec_ref(x_8); lean_dec(x_7); lean_dec_ref(x_6); lean_dec(x_5); lean_dec_ref(x_4); -lean_dec_ref(x_3); +lean_dec(x_3); lean_dec_ref(x_2); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_13 = l_Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28___closed__13; +x_14 = lean_unsigned_to_nat(1u); +x_15 = lean_nat_add(x_2, x_14); +x_16 = lean_name_append_index_after(x_13, x_15); +x_17 = lean_apply_11(x_1, x_16, x_3, x_2, x_4, x_6, x_7, x_8, x_9, x_10, x_11, lean_box(0)); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec_ref(x_5); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; -x_9 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_9 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_7); lean_dec_ref(x_6); lean_dec(x_5); @@ -30182,11 +30236,11 @@ lean_dec_ref(x_2); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; -x_9 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_9 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_7); lean_dec_ref(x_6); lean_dec(x_5); @@ -30196,7 +30250,36 @@ lean_dec_ref(x_2); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__57___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = l_Array_mask___redArg(x_1, x_3); +x_13 = lean_expr_instantiate_rev(x_2, x_12); +lean_dec_ref(x_12); +x_14 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_14, 0, x_13); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_10); +lean_dec_ref(x_9); +lean_dec(x_8); +lean_dec_ref(x_7); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec_ref(x_4); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__58___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { lean_object* x_9; lean_object* x_10; lean_object* x_11; @@ -30209,11 +30292,11 @@ lean_ctor_set(x_11, 0, x_10); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__57___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__58___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; -x_9 = l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__57___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_9 = l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__58___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_7); lean_dec_ref(x_6); lean_dec(x_5); @@ -30224,18 +30307,18 @@ lean_dec(x_1); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); return x_11; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { lean_object* x_9; @@ -30244,11 +30327,11 @@ lean_ctor_set(x_9, 0, x_1); return x_9; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_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_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_9 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_7); lean_dec_ref(x_6); lean_dec(x_5); @@ -30258,7 +30341,7 @@ lean_dec_ref(x_2); return x_9; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__45___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__46___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_10; @@ -30266,11 +30349,11 @@ x_10 = l_Lean_FVarId_getUserName___redArg(x_1, x_5, x_7, x_8); return x_10; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__45___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__46___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__45___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_10 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__46___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_8); lean_dec_ref(x_7); lean_dec(x_6); @@ -30280,7 +30363,7 @@ lean_dec(x_2); return x_10; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__45(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___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__46(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_11; @@ -30303,7 +30386,7 @@ lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; x_13 = lean_array_uget(x_3, x_2); x_14 = l_Lean_Expr_fvarId_x21(x_13); lean_dec(x_13); -x_15 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__45___lam__0___boxed), 9, 1); +x_15 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__46___lam__0___boxed), 9, 1); lean_closure_set(x_15, 0, x_14); lean_inc(x_9); lean_inc_ref(x_8); @@ -30354,27 +30437,27 @@ return x_26; } } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__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_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__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: { size_t x_10; size_t x_11; lean_object* x_12; x_10 = lean_array_size(x_1); x_11 = 0; -x_12 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__45(x_10, x_11, x_1, x_3, x_4, x_5, x_6, x_7, x_8); +x_12 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__46(x_10, x_11, x_1, x_3, x_4, x_5, x_6, x_7, x_8); return x_12; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_10 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_4); lean_dec_ref(x_2); return x_10; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_11; @@ -30382,11 +30465,11 @@ x_11 = l_Lean_Meta_instantiateLambda(x_1, x_2, x_6, x_7, x_8, x_9); return x_11; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); @@ -30394,7 +30477,7 @@ lean_dec_ref(x_2); return x_11; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, uint8_t x_7, uint8_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, uint8_t x_7, uint8_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { lean_object* x_16; @@ -30467,18 +30550,18 @@ return x_16; } } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { uint8_t x_16; uint8_t x_17; lean_object* x_18; x_16 = lean_unbox(x_7); x_17 = lean_unbox(x_8); -x_18 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__4(x_1, x_2, x_3, x_4, x_5, x_6, x_16, x_17, x_9, x_10, x_11, x_12, x_13, x_14); +x_18 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__4(x_1, x_2, x_3, x_4, x_5, x_6, x_16, x_17, x_9, x_10, x_11, x_12, x_13, x_14); lean_dec_ref(x_6); return x_18; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__2(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, uint8_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__2(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, uint8_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { lean_object* x_18; @@ -30488,7 +30571,7 @@ lean_inc(x_14); lean_inc_ref(x_13); lean_inc(x_12); lean_inc_ref(x_11); -x_18 = l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__48___redArg(x_1, x_2, x_3, x_11, x_12, x_13, x_14, x_15, x_16); +x_18 = l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__49___redArg(x_1, x_2, x_3, x_11, x_12, x_13, x_14, x_15, x_16); if (lean_obj_tag(x_18) == 0) { lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; @@ -30498,7 +30581,7 @@ lean_dec_ref(x_18); x_20 = lean_box(x_3); x_21 = lean_box(x_8); lean_inc_ref(x_7); -x_22 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__4___boxed), 15, 8); +x_22 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__4___boxed), 15, 8); lean_closure_set(x_22, 0, x_4); lean_closure_set(x_22, 1, x_5); lean_closure_set(x_22, 2, x_6); @@ -30507,7 +30590,7 @@ lean_closure_set(x_22, 4, x_7); lean_closure_set(x_22, 5, x_9); lean_closure_set(x_22, 6, x_20); lean_closure_set(x_22, 7, x_21); -x_23 = l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__50___redArg(x_7, x_19, x_22, x_11, x_12, x_13, x_14, x_15, x_16); +x_23 = l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__51___redArg(x_7, x_19, x_22, x_11, x_12, x_13, x_14, x_15, x_16); lean_dec(x_19); lean_dec_ref(x_7); return x_23; @@ -30545,7 +30628,7 @@ return x_26; } } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__2___boxed(lean_object** _args) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__2___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -30568,22 +30651,22 @@ lean_object* x_17 = _args[16]; uint8_t x_18; uint8_t x_19; lean_object* x_20; x_18 = lean_unbox(x_3); x_19 = lean_unbox(x_8); -x_20 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__2(x_1, x_2, x_18, x_4, x_5, x_6, x_7, x_19, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +x_20 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__2(x_1, x_2, x_18, x_4, x_5, x_6, x_7, x_19, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); return x_20; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__5(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__5(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_inc_ref(x_8); lean_inc_ref(x_1); -x_17 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__1___boxed), 10, 2); +x_17 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__1___boxed), 10, 2); lean_closure_set(x_17, 0, x_1); lean_closure_set(x_17, 1, x_8); x_18 = lean_box(x_3); x_19 = lean_box(x_6); -x_20 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__2___boxed), 17, 8); +x_20 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__2___boxed), 17, 8); lean_closure_set(x_20, 0, x_1); lean_closure_set(x_20, 1, x_2); lean_closure_set(x_20, 2, x_18); @@ -30598,17 +30681,17 @@ x_22 = l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Elab_Tactic_Do return x_22; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { uint8_t x_17; uint8_t x_18; lean_object* x_19; x_17 = lean_unbox(x_3); x_18 = lean_unbox(x_6); -x_19 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__5(x_1, x_2, x_17, x_4, x_5, x_18, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_19 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__5(x_1, x_2, x_17, x_4, x_5, x_18, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); return x_19; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { lean_object* x_16; @@ -30684,16 +30767,16 @@ return x_33; } } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { uint8_t x_16; lean_object* x_17; x_16 = lean_unbox(x_4); -x_17 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__6(x_1, x_2, x_3, x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +x_17 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__6(x_1, x_2, x_3, x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); return x_17; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg(lean_object* x_1, uint8_t x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg(lean_object* x_1, uint8_t x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { lean_object* x_15; uint8_t x_35; @@ -30753,7 +30836,7 @@ if (x_52 == 0) lean_object* x_53; lean_object* x_54; x_53 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_53, 0, x_7); -x_54 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__0___boxed), 8, 1); +x_54 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_54, 0, x_53); x_15 = x_54; goto block_34; @@ -30789,7 +30872,7 @@ lean_dec(x_50); lean_dec_ref(x_49); x_65 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_65, 0, x_7); -x_66 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__0___boxed), 8, 1); +x_66 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_66, 0, x_65); x_15 = x_66; goto block_34; @@ -30826,7 +30909,7 @@ lean_dec(x_50); lean_dec_ref(x_49); x_76 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_76, 0, x_7); -x_77 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__0___boxed), 8, 1); +x_77 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_77, 0, x_76); x_15 = x_77; goto block_34; @@ -30850,7 +30933,7 @@ x_80 = lean_ctor_get(x_41, 1); lean_dec(x_80); x_81 = lean_ctor_get(x_41, 0); lean_dec(x_81); -x_82 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__3___boxed), 9, 0); +x_82 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__3___boxed), 9, 0); x_83 = lean_array_fget(x_49, x_50); lean_dec(x_50); lean_dec_ref(x_49); @@ -30864,7 +30947,7 @@ lean_inc(x_5); lean_inc(x_6); lean_inc_ref(x_3); lean_inc(x_85); -x_88 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__5___boxed), 16, 7); +x_88 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__5___boxed), 16, 7); lean_closure_set(x_88, 0, x_85); lean_closure_set(x_88, 1, x_82); lean_closure_set(x_88, 2, x_86); @@ -30878,7 +30961,7 @@ lean_ctor_set(x_41, 1, x_89); x_90 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_90, 0, x_84); x_91 = lean_box(x_2); -x_92 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__6___boxed), 15, 8); +x_92 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__6___boxed), 15, 8); lean_closure_set(x_92, 0, x_83); lean_closure_set(x_92, 1, x_90); lean_closure_set(x_92, 2, x_88); @@ -30894,7 +30977,7 @@ else { lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_dec(x_41); -x_93 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__3___boxed), 9, 0); +x_93 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__3___boxed), 9, 0); x_94 = lean_array_fget(x_49, x_50); lean_dec(x_50); lean_dec_ref(x_49); @@ -30908,7 +30991,7 @@ lean_inc(x_5); lean_inc(x_6); lean_inc_ref(x_3); lean_inc(x_96); -x_99 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__5___boxed), 16, 7); +x_99 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__5___boxed), 16, 7); lean_closure_set(x_99, 0, x_96); lean_closure_set(x_99, 1, x_93); lean_closure_set(x_99, 2, x_97); @@ -30925,7 +31008,7 @@ lean_ctor_set(x_101, 2, x_73); x_102 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_102, 0, x_95); x_103 = lean_box(x_2); -x_104 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__6___boxed), 15, 8); +x_104 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__6___boxed), 15, 8); lean_closure_set(x_104, 0, x_94); lean_closure_set(x_104, 1, x_102); lean_closure_set(x_104, 2, x_99); @@ -30963,7 +31046,7 @@ lean_dec_ref(x_49); lean_ctor_set(x_37, 0, x_109); x_111 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_111, 0, x_7); -x_112 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__0___boxed), 8, 1); +x_112 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_112, 0, x_111); x_15 = x_112; goto block_34; @@ -30986,7 +31069,7 @@ if (lean_is_exclusive(x_41)) { lean_dec_ref(x_41); x_113 = lean_box(0); } -x_114 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__3___boxed), 9, 0); +x_114 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__3___boxed), 9, 0); x_115 = lean_array_fget(x_49, x_50); lean_dec(x_50); lean_dec_ref(x_49); @@ -31000,7 +31083,7 @@ lean_inc(x_5); lean_inc(x_6); lean_inc_ref(x_3); lean_inc(x_117); -x_120 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__5___boxed), 16, 7); +x_120 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__5___boxed), 16, 7); lean_closure_set(x_120, 0, x_117); lean_closure_set(x_120, 1, x_114); lean_closure_set(x_120, 2, x_118); @@ -31021,7 +31104,7 @@ lean_ctor_set(x_122, 2, x_107); x_123 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_123, 0, x_116); x_124 = lean_box(x_2); -x_125 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__6___boxed), 15, 8); +x_125 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__6___boxed), 15, 8); lean_closure_set(x_125, 0, x_115); lean_closure_set(x_125, 1, x_123); lean_closure_set(x_125, 2, x_120); @@ -31059,7 +31142,7 @@ lean_dec_ref(x_49); lean_ctor_set(x_38, 0, x_131); x_133 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_133, 0, x_7); -x_134 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__0___boxed), 8, 1); +x_134 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_134, 0, x_133); x_15 = x_134; goto block_34; @@ -31104,7 +31187,7 @@ lean_ctor_set(x_38, 0, x_131); lean_ctor_set(x_37, 0, x_140); x_142 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_142, 0, x_7); -x_143 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__0___boxed), 8, 1); +x_143 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_143, 0, x_142); x_15 = x_143; goto block_34; @@ -31127,7 +31210,7 @@ if (lean_is_exclusive(x_41)) { lean_dec_ref(x_41); x_144 = lean_box(0); } -x_145 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__3___boxed), 9, 0); +x_145 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__3___boxed), 9, 0); x_146 = lean_array_fget(x_49, x_50); lean_dec(x_50); lean_dec_ref(x_49); @@ -31141,7 +31224,7 @@ lean_inc(x_5); lean_inc(x_6); lean_inc_ref(x_3); lean_inc(x_148); -x_151 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__5___boxed), 16, 7); +x_151 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__5___boxed), 16, 7); lean_closure_set(x_151, 0, x_148); lean_closure_set(x_151, 1, x_145); lean_closure_set(x_151, 2, x_149); @@ -31162,7 +31245,7 @@ lean_ctor_set(x_153, 2, x_138); x_154 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_154, 0, x_147); x_155 = lean_box(x_2); -x_156 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__6___boxed), 15, 8); +x_156 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__6___boxed), 15, 8); lean_closure_set(x_156, 0, x_146); lean_closure_set(x_156, 1, x_154); lean_closure_set(x_156, 2, x_151); @@ -31197,7 +31280,7 @@ lean_ctor_set(x_162, 1, x_157); lean_ctor_set(x_37, 1, x_162); x_163 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_163, 0, x_7); -x_164 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__0___boxed), 8, 1); +x_164 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_164, 0, x_163); x_15 = x_164; goto block_34; @@ -31243,7 +31326,7 @@ lean_ctor_set(x_173, 1, x_157); lean_ctor_set(x_37, 1, x_173); x_174 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_174, 0, x_7); -x_175 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__0___boxed), 8, 1); +x_175 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_175, 0, x_174); x_15 = x_175; goto block_34; @@ -31291,7 +31374,7 @@ lean_ctor_set(x_37, 1, x_183); lean_ctor_set(x_37, 0, x_181); x_184 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_184, 0, x_7); -x_185 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__0___boxed), 8, 1); +x_185 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_185, 0, x_184); x_15 = x_185; goto block_34; @@ -31313,7 +31396,7 @@ if (lean_is_exclusive(x_41)) { lean_dec_ref(x_41); x_186 = lean_box(0); } -x_187 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__3___boxed), 9, 0); +x_187 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__3___boxed), 9, 0); x_188 = lean_array_fget(x_158, x_159); lean_dec(x_159); lean_dec_ref(x_158); @@ -31327,7 +31410,7 @@ lean_inc(x_5); lean_inc(x_6); lean_inc_ref(x_3); lean_inc(x_190); -x_193 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__5___boxed), 16, 7); +x_193 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__5___boxed), 16, 7); lean_closure_set(x_193, 0, x_190); lean_closure_set(x_193, 1, x_187); lean_closure_set(x_193, 2, x_191); @@ -31348,7 +31431,7 @@ lean_ctor_set(x_195, 2, x_179); x_196 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_196, 0, x_189); x_197 = lean_box(x_2); -x_198 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__6___boxed), 15, 8); +x_198 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__6___boxed), 15, 8); lean_closure_set(x_198, 0, x_188); lean_closure_set(x_198, 1, x_196); lean_closure_set(x_198, 2, x_193); @@ -31400,7 +31483,7 @@ lean_ctor_set(x_207, 1, x_206); lean_ctor_set(x_7, 1, x_207); x_208 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_208, 0, x_7); -x_209 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__0___boxed), 8, 1); +x_209 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_209, 0, x_208); x_15 = x_209; goto block_34; @@ -31453,7 +31536,7 @@ lean_ctor_set(x_219, 1, x_218); lean_ctor_set(x_7, 1, x_219); x_220 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_220, 0, x_7); -x_221 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__0___boxed), 8, 1); +x_221 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_221, 0, x_220); x_15 = x_221; goto block_34; @@ -31507,7 +31590,7 @@ lean_ctor_set(x_230, 1, x_229); lean_ctor_set(x_7, 1, x_230); x_231 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_231, 0, x_7); -x_232 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__0___boxed), 8, 1); +x_232 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_232, 0, x_231); x_15 = x_232; goto block_34; @@ -31529,7 +31612,7 @@ if (lean_is_exclusive(x_41)) { lean_dec_ref(x_41); x_233 = lean_box(0); } -x_234 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__3___boxed), 9, 0); +x_234 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__3___boxed), 9, 0); x_235 = lean_array_fget(x_202, x_203); lean_dec(x_203); lean_dec_ref(x_202); @@ -31543,7 +31626,7 @@ lean_inc(x_5); lean_inc(x_6); lean_inc_ref(x_3); lean_inc(x_237); -x_240 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__5___boxed), 16, 7); +x_240 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__5___boxed), 16, 7); lean_closure_set(x_240, 0, x_237); lean_closure_set(x_240, 1, x_234); lean_closure_set(x_240, 2, x_238); @@ -31564,7 +31647,7 @@ lean_ctor_set(x_242, 2, x_225); x_243 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_243, 0, x_236); x_244 = lean_box(x_2); -x_245 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__6___boxed), 15, 8); +x_245 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__6___boxed), 15, 8); lean_closure_set(x_245, 0, x_235); lean_closure_set(x_245, 1, x_243); lean_closure_set(x_245, 2, x_240); @@ -31632,7 +31715,7 @@ lean_ctor_set(x_257, 0, x_246); lean_ctor_set(x_257, 1, x_256); x_258 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_258, 0, x_257); -x_259 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__0___boxed), 8, 1); +x_259 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_259, 0, x_258); x_15 = x_259; goto block_34; @@ -31691,7 +31774,7 @@ lean_ctor_set(x_270, 0, x_246); lean_ctor_set(x_270, 1, x_269); x_271 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_271, 0, x_270); -x_272 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__0___boxed), 8, 1); +x_272 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_272, 0, x_271); x_15 = x_272; goto block_34; @@ -31751,7 +31834,7 @@ lean_ctor_set(x_282, 0, x_246); lean_ctor_set(x_282, 1, x_281); x_283 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_283, 0, x_282); -x_284 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__0___boxed), 8, 1); +x_284 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_284, 0, x_283); x_15 = x_284; goto block_34; @@ -31773,7 +31856,7 @@ if (lean_is_exclusive(x_246)) { lean_dec_ref(x_246); x_285 = lean_box(0); } -x_286 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__3___boxed), 9, 0); +x_286 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__3___boxed), 9, 0); x_287 = lean_array_fget(x_251, x_252); lean_dec(x_252); lean_dec_ref(x_251); @@ -31787,7 +31870,7 @@ lean_inc(x_5); lean_inc(x_6); lean_inc_ref(x_3); lean_inc(x_289); -x_292 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__5___boxed), 16, 7); +x_292 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__5___boxed), 16, 7); lean_closure_set(x_292, 0, x_289); lean_closure_set(x_292, 1, x_286); lean_closure_set(x_292, 2, x_290); @@ -31808,7 +31891,7 @@ lean_ctor_set(x_294, 2, x_276); x_295 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_295, 0, x_288); x_296 = lean_box(x_2); -x_297 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___lam__6___boxed), 15, 8); +x_297 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___lam__6___boxed), 15, 8); lean_closure_set(x_297, 0, x_287); lean_closure_set(x_297, 1, x_295); lean_closure_set(x_297, 2, x_292); @@ -31946,18 +32029,18 @@ return x_33; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_10 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_4); lean_dec_ref(x_3); lean_dec(x_2); return x_10; } } -static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__0() { +static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__0() { _start: { lean_object* x_1; @@ -31965,16 +32048,16 @@ x_1 = lean_mk_string_unchecked("failed to transform matcher, type error when con return x_1; } } -static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__1() { +static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__1() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__0; +x_1 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__0; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { lean_object* x_9; lean_object* x_10; @@ -31985,11 +32068,11 @@ lean_ctor_set(x_10, 0, x_9); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_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_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_9 = l_Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_7); lean_dec_ref(x_6); lean_dec(x_5); @@ -32000,7 +32083,7 @@ lean_dec(x_1); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { lean_object* x_9; lean_object* x_10; @@ -32011,11 +32094,11 @@ lean_ctor_set(x_10, 0, x_9); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_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_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_9 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_7); lean_dec_ref(x_6); lean_dec(x_5); @@ -32026,7 +32109,7 @@ lean_dec(x_1); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { lean_object* x_9; lean_object* x_10; @@ -32037,11 +32120,11 @@ lean_ctor_set(x_10, 0, x_9); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_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_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_9 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_7); lean_dec_ref(x_6); lean_dec(x_5); @@ -32052,11 +32135,11 @@ lean_dec(x_1); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_10 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_8); lean_dec_ref(x_7); lean_dec(x_6); @@ -32067,7 +32150,7 @@ lean_dec(x_2); return x_10; } } -static lean_object* _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__0() { +static lean_object* _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__0() { _start: { lean_object* x_1; @@ -32075,7 +32158,7 @@ x_1 = lean_mk_string_unchecked("unsolvedGoals", 13, 13); return x_1; } } -static lean_object* _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__1() { +static lean_object* _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__1() { _start: { lean_object* x_1; @@ -32083,7 +32166,7 @@ x_1 = lean_mk_string_unchecked("synthPlaceholder", 16, 16); return x_1; } } -static lean_object* _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__2() { +static lean_object* _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__2() { _start: { lean_object* x_1; @@ -32091,7 +32174,7 @@ x_1 = lean_mk_string_unchecked("lean", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__3() { +static lean_object* _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__3() { _start: { lean_object* x_1; @@ -32099,7 +32182,7 @@ x_1 = lean_mk_string_unchecked("inductionWithNoAlts", 19, 19); return x_1; } } -static lean_object* _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__4() { +static lean_object* _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__4() { _start: { lean_object* x_1; @@ -32107,7 +32190,7 @@ x_1 = lean_mk_string_unchecked("_namedError", 11, 11); return x_1; } } -static lean_object* _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__5() { +static lean_object* _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__5() { _start: { lean_object* x_1; @@ -32115,7 +32198,7 @@ x_1 = lean_mk_string_unchecked("trace", 5, 5); return x_1; } } -LEAN_EXPORT uint8_t l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2(uint8_t x_1, uint8_t x_2, lean_object* x_3) { +LEAN_EXPORT uint8_t l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2(uint8_t x_1, uint8_t x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 1) @@ -32147,7 +32230,7 @@ return x_1; else { lean_object* x_12; uint8_t x_13; -x_12 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__0; +x_12 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__0; x_13 = lean_string_dec_eq(x_6, x_12); if (x_13 == 0) { @@ -32162,7 +32245,7 @@ return x_2; else { lean_object* x_14; uint8_t x_15; -x_14 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__1; +x_14 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__1; x_15 = lean_string_dec_eq(x_6, x_14); if (x_15 == 0) { @@ -32184,7 +32267,7 @@ lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint x_17 = lean_ctor_get(x_3, 1); x_18 = lean_ctor_get(x_4, 1); x_19 = lean_ctor_get(x_5, 1); -x_20 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__2; +x_20 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__2; x_21 = lean_string_dec_eq(x_19, x_20); if (x_21 == 0) { @@ -32193,7 +32276,7 @@ return x_1; else { lean_object* x_22; uint8_t x_23; -x_22 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__3; +x_22 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__3; x_23 = lean_string_dec_eq(x_18, x_22); if (x_23 == 0) { @@ -32202,7 +32285,7 @@ return x_1; else { lean_object* x_24; uint8_t x_25; -x_24 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__4; +x_24 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__4; x_25 = lean_string_dec_eq(x_17, x_24); if (x_25 == 0) { @@ -32230,7 +32313,7 @@ case 0: { lean_object* x_26; lean_object* x_27; uint8_t x_28; x_26 = lean_ctor_get(x_3, 1); -x_27 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__5; +x_27 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__5; x_28 = lean_string_dec_eq(x_26, x_27); if (x_28 == 0) { @@ -32253,19 +32336,19 @@ return x_1; } } } -LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { uint8_t x_4; uint8_t x_5; uint8_t x_6; lean_object* x_7; x_4 = lean_unbox(x_1); x_5 = lean_unbox(x_2); -x_6 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2(x_4, x_5, x_3); +x_6 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2(x_4, x_5, x_3); lean_dec(x_3); x_7 = lean_box(x_6); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, uint8_t x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, uint8_t x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_52; @@ -32282,7 +32365,7 @@ else lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; x_53 = lean_box(x_5); x_54 = lean_box(x_52); -x_55 = lean_alloc_closure((void*)(l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___boxed), 3, 2); +x_55 = lean_alloc_closure((void*)(l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___boxed), 3, 2); lean_closure_set(x_55, 0, x_53); lean_closure_set(x_55, 1, x_54); lean_inc_ref(x_1); @@ -32405,14 +32488,14 @@ return x_50; } } } -LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { uint8_t x_17; uint8_t x_18; uint8_t x_19; lean_object* x_20; x_17 = lean_unbox(x_5); x_18 = lean_unbox(x_6); x_19 = lean_unbox(x_7); -x_20 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__3(x_1, x_2, x_3, x_4, x_17, x_18, x_19, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_20 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__3(x_1, x_2, x_3, x_4, x_17, x_18, x_19, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); lean_dec(x_15); lean_dec_ref(x_14); lean_dec(x_13); @@ -32423,7 +32506,7 @@ lean_dec(x_9); return x_20; } } -static lean_object* _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___closed__1() { +static lean_object* _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___closed__1() { _start: { lean_object* x_1; @@ -32431,14 +32514,14 @@ x_1 = l_Lean_warningAsError; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_12; lean_object* x_13; uint8_t x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_46; uint8_t x_47; lean_object* x_48; uint8_t x_49; lean_object* x_50; lean_object* x_51; lean_object* x_55; uint8_t x_56; uint8_t x_57; lean_object* x_58; uint8_t x_59; lean_object* x_70; uint8_t x_71; uint8_t x_72; lean_object* x_73; uint8_t x_74; uint8_t x_75; uint8_t x_77; uint8_t x_90; -x_12 = lean_alloc_closure((void*)(l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__0___boxed), 8, 0); -x_13 = lean_alloc_closure((void*)(l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__1___boxed), 8, 0); -x_55 = lean_alloc_closure((void*)(l_Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59___lam__0___boxed), 8, 0); -x_70 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___closed__0; +lean_object* x_12; lean_object* x_13; uint8_t x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_46; uint8_t x_47; uint8_t x_48; uint8_t x_49; lean_object* x_50; lean_object* x_51; lean_object* x_55; uint8_t x_56; lean_object* x_57; uint8_t x_58; uint8_t x_59; lean_object* x_70; uint8_t x_71; lean_object* x_72; uint8_t x_73; uint8_t x_74; uint8_t x_75; uint8_t x_77; uint8_t x_90; +x_12 = lean_alloc_closure((void*)(l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__0___boxed), 8, 0); +x_13 = lean_alloc_closure((void*)(l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__1___boxed), 8, 0); +x_55 = lean_alloc_closure((void*)(l_Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60___lam__0___boxed), 8, 0); +x_70 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___closed__0; x_71 = 2; x_90 = l_Lean_instBEqMessageSeverity_beq(x_3, x_71); if (x_90 == 0) @@ -32470,7 +32553,7 @@ x_20 = lean_ctor_get(x_19, 0); lean_inc(x_20); lean_dec_ref(x_19); x_21 = l___private_Lean_Log_0__Lean_MessageData_appendDescriptionWidgetIfNamed(x_2); -x_22 = lean_alloc_closure((void*)(l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__4___boxed), 9, 1); +x_22 = lean_alloc_closure((void*)(l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__4___boxed), 9, 1); lean_closure_set(x_22, 0, x_21); lean_inc(x_10); lean_inc_ref(x_9); @@ -32497,8 +32580,8 @@ x_26 = lean_ctor_get(x_25, 0); lean_inc(x_26); lean_dec_ref(x_25); lean_inc(x_20); -x_27 = l_Lean_FileMap_toPosition(x_20, x_16); -lean_dec(x_16); +x_27 = l_Lean_FileMap_toPosition(x_20, x_17); +lean_dec(x_17); x_28 = l_Lean_FileMap_toPosition(x_20, x_18); lean_dec(x_18); x_29 = lean_alloc_ctor(1, 1, 0); @@ -32507,7 +32590,7 @@ x_30 = l_Lean_addTrace___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab x_31 = lean_box(x_15); x_32 = lean_box(x_14); x_33 = lean_box(x_4); -x_34 = lean_alloc_closure((void*)(l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__3___boxed), 16, 8); +x_34 = lean_alloc_closure((void*)(l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__3___boxed), 16, 8); lean_closure_set(x_34, 0, x_24); lean_closure_set(x_34, 1, x_26); lean_closure_set(x_34, 2, x_27); @@ -32525,7 +32608,7 @@ uint8_t x_36; lean_dec(x_24); lean_dec(x_20); lean_dec(x_18); -lean_dec(x_16); +lean_dec(x_17); lean_dec(x_10); lean_dec_ref(x_9); lean_dec(x_8); @@ -32553,7 +32636,7 @@ else uint8_t x_39; lean_dec(x_20); lean_dec(x_18); -lean_dec(x_16); +lean_dec(x_17); lean_dec_ref(x_12); lean_dec(x_10); lean_dec_ref(x_9); @@ -32581,7 +32664,7 @@ else { uint8_t x_42; lean_dec(x_18); -lean_dec(x_16); +lean_dec(x_17); lean_dec_ref(x_12); lean_dec(x_10); lean_dec_ref(x_9); @@ -32614,10 +32697,10 @@ lean_dec(x_50); if (lean_obj_tag(x_52) == 0) { lean_inc(x_51); -x_14 = x_46; -x_15 = x_47; -x_16 = x_51; -x_17 = lean_box(0); +x_14 = x_47; +x_15 = x_48; +x_16 = lean_box(0); +x_17 = x_51; x_18 = x_51; goto block_45; } @@ -32627,10 +32710,10 @@ lean_object* x_53; x_53 = lean_ctor_get(x_52, 0); lean_inc(x_53); lean_dec_ref(x_52); -x_14 = x_46; -x_15 = x_47; -x_16 = x_51; -x_17 = lean_box(0); +x_14 = x_47; +x_15 = x_48; +x_16 = lean_box(0); +x_17 = x_51; x_18 = x_53; goto block_45; } @@ -32652,15 +32735,15 @@ lean_inc(x_61); lean_dec_ref(x_60); x_62 = l_Lean_replaceRef(x_1, x_61); lean_dec(x_61); -x_63 = l_Lean_Syntax_getPos_x3f(x_62, x_57); +x_63 = l_Lean_Syntax_getPos_x3f(x_62, x_58); if (lean_obj_tag(x_63) == 0) { lean_object* x_64; x_64 = lean_unsigned_to_nat(0u); -x_46 = x_59; -x_47 = x_56; -x_48 = lean_box(0); -x_49 = x_57; +x_46 = lean_box(0); +x_47 = x_59; +x_48 = x_56; +x_49 = x_58; x_50 = x_62; x_51 = x_64; goto block_54; @@ -32671,10 +32754,10 @@ lean_object* x_65; x_65 = lean_ctor_get(x_63, 0); lean_inc(x_65); lean_dec_ref(x_63); -x_46 = x_59; -x_47 = x_56; -x_48 = lean_box(0); -x_49 = x_57; +x_46 = lean_box(0); +x_47 = x_59; +x_48 = x_56; +x_49 = x_58; x_50 = x_62; x_51 = x_65; goto block_54; @@ -32712,17 +32795,17 @@ return x_68; { if (x_75 == 0) { -x_56 = x_72; -x_57 = x_74; -x_58 = lean_box(0); +x_56 = x_73; +x_57 = lean_box(0); +x_58 = x_74; x_59 = x_3; goto block_69; } else { -x_56 = x_72; -x_57 = x_74; -x_58 = lean_box(0); +x_56 = x_73; +x_57 = lean_box(0); +x_58 = x_74; x_59 = x_71; goto block_69; } @@ -32749,8 +32832,8 @@ x_81 = l_Lean_instBEqMessageSeverity_beq(x_3, x_80); if (x_81 == 0) { lean_dec(x_79); -x_72 = x_77; -x_73 = lean_box(0); +x_72 = lean_box(0); +x_73 = x_77; x_74 = x_77; x_75 = x_81; goto block_76; @@ -32758,11 +32841,11 @@ goto block_76; else { lean_object* x_82; uint8_t x_83; -x_82 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___closed__1; +x_82 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___closed__1; x_83 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13_spec__27(x_79, x_82); lean_dec(x_79); -x_72 = x_77; -x_73 = lean_box(0); +x_72 = lean_box(0); +x_73 = x_77; x_74 = x_77; x_75 = x_83; goto block_76; @@ -32817,11 +32900,11 @@ return x_88; } } } -LEAN_EXPORT lean_object* l_Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59(lean_object* x_1, uint8_t x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60(lean_object* x_1, uint8_t x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_11; lean_object* x_12; -x_11 = lean_alloc_closure((void*)(l_Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59___lam__0___boxed), 8, 0); +x_11 = lean_alloc_closure((void*)(l_Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60___lam__0___boxed), 8, 0); lean_inc(x_9); lean_inc_ref(x_8); lean_inc(x_7); @@ -32834,7 +32917,7 @@ lean_object* x_13; lean_object* x_14; x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); lean_dec_ref(x_12); -x_14 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70(x_13, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_14 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71(x_13, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_13); return x_14; } @@ -32865,28 +32948,28 @@ return x_17; } } } -LEAN_EXPORT lean_object* l_Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52(lean_object* x_1, lean_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_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53(lean_object* x_1, lean_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_9; uint8_t x_10; lean_object* x_11; x_9 = 2; x_10 = 0; -x_11 = l_Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59(x_1, x_9, x_10, x_2, x_3, x_4, x_5, x_6, x_7); +x_11 = l_Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60(x_1, x_9, x_10, x_2, x_3, x_4, x_5, x_6, x_7); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_10 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_4); lean_dec_ref(x_3); lean_dec(x_2); return x_10; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { lean_object* x_9; @@ -32895,11 +32978,11 @@ lean_ctor_set(x_9, 0, x_1); return x_9; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_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_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_9 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_7); lean_dec_ref(x_6); lean_dec(x_5); @@ -32909,7 +32992,7 @@ lean_dec_ref(x_2); return x_9; } } -static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__0() { +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__0() { _start: { lean_object* x_1; @@ -32917,7 +33000,7 @@ x_1 = lean_mk_string_unchecked("assertion violation: altInfo.numOverlaps = 0\n return x_1; } } -static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__1() { +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__1() { _start: { lean_object* x_1; @@ -32925,7 +33008,7 @@ x_1 = lean_mk_string_unchecked("Lean.Meta.MatcherApp.transform", 30, 30); return x_1; } } -static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__0() { +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__0() { _start: { lean_object* x_1; @@ -32933,20 +33016,20 @@ x_1 = lean_mk_string_unchecked("Lean.Meta.Match.MatcherApp.Transform", 36, 36); return x_1; } } -static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__1() { +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__0; +x_1 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__0; x_2 = lean_unsigned_to_nat(6u); x_3 = lean_unsigned_to_nat(319u); -x_4 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__1; -x_5 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__0; +x_4 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__1; +x_5 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); return x_6; } } -static lean_object* _init_l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___closed__0() { +static lean_object* _init_l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___closed__0() { _start: { lean_object* x_1; @@ -32954,7 +33037,7 @@ x_1 = l_Subarray_empty(lean_box(0)); return x_1; } } -static lean_object* _init_l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___closed__1() { +static lean_object* _init_l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___closed__1() { _start: { lean_object* x_1; @@ -32962,7 +33045,7 @@ x_1 = l_Array_empty(lean_box(0)); return x_1; } } -LEAN_EXPORT lean_object* l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53(lean_object* x_1, lean_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_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { lean_object* x_9; uint8_t x_10; @@ -33048,8 +33131,8 @@ lean_ctor_set(x_29, 0, x_41); lean_ctor_set(x_27, 1, x_38); x_45 = l_ReaderT_instMonad___redArg(x_27); x_46 = l_ReaderT_instMonad___redArg(x_45); -x_47 = l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___closed__0; -x_48 = l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___closed__1; +x_47 = l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___closed__0; +x_48 = l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___closed__1; x_49 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_49, 0, x_47); lean_ctor_set(x_49, 1, x_48); @@ -33110,8 +33193,8 @@ lean_ctor_set(x_27, 1, x_63); lean_ctor_set(x_27, 0, x_70); x_71 = l_ReaderT_instMonad___redArg(x_27); x_72 = l_ReaderT_instMonad___redArg(x_71); -x_73 = l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___closed__0; -x_74 = l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___closed__1; +x_73 = l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___closed__0; +x_74 = l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___closed__1; x_75 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_75, 0, x_73); lean_ctor_set(x_75, 1, x_74); @@ -33191,8 +33274,8 @@ lean_ctor_set(x_99, 0, x_98); lean_ctor_set(x_99, 1, x_91); x_100 = l_ReaderT_instMonad___redArg(x_99); x_101 = l_ReaderT_instMonad___redArg(x_100); -x_102 = l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___closed__0; -x_103 = l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___closed__1; +x_102 = l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___closed__0; +x_103 = l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___closed__1; x_104 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_104, 0, x_102); lean_ctor_set(x_104, 1, x_103); @@ -33317,8 +33400,8 @@ lean_ctor_set(x_143, 0, x_142); lean_ctor_set(x_143, 1, x_135); x_144 = l_ReaderT_instMonad___redArg(x_143); x_145 = l_ReaderT_instMonad___redArg(x_144); -x_146 = l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___closed__0; -x_147 = l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___closed__1; +x_146 = l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___closed__0; +x_147 = l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___closed__1; x_148 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_148, 0, x_146); lean_ctor_set(x_148, 1, x_147); @@ -33461,8 +33544,8 @@ lean_ctor_set(x_190, 0, x_189); lean_ctor_set(x_190, 1, x_182); x_191 = l_ReaderT_instMonad___redArg(x_190); x_192 = l_ReaderT_instMonad___redArg(x_191); -x_193 = l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___closed__0; -x_194 = l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___closed__1; +x_193 = l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___closed__0; +x_194 = l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___closed__1; x_195 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_195, 0, x_193); lean_ctor_set(x_195, 1, x_194); @@ -33487,15 +33570,15 @@ return x_203; } } } -LEAN_EXPORT lean_object* l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_9 = l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54(x_1, x_2, x_3, x_4, x_5, x_6, x_7); return x_9; } } -static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__2() { +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__2() { _start: { lean_object* x_1; @@ -33503,20 +33586,20 @@ x_1 = lean_mk_string_unchecked("assertion violation: ys.size == splitterAltInfo. return x_1; } } -static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__3() { +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___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_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__2; +x_1 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__2; x_2 = lean_unsigned_to_nat(8u); x_3 = lean_unsigned_to_nat(321u); -x_4 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__1; -x_5 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__0; +x_4 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__1; +x_5 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); return x_6; } } -LEAN_EXPORT lean_object* l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__54(lean_object* x_1, lean_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_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__55(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { lean_object* x_9; uint8_t x_10; @@ -33951,7 +34034,7 @@ return x_168; } } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_11; @@ -33959,11 +34042,11 @@ x_11 = l_Lean_Meta_instantiateForall(x_1, x_2, x_6, x_7, x_8, x_9); return x_11; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; -x_11 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); @@ -33971,7 +34054,7 @@ lean_dec_ref(x_2); return x_11; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__2(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__2(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { lean_object* x_14; @@ -33979,14 +34062,14 @@ x_14 = l_Lean_Meta_mkLambdaFVars(x_1, x_2, x_3, x_4, x_3, x_4, x_5, x_9, x_10, x return x_14; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { uint8_t x_14; uint8_t x_15; uint8_t x_16; lean_object* x_17; x_14 = lean_unbox(x_3); x_15 = lean_unbox(x_4); x_16 = lean_unbox(x_5); -x_17 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__2(x_1, x_2, x_14, x_15, x_16, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_17 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__2(x_1, x_2, x_14, x_15, x_16, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_12); lean_dec_ref(x_11); lean_dec(x_10); @@ -33998,7 +34081,7 @@ lean_dec_ref(x_1); return x_17; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_11; @@ -34006,11 +34089,11 @@ x_11 = l_Lean_Meta_instantiateLambda(x_1, x_2, x_6, x_7, x_8, x_9); return x_11; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); @@ -34018,7 +34101,7 @@ lean_dec_ref(x_2); return x_11; } } -static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__3___closed__0() { +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__3___closed__0() { _start: { lean_object* x_1; @@ -34026,22 +34109,22 @@ x_1 = lean_mk_string_unchecked("unexpected matcher application, insufficient num return x_1; } } -static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__3___closed__1() { +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__3___closed__1() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__3___closed__0; +x_1 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__3___closed__0; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, uint8_t x_8, uint8_t x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, uint8_t x_8, uint8_t x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { _start: { lean_object* x_19; lean_object* x_20; lean_object* x_34; lean_object* x_35; x_19 = l_Array_append___redArg(x_1, x_2); lean_inc_ref(x_19); -x_34 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__1___boxed), 10, 2); +x_34 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__1___boxed), 10, 2); lean_closure_set(x_34, 0, x_3); lean_closure_set(x_34, 1, x_19); lean_inc(x_17); @@ -34080,7 +34163,7 @@ if (x_37 == 0) { lean_object* x_38; lean_object* x_39; lean_dec_ref(x_35); -x_38 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__3___closed__1; +x_38 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__3___closed__1; lean_inc(x_17); lean_inc_ref(x_16); lean_inc(x_15); @@ -34125,7 +34208,7 @@ x_27 = 1; x_28 = lean_box(x_8); x_29 = lean_box(x_9); x_30 = lean_box(x_27); -x_31 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__2___boxed), 13, 5); +x_31 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__2___boxed), 13, 5); lean_closure_set(x_31, 0, x_26); lean_closure_set(x_31, 1, x_23); lean_closure_set(x_31, 2, x_28); @@ -34165,7 +34248,7 @@ return x_20; } } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__3___boxed(lean_object** _args) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__3___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -34189,20 +34272,20 @@ lean_object* x_18 = _args[17]; uint8_t x_19; uint8_t x_20; lean_object* x_21; x_19 = lean_unbox(x_8); x_20 = lean_unbox(x_9); -x_21 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_19, x_20, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +x_21 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_19, x_20, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); lean_dec_ref(x_10); lean_dec_ref(x_7); lean_dec_ref(x_2); return x_21; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, uint8_t x_7, uint8_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, uint8_t x_7, uint8_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { _start: { lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; x_19 = lean_box(x_7); x_20 = lean_box(x_8); -x_21 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__3___boxed), 18, 9); +x_21 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__3___boxed), 18, 9); lean_closure_set(x_21, 0, x_1); lean_closure_set(x_21, 1, x_10); lean_closure_set(x_21, 2, x_2); @@ -34218,7 +34301,7 @@ x_23 = l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Elab_Tactic_Do return x_23; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__4___boxed(lean_object** _args) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__4___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -34242,17 +34325,17 @@ lean_object* x_18 = _args[17]; uint8_t x_19; uint8_t x_20; lean_object* x_21; x_19 = lean_unbox(x_7); x_20 = lean_unbox(x_8); -x_21 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__4(x_1, x_2, x_3, x_4, x_5, x_6, x_19, x_20, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +x_21 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__4(x_1, x_2, x_3, x_4, x_5, x_6, x_19, x_20, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); return x_21; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { _start: { lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; x_19 = lean_box(x_6); x_20 = lean_box(x_7); -x_21 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__4___boxed), 18, 9); +x_21 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__4___boxed), 18, 9); lean_closure_set(x_21, 0, x_1); lean_closure_set(x_21, 1, x_2); lean_closure_set(x_21, 2, x_3); @@ -34268,7 +34351,7 @@ x_23 = l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Elab_Tactic_Do return x_23; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__6___boxed(lean_object** _args) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__6___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -34292,11 +34375,11 @@ lean_object* x_18 = _args[17]; uint8_t x_19; uint8_t x_20; lean_object* x_21; x_19 = lean_unbox(x_6); x_20 = lean_unbox(x_7); -x_21 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__6(x_1, x_2, x_3, x_4, x_5, x_19, x_20, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +x_21 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__6(x_1, x_2, x_3, x_4, x_5, x_19, x_20, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); return x_21; } } -static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__5() { +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__5() { _start: { lean_object* x_1; @@ -34304,7 +34387,7 @@ x_1 = lean_mk_string_unchecked("const", 5, 5); return x_1; } } -static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__4() { +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__4() { _start: { lean_object* x_1; @@ -34312,17 +34395,17 @@ x_1 = lean_mk_string_unchecked("Function", 8, 8); return x_1; } } -static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__6() { +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__5; -x_2 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__4; +x_1 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__5; +x_2 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__4; x_3 = l_Lean_Name_mkStr2(x_2, x_1); return x_3; } } -static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__7() { +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -34332,17 +34415,17 @@ x_3 = l_Lean_mkConst(x_2, x_1); return x_3; } } -static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__9() { +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__7; -x_2 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__8; +x_1 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__7; +x_2 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__8; x_3 = lean_array_push(x_2, x_1); return x_3; } } -static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__10() { +static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -34352,7 +34435,7 @@ x_3 = l_Lean_mkConst(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_11; @@ -34360,11 +34443,11 @@ x_11 = l_Lean_Meta_instantiateForall(x_1, x_2, x_6, x_7, x_8, x_9); return x_11; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; -x_11 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); @@ -34372,7 +34455,7 @@ lean_dec_ref(x_2); return x_11; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, uint8_t x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, uint8_t x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19) { _start: { lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; uint8_t x_25; @@ -34392,15 +34475,15 @@ lean_dec(x_5); lean_dec_ref(x_4); lean_dec_ref(x_3); lean_dec_ref(x_2); -x_26 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__3; -x_27 = l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__54(x_26, x_14, x_15, x_16, x_17, x_18, x_19); +x_26 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__3; +x_27 = l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__55(x_26, x_14, x_15, x_16, x_17, x_18, x_19); return x_27; } else { lean_object* x_28; lean_object* x_29; lean_inc_ref(x_12); -x_28 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__5___boxed), 10, 2); +x_28 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__5___boxed), 10, 2); lean_closure_set(x_28, 0, x_2); lean_closure_set(x_28, 1, x_12); lean_inc(x_19); @@ -34423,7 +34506,7 @@ if (lean_is_exclusive(x_29)) { } x_32 = lean_box(x_6); x_33 = lean_box(x_7); -x_34 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__6___boxed), 18, 9); +x_34 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__6___boxed), 18, 9); lean_closure_set(x_34, 0, x_13); lean_closure_set(x_34, 1, x_3); lean_closure_set(x_34, 2, x_4); @@ -34448,10 +34531,10 @@ goto block_51; else { lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_52 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__10; +x_52 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__10; x_53 = lean_mk_empty_array_with_capacity(x_11); x_54 = lean_array_push(x_53, x_52); -x_55 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__8___boxed), 10, 2); +x_55 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__8___boxed), 10, 2); lean_closure_set(x_55, 0, x_30); lean_closure_set(x_55, 1, x_54); lean_inc(x_19); @@ -34525,8 +34608,8 @@ lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean x_45 = lean_ctor_get(x_44, 0); lean_inc(x_45); lean_dec_ref(x_44); -x_46 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__6; -x_47 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__9; +x_46 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__6; +x_47 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__9; x_48 = lean_array_push(x_47, x_45); x_49 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___lam__11___boxed), 10, 2); lean_closure_set(x_49, 0, x_46); @@ -34568,7 +34651,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___boxed(lean_object** _args) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -34595,13 +34678,13 @@ uint8_t x_21; uint8_t x_22; uint8_t x_23; lean_object* x_24; x_21 = lean_unbox(x_6); x_22 = lean_unbox(x_7); x_23 = lean_unbox(x_10); -x_24 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7(x_1, x_2, x_3, x_4, x_5, x_21, x_22, x_8, x_9, x_23, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); +x_24 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7(x_1, x_2, x_3, x_4, x_5, x_21, x_22, x_8, x_9, x_23, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); lean_dec(x_11); lean_dec_ref(x_1); return x_24; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__55___redArg___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__56___redArg___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_13; @@ -34609,21 +34692,21 @@ x_13 = lean_apply_9(x_1, x_4, x_5, x_2, x_3, x_8, x_9, x_10, x_11, lean_box(0)); return x_13; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__55___redArg___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__56___redArg___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { lean_object* x_13; -x_13 = l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__55___redArg___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_13 = l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__56___redArg___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec_ref(x_7); lean_dec_ref(x_6); return x_13; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__55___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__56___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_11; lean_object* x_12; -x_11 = lean_alloc_closure((void*)(l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__55___redArg___lam__0___boxed), 12, 3); +x_11 = lean_alloc_closure((void*)(l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__56___redArg___lam__0___boxed), 12, 3); lean_closure_set(x_11, 0, x_3); lean_closure_set(x_11, 1, x_4); lean_closure_set(x_11, 2, x_5); @@ -34653,11 +34736,11 @@ return x_15; } } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { lean_object* x_17; -x_17 = l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__55___redArg(x_1, x_2, x_3, x_10, x_11, x_12, x_13, x_14, x_15); +x_17 = l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__56___redArg(x_1, x_2, x_3, x_10, x_11, x_12, x_13, x_14, x_15); if (lean_obj_tag(x_17) == 0) { uint8_t x_18; @@ -34743,15 +34826,15 @@ return x_38; } } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { lean_object* x_17; -x_17 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_17 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); return x_17; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56_spec__64___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57_spec__65___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { lean_object* x_17; uint8_t x_37; @@ -34830,7 +34913,7 @@ if (x_62 == 0) lean_object* x_63; lean_object* x_64; x_63 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_63, 0, x_9); -x_64 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_64 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_64, 0, x_63); x_17 = x_64; goto block_36; @@ -34866,7 +34949,7 @@ lean_dec(x_60); lean_dec_ref(x_59); x_75 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_75, 0, x_9); -x_76 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_76 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_76, 0, x_75); x_17 = x_76; goto block_36; @@ -34903,7 +34986,7 @@ lean_dec(x_60); lean_dec_ref(x_59); x_86 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_86, 0, x_9); -x_87 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_87 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_87, 0, x_86); x_17 = x_87; goto block_36; @@ -34942,7 +35025,7 @@ lean_dec(x_60); lean_dec_ref(x_59); x_97 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_97, 0, x_9); -x_98 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_98 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_98, 0, x_97); x_17 = x_98; goto block_36; @@ -34983,7 +35066,7 @@ lean_dec(x_60); lean_dec_ref(x_59); x_108 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_108, 0, x_9); -x_109 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_109 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_109, 0, x_108); x_17 = x_109; goto block_36; @@ -35034,8 +35117,8 @@ lean_dec_ref(x_69); lean_dec(x_60); lean_dec_ref(x_59); lean_dec(x_57); -x_118 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__1; -x_119 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___boxed), 8, 1); +x_118 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__1; +x_119 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___boxed), 8, 1); lean_closure_set(x_119, 0, x_118); x_17 = x_119; goto block_36; @@ -35061,7 +35144,7 @@ lean_inc(x_6); lean_inc(x_8); lean_inc_ref(x_3); lean_inc(x_123); -x_127 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___boxed), 20, 11); +x_127 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___boxed), 20, 11); lean_closure_set(x_127, 0, x_122); lean_closure_set(x_127, 1, x_120); lean_closure_set(x_127, 2, x_123); @@ -35076,7 +35159,7 @@ lean_closure_set(x_127, 10, x_72); x_128 = lean_nat_add(x_104, x_72); lean_dec(x_104); lean_ctor_set(x_45, 1, x_128); -x_129 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__9___boxed), 16, 9); +x_129 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__9___boxed), 16, 9); lean_closure_set(x_129, 0, x_121); lean_closure_set(x_129, 1, x_114); lean_closure_set(x_129, 2, x_127); @@ -35118,8 +35201,8 @@ lean_dec_ref(x_69); lean_dec(x_60); lean_dec_ref(x_59); lean_dec(x_57); -x_134 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__1; -x_135 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___boxed), 8, 1); +x_134 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__1; +x_135 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___boxed), 8, 1); lean_closure_set(x_135, 0, x_134); x_17 = x_135; goto block_36; @@ -35145,7 +35228,7 @@ lean_inc(x_6); lean_inc(x_8); lean_inc_ref(x_3); lean_inc(x_139); -x_143 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___boxed), 20, 11); +x_143 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___boxed), 20, 11); lean_closure_set(x_143, 0, x_138); lean_closure_set(x_143, 1, x_136); lean_closure_set(x_143, 2, x_139); @@ -35163,7 +35246,7 @@ x_145 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_145, 0, x_103); lean_ctor_set(x_145, 1, x_144); lean_ctor_set(x_145, 2, x_105); -x_146 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__9___boxed), 16, 9); +x_146 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__9___boxed), 16, 9); lean_closure_set(x_146, 0, x_137); lean_closure_set(x_146, 1, x_130); lean_closure_set(x_146, 2, x_143); @@ -35207,7 +35290,7 @@ lean_dec_ref(x_59); lean_ctor_set(x_39, 0, x_151); x_153 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_153, 0, x_9); -x_154 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_154 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_154, 0, x_153); x_17 = x_154; goto block_36; @@ -35257,8 +35340,8 @@ lean_dec_ref(x_69); lean_dec(x_60); lean_dec_ref(x_59); lean_dec(x_57); -x_160 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__1; -x_161 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___boxed), 8, 1); +x_160 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__1; +x_161 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___boxed), 8, 1); lean_closure_set(x_161, 0, x_160); x_17 = x_161; goto block_36; @@ -35284,7 +35367,7 @@ lean_inc(x_6); lean_inc(x_8); lean_inc_ref(x_3); lean_inc(x_165); -x_169 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___boxed), 20, 11); +x_169 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___boxed), 20, 11); lean_closure_set(x_169, 0, x_164); lean_closure_set(x_169, 1, x_162); lean_closure_set(x_169, 2, x_165); @@ -35306,7 +35389,7 @@ if (lean_is_scalar(x_155)) { lean_ctor_set(x_171, 0, x_147); lean_ctor_set(x_171, 1, x_170); lean_ctor_set(x_171, 2, x_149); -x_172 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__9___boxed), 16, 9); +x_172 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__9___boxed), 16, 9); lean_closure_set(x_172, 0, x_163); lean_closure_set(x_172, 1, x_156); lean_closure_set(x_172, 2, x_169); @@ -35349,7 +35432,7 @@ lean_dec_ref(x_59); lean_ctor_set(x_40, 0, x_177); x_179 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_179, 0, x_9); -x_180 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_180 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_180, 0, x_179); x_17 = x_180; goto block_36; @@ -35398,7 +35481,7 @@ lean_ctor_set(x_40, 0, x_177); lean_ctor_set(x_39, 0, x_186); x_188 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_188, 0, x_9); -x_189 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_189 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_189, 0, x_188); x_17 = x_189; goto block_36; @@ -35448,8 +35531,8 @@ lean_dec_ref(x_69); lean_dec(x_60); lean_dec_ref(x_59); lean_dec(x_57); -x_195 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__1; -x_196 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___boxed), 8, 1); +x_195 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__1; +x_196 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___boxed), 8, 1); lean_closure_set(x_196, 0, x_195); x_17 = x_196; goto block_36; @@ -35475,7 +35558,7 @@ lean_inc(x_6); lean_inc(x_8); lean_inc_ref(x_3); lean_inc(x_200); -x_204 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___boxed), 20, 11); +x_204 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___boxed), 20, 11); lean_closure_set(x_204, 0, x_199); lean_closure_set(x_204, 1, x_197); lean_closure_set(x_204, 2, x_200); @@ -35497,7 +35580,7 @@ if (lean_is_scalar(x_190)) { lean_ctor_set(x_206, 0, x_182); lean_ctor_set(x_206, 1, x_205); lean_ctor_set(x_206, 2, x_184); -x_207 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__9___boxed), 16, 9); +x_207 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__9___boxed), 16, 9); lean_closure_set(x_207, 0, x_198); lean_closure_set(x_207, 1, x_191); lean_closure_set(x_207, 2, x_204); @@ -35539,7 +35622,7 @@ lean_dec_ref(x_59); lean_ctor_set(x_41, 0, x_212); x_214 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_214, 0, x_9); -x_215 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_215 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_215, 0, x_214); x_17 = x_215; goto block_36; @@ -35586,7 +35669,7 @@ lean_ctor_set(x_41, 0, x_212); lean_ctor_set(x_40, 0, x_221); x_223 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_223, 0, x_9); -x_224 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_224 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_224, 0, x_223); x_17 = x_224; goto block_36; @@ -35636,7 +35719,7 @@ lean_ctor_set(x_40, 0, x_221); lean_ctor_set(x_39, 0, x_230); x_232 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_232, 0, x_9); -x_233 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_233 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_233, 0, x_232); x_17 = x_233; goto block_36; @@ -35686,8 +35769,8 @@ lean_dec_ref(x_69); lean_dec(x_60); lean_dec_ref(x_59); lean_dec(x_57); -x_239 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__1; -x_240 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___boxed), 8, 1); +x_239 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__1; +x_240 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___boxed), 8, 1); lean_closure_set(x_240, 0, x_239); x_17 = x_240; goto block_36; @@ -35713,7 +35796,7 @@ lean_inc(x_6); lean_inc(x_8); lean_inc_ref(x_3); lean_inc(x_244); -x_248 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___boxed), 20, 11); +x_248 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___boxed), 20, 11); lean_closure_set(x_248, 0, x_243); lean_closure_set(x_248, 1, x_241); lean_closure_set(x_248, 2, x_244); @@ -35735,7 +35818,7 @@ if (lean_is_scalar(x_234)) { lean_ctor_set(x_250, 0, x_226); lean_ctor_set(x_250, 1, x_249); lean_ctor_set(x_250, 2, x_228); -x_251 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__9___boxed), 16, 9); +x_251 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__9___boxed), 16, 9); lean_closure_set(x_251, 0, x_242); lean_closure_set(x_251, 1, x_235); lean_closure_set(x_251, 2, x_248); @@ -35777,7 +35860,7 @@ lean_dec_ref(x_59); lean_ctor_set(x_42, 0, x_257); x_259 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_259, 0, x_9); -x_260 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_260 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_260, 0, x_259); x_17 = x_260; goto block_36; @@ -35822,7 +35905,7 @@ lean_ctor_set(x_42, 0, x_257); lean_ctor_set(x_41, 0, x_266); x_268 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_268, 0, x_9); -x_269 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_269 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_269, 0, x_268); x_17 = x_269; goto block_36; @@ -35870,7 +35953,7 @@ lean_ctor_set(x_41, 0, x_266); lean_ctor_set(x_40, 0, x_275); x_277 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_277, 0, x_9); -x_278 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_278 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_278, 0, x_277); x_17 = x_278; goto block_36; @@ -35921,7 +36004,7 @@ lean_ctor_set(x_40, 0, x_275); lean_ctor_set(x_39, 0, x_284); x_286 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_286, 0, x_9); -x_287 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_287 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_287, 0, x_286); x_17 = x_287; goto block_36; @@ -35971,8 +36054,8 @@ lean_dec_ref(x_252); lean_dec(x_60); lean_dec_ref(x_59); lean_dec(x_57); -x_293 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__1; -x_294 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___boxed), 8, 1); +x_293 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__1; +x_294 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___boxed), 8, 1); lean_closure_set(x_294, 0, x_293); x_17 = x_294; goto block_36; @@ -35998,7 +36081,7 @@ lean_inc(x_6); lean_inc(x_8); lean_inc_ref(x_3); lean_inc(x_298); -x_302 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___boxed), 20, 11); +x_302 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___boxed), 20, 11); lean_closure_set(x_302, 0, x_297); lean_closure_set(x_302, 1, x_295); lean_closure_set(x_302, 2, x_298); @@ -36020,7 +36103,7 @@ if (lean_is_scalar(x_288)) { lean_ctor_set(x_304, 0, x_280); lean_ctor_set(x_304, 1, x_303); lean_ctor_set(x_304, 2, x_282); -x_305 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__9___boxed), 16, 9); +x_305 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__9___boxed), 16, 9); lean_closure_set(x_305, 0, x_296); lean_closure_set(x_305, 1, x_289); lean_closure_set(x_305, 2, x_302); @@ -36059,7 +36142,7 @@ lean_ctor_set(x_311, 1, x_306); lean_ctor_set(x_41, 1, x_311); x_312 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_312, 0, x_9); -x_313 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_313 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_313, 0, x_312); x_17 = x_313; goto block_36; @@ -36105,7 +36188,7 @@ lean_ctor_set(x_322, 1, x_306); lean_ctor_set(x_41, 1, x_322); x_323 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_323, 0, x_9); -x_324 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_324 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_324, 0, x_323); x_17 = x_324; goto block_36; @@ -36153,7 +36236,7 @@ lean_ctor_set(x_41, 1, x_332); lean_ctor_set(x_41, 0, x_330); x_333 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_333, 0, x_9); -x_334 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_334 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_334, 0, x_333); x_17 = x_334; goto block_36; @@ -36204,7 +36287,7 @@ lean_ctor_set(x_41, 0, x_330); lean_ctor_set(x_40, 0, x_340); x_343 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_343, 0, x_9); -x_344 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_344 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_344, 0, x_343); x_17 = x_344; goto block_36; @@ -36258,7 +36341,7 @@ lean_ctor_set(x_40, 0, x_340); lean_ctor_set(x_39, 0, x_350); x_353 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_353, 0, x_9); -x_354 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_354 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_354, 0, x_353); x_17 = x_354; goto block_36; @@ -36307,8 +36390,8 @@ lean_dec_ref(x_315); lean_dec(x_308); lean_dec_ref(x_307); lean_dec(x_306); -x_360 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__1; -x_361 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___boxed), 8, 1); +x_360 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__1; +x_361 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___boxed), 8, 1); lean_closure_set(x_361, 0, x_360); x_17 = x_361; goto block_36; @@ -36334,7 +36417,7 @@ lean_inc(x_6); lean_inc(x_8); lean_inc_ref(x_3); lean_inc(x_365); -x_369 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___boxed), 20, 11); +x_369 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___boxed), 20, 11); lean_closure_set(x_369, 0, x_364); lean_closure_set(x_369, 1, x_362); lean_closure_set(x_369, 2, x_365); @@ -36356,7 +36439,7 @@ if (lean_is_scalar(x_355)) { lean_ctor_set(x_371, 0, x_346); lean_ctor_set(x_371, 1, x_370); lean_ctor_set(x_371, 2, x_348); -x_372 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__9___boxed), 16, 9); +x_372 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__9___boxed), 16, 9); lean_closure_set(x_372, 0, x_363); lean_closure_set(x_372, 1, x_356); lean_closure_set(x_372, 2, x_369); @@ -36412,7 +36495,7 @@ lean_ctor_set(x_381, 1, x_380); lean_ctor_set(x_40, 1, x_381); x_382 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_382, 0, x_9); -x_383 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_383 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_383, 0, x_382); x_17 = x_383; goto block_36; @@ -36465,7 +36548,7 @@ lean_ctor_set(x_393, 1, x_392); lean_ctor_set(x_40, 1, x_393); x_394 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_394, 0, x_9); -x_395 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_395 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_395, 0, x_394); x_17 = x_395; goto block_36; @@ -36519,7 +36602,7 @@ lean_ctor_set(x_404, 1, x_403); lean_ctor_set(x_40, 1, x_404); x_405 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_405, 0, x_9); -x_406 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_406 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_406, 0, x_405); x_17 = x_406; goto block_36; @@ -36576,7 +36659,7 @@ lean_ctor_set(x_40, 1, x_415); lean_ctor_set(x_40, 0, x_412); x_416 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_416, 0, x_9); -x_417 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_417 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_417, 0, x_416); x_17 = x_417; goto block_36; @@ -36636,7 +36719,7 @@ lean_ctor_set(x_40, 0, x_412); lean_ctor_set(x_39, 0, x_423); x_427 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_427, 0, x_9); -x_428 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_428 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_428, 0, x_427); x_17 = x_428; goto block_36; @@ -36685,8 +36768,8 @@ lean_dec_ref(x_385); lean_dec(x_377); lean_dec_ref(x_376); lean_dec(x_374); -x_434 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__1; -x_435 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___boxed), 8, 1); +x_434 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__1; +x_435 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___boxed), 8, 1); lean_closure_set(x_435, 0, x_434); x_17 = x_435; goto block_36; @@ -36712,7 +36795,7 @@ lean_inc(x_6); lean_inc(x_8); lean_inc_ref(x_3); lean_inc(x_439); -x_443 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___boxed), 20, 11); +x_443 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___boxed), 20, 11); lean_closure_set(x_443, 0, x_438); lean_closure_set(x_443, 1, x_436); lean_closure_set(x_443, 2, x_439); @@ -36734,7 +36817,7 @@ if (lean_is_scalar(x_429)) { lean_ctor_set(x_445, 0, x_419); lean_ctor_set(x_445, 1, x_444); lean_ctor_set(x_445, 2, x_421); -x_446 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__9___boxed), 16, 9); +x_446 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__9___boxed), 16, 9); lean_closure_set(x_446, 0, x_437); lean_closure_set(x_446, 1, x_430); lean_closure_set(x_446, 2, x_443); @@ -36807,7 +36890,7 @@ lean_ctor_set(x_458, 1, x_457); lean_ctor_set(x_39, 1, x_458); x_459 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_459, 0, x_9); -x_460 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_460 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_460, 0, x_459); x_17 = x_460; goto block_36; @@ -36867,7 +36950,7 @@ lean_ctor_set(x_471, 1, x_470); lean_ctor_set(x_39, 1, x_471); x_472 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_472, 0, x_9); -x_473 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_473 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_473, 0, x_472); x_17 = x_473; goto block_36; @@ -36928,7 +37011,7 @@ lean_ctor_set(x_483, 1, x_482); lean_ctor_set(x_39, 1, x_483); x_484 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_484, 0, x_9); -x_485 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_485 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_485, 0, x_484); x_17 = x_485; goto block_36; @@ -36991,7 +37074,7 @@ lean_ctor_set(x_495, 1, x_494); lean_ctor_set(x_39, 1, x_495); x_496 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_496, 0, x_9); -x_497 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_497 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_497, 0, x_496); x_17 = x_497; goto block_36; @@ -37057,7 +37140,7 @@ lean_ctor_set(x_39, 1, x_507); lean_ctor_set(x_39, 0, x_503); x_508 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_508, 0, x_9); -x_509 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_509 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_509, 0, x_508); x_17 = x_509; goto block_36; @@ -37106,8 +37189,8 @@ lean_dec_ref(x_462); lean_dec(x_453); lean_dec_ref(x_452); lean_dec(x_450); -x_515 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__1; -x_516 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___boxed), 8, 1); +x_515 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__1; +x_516 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___boxed), 8, 1); lean_closure_set(x_516, 0, x_515); x_17 = x_516; goto block_36; @@ -37133,7 +37216,7 @@ lean_inc(x_6); lean_inc(x_8); lean_inc_ref(x_3); lean_inc(x_520); -x_524 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___boxed), 20, 11); +x_524 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___boxed), 20, 11); lean_closure_set(x_524, 0, x_519); lean_closure_set(x_524, 1, x_517); lean_closure_set(x_524, 2, x_520); @@ -37155,7 +37238,7 @@ if (lean_is_scalar(x_510)) { lean_ctor_set(x_526, 0, x_499); lean_ctor_set(x_526, 1, x_525); lean_ctor_set(x_526, 2, x_501); -x_527 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__9___boxed), 16, 9); +x_527 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__9___boxed), 16, 9); lean_closure_set(x_527, 0, x_518); lean_closure_set(x_527, 1, x_511); lean_closure_set(x_527, 2, x_524); @@ -37245,7 +37328,7 @@ lean_ctor_set(x_542, 1, x_541); lean_ctor_set(x_9, 1, x_542); x_543 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_543, 0, x_9); -x_544 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_544 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_544, 0, x_543); x_17 = x_544; goto block_36; @@ -37312,7 +37395,7 @@ lean_ctor_set(x_556, 1, x_555); lean_ctor_set(x_9, 1, x_556); x_557 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_557, 0, x_9); -x_558 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_558 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_558, 0, x_557); x_17 = x_558; goto block_36; @@ -37380,7 +37463,7 @@ lean_ctor_set(x_569, 1, x_568); lean_ctor_set(x_9, 1, x_569); x_570 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_570, 0, x_9); -x_571 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_571 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_571, 0, x_570); x_17 = x_571; goto block_36; @@ -37450,7 +37533,7 @@ lean_ctor_set(x_582, 1, x_581); lean_ctor_set(x_9, 1, x_582); x_583 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_583, 0, x_9); -x_584 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_584 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_584, 0, x_583); x_17 = x_584; goto block_36; @@ -37522,7 +37605,7 @@ lean_ctor_set(x_595, 1, x_594); lean_ctor_set(x_9, 1, x_595); x_596 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_596, 0, x_9); -x_597 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_597 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_597, 0, x_596); x_17 = x_597; goto block_36; @@ -37571,8 +37654,8 @@ lean_dec_ref(x_546); lean_dec(x_536); lean_dec_ref(x_535); lean_dec(x_533); -x_603 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__1; -x_604 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___boxed), 8, 1); +x_603 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__1; +x_604 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___boxed), 8, 1); lean_closure_set(x_604, 0, x_603); x_17 = x_604; goto block_36; @@ -37598,7 +37681,7 @@ lean_inc(x_6); lean_inc(x_8); lean_inc_ref(x_3); lean_inc(x_608); -x_612 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___boxed), 20, 11); +x_612 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___boxed), 20, 11); lean_closure_set(x_612, 0, x_607); lean_closure_set(x_612, 1, x_605); lean_closure_set(x_612, 2, x_608); @@ -37620,7 +37703,7 @@ if (lean_is_scalar(x_598)) { lean_ctor_set(x_614, 0, x_586); lean_ctor_set(x_614, 1, x_613); lean_ctor_set(x_614, 2, x_588); -x_615 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__9___boxed), 16, 9); +x_615 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__9___boxed), 16, 9); lean_closure_set(x_615, 0, x_606); lean_closure_set(x_615, 1, x_599); lean_closure_set(x_615, 2, x_612); @@ -37726,7 +37809,7 @@ lean_ctor_set(x_633, 0, x_616); lean_ctor_set(x_633, 1, x_632); x_634 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_634, 0, x_633); -x_635 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_635 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_635, 0, x_634); x_17 = x_635; goto block_36; @@ -37799,7 +37882,7 @@ lean_ctor_set(x_648, 0, x_616); lean_ctor_set(x_648, 1, x_647); x_649 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_649, 0, x_648); -x_650 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_650 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_650, 0, x_649); x_17 = x_650; goto block_36; @@ -37873,7 +37956,7 @@ lean_ctor_set(x_662, 0, x_616); lean_ctor_set(x_662, 1, x_661); x_663 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_663, 0, x_662); -x_664 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_664 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_664, 0, x_663); x_17 = x_664; goto block_36; @@ -37949,7 +38032,7 @@ lean_ctor_set(x_676, 0, x_616); lean_ctor_set(x_676, 1, x_675); x_677 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_677, 0, x_676); -x_678 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_678 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_678, 0, x_677); x_17 = x_678; goto block_36; @@ -38027,7 +38110,7 @@ lean_ctor_set(x_690, 0, x_616); lean_ctor_set(x_690, 1, x_689); x_691 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_691, 0, x_690); -x_692 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_692 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_692, 0, x_691); x_17 = x_692; goto block_36; @@ -38076,8 +38159,8 @@ lean_dec_ref(x_637); lean_dec(x_626); lean_dec_ref(x_625); lean_dec(x_623); -x_698 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__1; -x_699 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___boxed), 8, 1); +x_698 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__1; +x_699 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___boxed), 8, 1); lean_closure_set(x_699, 0, x_698); x_17 = x_699; goto block_36; @@ -38103,7 +38186,7 @@ lean_inc(x_6); lean_inc(x_8); lean_inc_ref(x_3); lean_inc(x_703); -x_707 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___boxed), 20, 11); +x_707 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___boxed), 20, 11); lean_closure_set(x_707, 0, x_702); lean_closure_set(x_707, 1, x_700); lean_closure_set(x_707, 2, x_703); @@ -38125,7 +38208,7 @@ if (lean_is_scalar(x_693)) { lean_ctor_set(x_709, 0, x_680); lean_ctor_set(x_709, 1, x_708); lean_ctor_set(x_709, 2, x_682); -x_710 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__9___boxed), 16, 9); +x_710 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__9___boxed), 16, 9); lean_closure_set(x_710, 0, x_701); lean_closure_set(x_710, 1, x_694); lean_closure_set(x_710, 2, x_707); @@ -38270,7 +38353,7 @@ return x_35; } } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { lean_object* x_17; uint8_t x_37; @@ -38349,7 +38432,7 @@ if (x_62 == 0) lean_object* x_63; lean_object* x_64; x_63 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_63, 0, x_9); -x_64 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_64 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_64, 0, x_63); x_17 = x_64; goto block_36; @@ -38385,7 +38468,7 @@ lean_dec(x_60); lean_dec_ref(x_59); x_75 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_75, 0, x_9); -x_76 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_76 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_76, 0, x_75); x_17 = x_76; goto block_36; @@ -38422,7 +38505,7 @@ lean_dec(x_60); lean_dec_ref(x_59); x_86 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_86, 0, x_9); -x_87 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_87 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_87, 0, x_86); x_17 = x_87; goto block_36; @@ -38461,7 +38544,7 @@ lean_dec(x_60); lean_dec_ref(x_59); x_97 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_97, 0, x_9); -x_98 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_98 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_98, 0, x_97); x_17 = x_98; goto block_36; @@ -38502,7 +38585,7 @@ lean_dec(x_60); lean_dec_ref(x_59); x_108 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_108, 0, x_9); -x_109 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_109 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_109, 0, x_108); x_17 = x_109; goto block_36; @@ -38553,8 +38636,8 @@ lean_dec_ref(x_69); lean_dec(x_60); lean_dec_ref(x_59); lean_dec(x_57); -x_118 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__1; -x_119 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___boxed), 8, 1); +x_118 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__1; +x_119 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___boxed), 8, 1); lean_closure_set(x_119, 0, x_118); x_17 = x_119; goto block_36; @@ -38580,7 +38663,7 @@ lean_inc(x_6); lean_inc(x_8); lean_inc_ref(x_3); lean_inc(x_123); -x_127 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___boxed), 20, 11); +x_127 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___boxed), 20, 11); lean_closure_set(x_127, 0, x_122); lean_closure_set(x_127, 1, x_120); lean_closure_set(x_127, 2, x_123); @@ -38595,7 +38678,7 @@ lean_closure_set(x_127, 10, x_72); x_128 = lean_nat_add(x_104, x_72); lean_dec(x_104); lean_ctor_set(x_45, 1, x_128); -x_129 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__9___boxed), 16, 9); +x_129 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__9___boxed), 16, 9); lean_closure_set(x_129, 0, x_121); lean_closure_set(x_129, 1, x_114); lean_closure_set(x_129, 2, x_127); @@ -38637,8 +38720,8 @@ lean_dec_ref(x_69); lean_dec(x_60); lean_dec_ref(x_59); lean_dec(x_57); -x_134 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__1; -x_135 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___boxed), 8, 1); +x_134 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__1; +x_135 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___boxed), 8, 1); lean_closure_set(x_135, 0, x_134); x_17 = x_135; goto block_36; @@ -38664,7 +38747,7 @@ lean_inc(x_6); lean_inc(x_8); lean_inc_ref(x_3); lean_inc(x_139); -x_143 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___boxed), 20, 11); +x_143 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___boxed), 20, 11); lean_closure_set(x_143, 0, x_138); lean_closure_set(x_143, 1, x_136); lean_closure_set(x_143, 2, x_139); @@ -38682,7 +38765,7 @@ x_145 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_145, 0, x_103); lean_ctor_set(x_145, 1, x_144); lean_ctor_set(x_145, 2, x_105); -x_146 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__9___boxed), 16, 9); +x_146 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__9___boxed), 16, 9); lean_closure_set(x_146, 0, x_137); lean_closure_set(x_146, 1, x_130); lean_closure_set(x_146, 2, x_143); @@ -38726,7 +38809,7 @@ lean_dec_ref(x_59); lean_ctor_set(x_39, 0, x_151); x_153 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_153, 0, x_9); -x_154 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_154 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_154, 0, x_153); x_17 = x_154; goto block_36; @@ -38776,8 +38859,8 @@ lean_dec_ref(x_69); lean_dec(x_60); lean_dec_ref(x_59); lean_dec(x_57); -x_160 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__1; -x_161 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___boxed), 8, 1); +x_160 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__1; +x_161 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___boxed), 8, 1); lean_closure_set(x_161, 0, x_160); x_17 = x_161; goto block_36; @@ -38803,7 +38886,7 @@ lean_inc(x_6); lean_inc(x_8); lean_inc_ref(x_3); lean_inc(x_165); -x_169 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___boxed), 20, 11); +x_169 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___boxed), 20, 11); lean_closure_set(x_169, 0, x_164); lean_closure_set(x_169, 1, x_162); lean_closure_set(x_169, 2, x_165); @@ -38825,7 +38908,7 @@ if (lean_is_scalar(x_155)) { lean_ctor_set(x_171, 0, x_147); lean_ctor_set(x_171, 1, x_170); lean_ctor_set(x_171, 2, x_149); -x_172 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__9___boxed), 16, 9); +x_172 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__9___boxed), 16, 9); lean_closure_set(x_172, 0, x_163); lean_closure_set(x_172, 1, x_156); lean_closure_set(x_172, 2, x_169); @@ -38868,7 +38951,7 @@ lean_dec_ref(x_59); lean_ctor_set(x_40, 0, x_177); x_179 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_179, 0, x_9); -x_180 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_180 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_180, 0, x_179); x_17 = x_180; goto block_36; @@ -38917,7 +39000,7 @@ lean_ctor_set(x_40, 0, x_177); lean_ctor_set(x_39, 0, x_186); x_188 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_188, 0, x_9); -x_189 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_189 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_189, 0, x_188); x_17 = x_189; goto block_36; @@ -38967,8 +39050,8 @@ lean_dec_ref(x_69); lean_dec(x_60); lean_dec_ref(x_59); lean_dec(x_57); -x_195 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__1; -x_196 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___boxed), 8, 1); +x_195 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__1; +x_196 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___boxed), 8, 1); lean_closure_set(x_196, 0, x_195); x_17 = x_196; goto block_36; @@ -38994,7 +39077,7 @@ lean_inc(x_6); lean_inc(x_8); lean_inc_ref(x_3); lean_inc(x_200); -x_204 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___boxed), 20, 11); +x_204 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___boxed), 20, 11); lean_closure_set(x_204, 0, x_199); lean_closure_set(x_204, 1, x_197); lean_closure_set(x_204, 2, x_200); @@ -39016,7 +39099,7 @@ if (lean_is_scalar(x_190)) { lean_ctor_set(x_206, 0, x_182); lean_ctor_set(x_206, 1, x_205); lean_ctor_set(x_206, 2, x_184); -x_207 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__9___boxed), 16, 9); +x_207 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__9___boxed), 16, 9); lean_closure_set(x_207, 0, x_198); lean_closure_set(x_207, 1, x_191); lean_closure_set(x_207, 2, x_204); @@ -39058,7 +39141,7 @@ lean_dec_ref(x_59); lean_ctor_set(x_41, 0, x_212); x_214 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_214, 0, x_9); -x_215 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_215 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_215, 0, x_214); x_17 = x_215; goto block_36; @@ -39105,7 +39188,7 @@ lean_ctor_set(x_41, 0, x_212); lean_ctor_set(x_40, 0, x_221); x_223 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_223, 0, x_9); -x_224 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_224 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_224, 0, x_223); x_17 = x_224; goto block_36; @@ -39155,7 +39238,7 @@ lean_ctor_set(x_40, 0, x_221); lean_ctor_set(x_39, 0, x_230); x_232 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_232, 0, x_9); -x_233 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_233 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_233, 0, x_232); x_17 = x_233; goto block_36; @@ -39205,8 +39288,8 @@ lean_dec_ref(x_69); lean_dec(x_60); lean_dec_ref(x_59); lean_dec(x_57); -x_239 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__1; -x_240 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___boxed), 8, 1); +x_239 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__1; +x_240 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___boxed), 8, 1); lean_closure_set(x_240, 0, x_239); x_17 = x_240; goto block_36; @@ -39232,7 +39315,7 @@ lean_inc(x_6); lean_inc(x_8); lean_inc_ref(x_3); lean_inc(x_244); -x_248 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___boxed), 20, 11); +x_248 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___boxed), 20, 11); lean_closure_set(x_248, 0, x_243); lean_closure_set(x_248, 1, x_241); lean_closure_set(x_248, 2, x_244); @@ -39254,7 +39337,7 @@ if (lean_is_scalar(x_234)) { lean_ctor_set(x_250, 0, x_226); lean_ctor_set(x_250, 1, x_249); lean_ctor_set(x_250, 2, x_228); -x_251 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__9___boxed), 16, 9); +x_251 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__9___boxed), 16, 9); lean_closure_set(x_251, 0, x_242); lean_closure_set(x_251, 1, x_235); lean_closure_set(x_251, 2, x_248); @@ -39296,7 +39379,7 @@ lean_dec_ref(x_59); lean_ctor_set(x_42, 0, x_257); x_259 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_259, 0, x_9); -x_260 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_260 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_260, 0, x_259); x_17 = x_260; goto block_36; @@ -39341,7 +39424,7 @@ lean_ctor_set(x_42, 0, x_257); lean_ctor_set(x_41, 0, x_266); x_268 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_268, 0, x_9); -x_269 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_269 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_269, 0, x_268); x_17 = x_269; goto block_36; @@ -39389,7 +39472,7 @@ lean_ctor_set(x_41, 0, x_266); lean_ctor_set(x_40, 0, x_275); x_277 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_277, 0, x_9); -x_278 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_278 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_278, 0, x_277); x_17 = x_278; goto block_36; @@ -39440,7 +39523,7 @@ lean_ctor_set(x_40, 0, x_275); lean_ctor_set(x_39, 0, x_284); x_286 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_286, 0, x_9); -x_287 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_287 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_287, 0, x_286); x_17 = x_287; goto block_36; @@ -39490,8 +39573,8 @@ lean_dec_ref(x_252); lean_dec(x_60); lean_dec_ref(x_59); lean_dec(x_57); -x_293 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__1; -x_294 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___boxed), 8, 1); +x_293 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__1; +x_294 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___boxed), 8, 1); lean_closure_set(x_294, 0, x_293); x_17 = x_294; goto block_36; @@ -39517,7 +39600,7 @@ lean_inc(x_6); lean_inc(x_8); lean_inc_ref(x_3); lean_inc(x_298); -x_302 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___boxed), 20, 11); +x_302 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___boxed), 20, 11); lean_closure_set(x_302, 0, x_297); lean_closure_set(x_302, 1, x_295); lean_closure_set(x_302, 2, x_298); @@ -39539,7 +39622,7 @@ if (lean_is_scalar(x_288)) { lean_ctor_set(x_304, 0, x_280); lean_ctor_set(x_304, 1, x_303); lean_ctor_set(x_304, 2, x_282); -x_305 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__9___boxed), 16, 9); +x_305 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__9___boxed), 16, 9); lean_closure_set(x_305, 0, x_296); lean_closure_set(x_305, 1, x_289); lean_closure_set(x_305, 2, x_302); @@ -39578,7 +39661,7 @@ lean_ctor_set(x_311, 1, x_306); lean_ctor_set(x_41, 1, x_311); x_312 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_312, 0, x_9); -x_313 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_313 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_313, 0, x_312); x_17 = x_313; goto block_36; @@ -39624,7 +39707,7 @@ lean_ctor_set(x_322, 1, x_306); lean_ctor_set(x_41, 1, x_322); x_323 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_323, 0, x_9); -x_324 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_324 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_324, 0, x_323); x_17 = x_324; goto block_36; @@ -39672,7 +39755,7 @@ lean_ctor_set(x_41, 1, x_332); lean_ctor_set(x_41, 0, x_330); x_333 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_333, 0, x_9); -x_334 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_334 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_334, 0, x_333); x_17 = x_334; goto block_36; @@ -39723,7 +39806,7 @@ lean_ctor_set(x_41, 0, x_330); lean_ctor_set(x_40, 0, x_340); x_343 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_343, 0, x_9); -x_344 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_344 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_344, 0, x_343); x_17 = x_344; goto block_36; @@ -39777,7 +39860,7 @@ lean_ctor_set(x_40, 0, x_340); lean_ctor_set(x_39, 0, x_350); x_353 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_353, 0, x_9); -x_354 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_354 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_354, 0, x_353); x_17 = x_354; goto block_36; @@ -39826,8 +39909,8 @@ lean_dec_ref(x_315); lean_dec(x_308); lean_dec_ref(x_307); lean_dec(x_306); -x_360 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__1; -x_361 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___boxed), 8, 1); +x_360 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__1; +x_361 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___boxed), 8, 1); lean_closure_set(x_361, 0, x_360); x_17 = x_361; goto block_36; @@ -39853,7 +39936,7 @@ lean_inc(x_6); lean_inc(x_8); lean_inc_ref(x_3); lean_inc(x_365); -x_369 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___boxed), 20, 11); +x_369 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___boxed), 20, 11); lean_closure_set(x_369, 0, x_364); lean_closure_set(x_369, 1, x_362); lean_closure_set(x_369, 2, x_365); @@ -39875,7 +39958,7 @@ if (lean_is_scalar(x_355)) { lean_ctor_set(x_371, 0, x_346); lean_ctor_set(x_371, 1, x_370); lean_ctor_set(x_371, 2, x_348); -x_372 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__9___boxed), 16, 9); +x_372 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__9___boxed), 16, 9); lean_closure_set(x_372, 0, x_363); lean_closure_set(x_372, 1, x_356); lean_closure_set(x_372, 2, x_369); @@ -39931,7 +40014,7 @@ lean_ctor_set(x_381, 1, x_380); lean_ctor_set(x_40, 1, x_381); x_382 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_382, 0, x_9); -x_383 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_383 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_383, 0, x_382); x_17 = x_383; goto block_36; @@ -39984,7 +40067,7 @@ lean_ctor_set(x_393, 1, x_392); lean_ctor_set(x_40, 1, x_393); x_394 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_394, 0, x_9); -x_395 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_395 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_395, 0, x_394); x_17 = x_395; goto block_36; @@ -40038,7 +40121,7 @@ lean_ctor_set(x_404, 1, x_403); lean_ctor_set(x_40, 1, x_404); x_405 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_405, 0, x_9); -x_406 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_406 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_406, 0, x_405); x_17 = x_406; goto block_36; @@ -40095,7 +40178,7 @@ lean_ctor_set(x_40, 1, x_415); lean_ctor_set(x_40, 0, x_412); x_416 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_416, 0, x_9); -x_417 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_417 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_417, 0, x_416); x_17 = x_417; goto block_36; @@ -40155,7 +40238,7 @@ lean_ctor_set(x_40, 0, x_412); lean_ctor_set(x_39, 0, x_423); x_427 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_427, 0, x_9); -x_428 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_428 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_428, 0, x_427); x_17 = x_428; goto block_36; @@ -40204,8 +40287,8 @@ lean_dec_ref(x_385); lean_dec(x_377); lean_dec_ref(x_376); lean_dec(x_374); -x_434 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__1; -x_435 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___boxed), 8, 1); +x_434 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__1; +x_435 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___boxed), 8, 1); lean_closure_set(x_435, 0, x_434); x_17 = x_435; goto block_36; @@ -40231,7 +40314,7 @@ lean_inc(x_6); lean_inc(x_8); lean_inc_ref(x_3); lean_inc(x_439); -x_443 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___boxed), 20, 11); +x_443 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___boxed), 20, 11); lean_closure_set(x_443, 0, x_438); lean_closure_set(x_443, 1, x_436); lean_closure_set(x_443, 2, x_439); @@ -40253,7 +40336,7 @@ if (lean_is_scalar(x_429)) { lean_ctor_set(x_445, 0, x_419); lean_ctor_set(x_445, 1, x_444); lean_ctor_set(x_445, 2, x_421); -x_446 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__9___boxed), 16, 9); +x_446 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__9___boxed), 16, 9); lean_closure_set(x_446, 0, x_437); lean_closure_set(x_446, 1, x_430); lean_closure_set(x_446, 2, x_443); @@ -40326,7 +40409,7 @@ lean_ctor_set(x_458, 1, x_457); lean_ctor_set(x_39, 1, x_458); x_459 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_459, 0, x_9); -x_460 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_460 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_460, 0, x_459); x_17 = x_460; goto block_36; @@ -40386,7 +40469,7 @@ lean_ctor_set(x_471, 1, x_470); lean_ctor_set(x_39, 1, x_471); x_472 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_472, 0, x_9); -x_473 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_473 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_473, 0, x_472); x_17 = x_473; goto block_36; @@ -40447,7 +40530,7 @@ lean_ctor_set(x_483, 1, x_482); lean_ctor_set(x_39, 1, x_483); x_484 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_484, 0, x_9); -x_485 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_485 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_485, 0, x_484); x_17 = x_485; goto block_36; @@ -40510,7 +40593,7 @@ lean_ctor_set(x_495, 1, x_494); lean_ctor_set(x_39, 1, x_495); x_496 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_496, 0, x_9); -x_497 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_497 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_497, 0, x_496); x_17 = x_497; goto block_36; @@ -40576,7 +40659,7 @@ lean_ctor_set(x_39, 1, x_507); lean_ctor_set(x_39, 0, x_503); x_508 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_508, 0, x_9); -x_509 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_509 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_509, 0, x_508); x_17 = x_509; goto block_36; @@ -40625,8 +40708,8 @@ lean_dec_ref(x_462); lean_dec(x_453); lean_dec_ref(x_452); lean_dec(x_450); -x_515 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__1; -x_516 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___boxed), 8, 1); +x_515 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__1; +x_516 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___boxed), 8, 1); lean_closure_set(x_516, 0, x_515); x_17 = x_516; goto block_36; @@ -40652,7 +40735,7 @@ lean_inc(x_6); lean_inc(x_8); lean_inc_ref(x_3); lean_inc(x_520); -x_524 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___boxed), 20, 11); +x_524 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___boxed), 20, 11); lean_closure_set(x_524, 0, x_519); lean_closure_set(x_524, 1, x_517); lean_closure_set(x_524, 2, x_520); @@ -40674,7 +40757,7 @@ if (lean_is_scalar(x_510)) { lean_ctor_set(x_526, 0, x_499); lean_ctor_set(x_526, 1, x_525); lean_ctor_set(x_526, 2, x_501); -x_527 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__9___boxed), 16, 9); +x_527 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__9___boxed), 16, 9); lean_closure_set(x_527, 0, x_518); lean_closure_set(x_527, 1, x_511); lean_closure_set(x_527, 2, x_524); @@ -40764,7 +40847,7 @@ lean_ctor_set(x_542, 1, x_541); lean_ctor_set(x_9, 1, x_542); x_543 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_543, 0, x_9); -x_544 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_544 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_544, 0, x_543); x_17 = x_544; goto block_36; @@ -40831,7 +40914,7 @@ lean_ctor_set(x_556, 1, x_555); lean_ctor_set(x_9, 1, x_556); x_557 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_557, 0, x_9); -x_558 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_558 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_558, 0, x_557); x_17 = x_558; goto block_36; @@ -40899,7 +40982,7 @@ lean_ctor_set(x_569, 1, x_568); lean_ctor_set(x_9, 1, x_569); x_570 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_570, 0, x_9); -x_571 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_571 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_571, 0, x_570); x_17 = x_571; goto block_36; @@ -40969,7 +41052,7 @@ lean_ctor_set(x_582, 1, x_581); lean_ctor_set(x_9, 1, x_582); x_583 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_583, 0, x_9); -x_584 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_584 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_584, 0, x_583); x_17 = x_584; goto block_36; @@ -41041,7 +41124,7 @@ lean_ctor_set(x_595, 1, x_594); lean_ctor_set(x_9, 1, x_595); x_596 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_596, 0, x_9); -x_597 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_597 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_597, 0, x_596); x_17 = x_597; goto block_36; @@ -41090,8 +41173,8 @@ lean_dec_ref(x_546); lean_dec(x_536); lean_dec_ref(x_535); lean_dec(x_533); -x_603 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__1; -x_604 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___boxed), 8, 1); +x_603 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__1; +x_604 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___boxed), 8, 1); lean_closure_set(x_604, 0, x_603); x_17 = x_604; goto block_36; @@ -41117,7 +41200,7 @@ lean_inc(x_6); lean_inc(x_8); lean_inc_ref(x_3); lean_inc(x_608); -x_612 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___boxed), 20, 11); +x_612 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___boxed), 20, 11); lean_closure_set(x_612, 0, x_607); lean_closure_set(x_612, 1, x_605); lean_closure_set(x_612, 2, x_608); @@ -41139,7 +41222,7 @@ if (lean_is_scalar(x_598)) { lean_ctor_set(x_614, 0, x_586); lean_ctor_set(x_614, 1, x_613); lean_ctor_set(x_614, 2, x_588); -x_615 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__9___boxed), 16, 9); +x_615 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__9___boxed), 16, 9); lean_closure_set(x_615, 0, x_606); lean_closure_set(x_615, 1, x_599); lean_closure_set(x_615, 2, x_612); @@ -41245,7 +41328,7 @@ lean_ctor_set(x_633, 0, x_616); lean_ctor_set(x_633, 1, x_632); x_634 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_634, 0, x_633); -x_635 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_635 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_635, 0, x_634); x_17 = x_635; goto block_36; @@ -41318,7 +41401,7 @@ lean_ctor_set(x_648, 0, x_616); lean_ctor_set(x_648, 1, x_647); x_649 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_649, 0, x_648); -x_650 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_650 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_650, 0, x_649); x_17 = x_650; goto block_36; @@ -41392,7 +41475,7 @@ lean_ctor_set(x_662, 0, x_616); lean_ctor_set(x_662, 1, x_661); x_663 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_663, 0, x_662); -x_664 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_664 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_664, 0, x_663); x_17 = x_664; goto block_36; @@ -41468,7 +41551,7 @@ lean_ctor_set(x_676, 0, x_616); lean_ctor_set(x_676, 1, x_675); x_677 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_677, 0, x_676); -x_678 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_678 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_678, 0, x_677); x_17 = x_678; goto block_36; @@ -41546,7 +41629,7 @@ lean_ctor_set(x_690, 0, x_616); lean_ctor_set(x_690, 1, x_689); x_691 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_691, 0, x_690); -x_692 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__0___boxed), 8, 1); +x_692 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__0___boxed), 8, 1); lean_closure_set(x_692, 0, x_691); x_17 = x_692; goto block_36; @@ -41595,8 +41678,8 @@ lean_dec_ref(x_637); lean_dec(x_626); lean_dec_ref(x_625); lean_dec(x_623); -x_698 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__1; -x_699 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___boxed), 8, 1); +x_698 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__1; +x_699 = lean_alloc_closure((void*)(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___boxed), 8, 1); lean_closure_set(x_699, 0, x_698); x_17 = x_699; goto block_36; @@ -41622,7 +41705,7 @@ lean_inc(x_6); lean_inc(x_8); lean_inc_ref(x_3); lean_inc(x_703); -x_707 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___boxed), 20, 11); +x_707 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___boxed), 20, 11); lean_closure_set(x_707, 0, x_702); lean_closure_set(x_707, 1, x_700); lean_closure_set(x_707, 2, x_703); @@ -41644,7 +41727,7 @@ if (lean_is_scalar(x_693)) { lean_ctor_set(x_709, 0, x_680); lean_ctor_set(x_709, 1, x_708); lean_ctor_set(x_709, 2, x_682); -x_710 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__9___boxed), 16, 9); +x_710 = lean_alloc_closure((void*)(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__9___boxed), 16, 9); lean_closure_set(x_710, 0, x_701); lean_closure_set(x_710, 1, x_694); lean_closure_set(x_710, 2, x_707); @@ -41711,7 +41794,7 @@ lean_dec_ref(x_20); x_23 = lean_unsigned_to_nat(1u); x_24 = lean_nat_add(x_8, x_23); lean_dec(x_8); -x_25 = l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56_spec__64___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_24, x_22, x_10, x_11, x_12, x_13, x_14, x_15); +x_25 = l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57_spec__65___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_24, x_22, x_10, x_11, x_12, x_13, x_14, x_15); return x_25; } } @@ -41750,7 +41833,7 @@ lean_dec_ref(x_26); x_30 = lean_unsigned_to_nat(1u); x_31 = lean_nat_add(x_8, x_30); lean_dec(x_8); -x_32 = l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56_spec__64___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_31, x_29, x_10, x_11, x_12, x_13, x_14, x_15); +x_32 = l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57_spec__65___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_31, x_29, x_10, x_11, x_12, x_13, x_14, x_15); return x_32; } } @@ -41787,71 +41870,71 @@ return x_35; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_10 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_4); lean_dec_ref(x_3); lean_dec(x_2); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; -x_11 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_10 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_4); lean_dec_ref(x_3); lean_dec(x_2); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; -x_11 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_10 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_4); lean_dec_ref(x_3); lean_dec(x_2); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; -x_9 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_9 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_3); return x_9; } } -static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__2() { +static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__2() { _start: { lean_object* x_1; @@ -41859,16 +41942,16 @@ x_1 = lean_mk_string_unchecked("failed to transform matcher, type error when con return x_1; } } -static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__3() { +static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__2; +x_1 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__2; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__4() { +static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__4() { _start: { lean_object* x_1; @@ -41876,16 +41959,16 @@ x_1 = lean_mk_string_unchecked("\nfailed with", 12, 12); return x_1; } } -static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__5() { +static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__4; +x_1 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__4; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__10(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__10(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; @@ -41896,7 +41979,7 @@ lean_ctor_set(x_4, 1, x_3); return x_4; } } -static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__6() { +static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -41908,7 +41991,7 @@ lean_ctor_set(x_3, 1, x_1); return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__49___lam__0(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__50___lam__0(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { if (x_1 == 0) @@ -41925,19 +42008,19 @@ return x_12; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__49___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__50___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { uint8_t x_11; lean_object* x_12; x_11 = lean_unbox(x_1); -x_12 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__49___lam__0(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_12 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__50___lam__0(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); return x_12; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__49(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__50(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_12; lean_object* x_13; uint8_t x_18; @@ -42019,7 +42102,7 @@ x_38 = lean_ctor_get(x_35, 0); lean_inc(x_38); lean_dec_ref(x_35); x_39 = lean_array_uget(x_1, x_3); -x_40 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__49___lam__0___boxed), 10, 2); +x_40 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__50___lam__0___boxed), 10, 2); lean_closure_set(x_40, 0, x_38); lean_closure_set(x_40, 1, x_39); lean_inc(x_10); @@ -42100,7 +42183,7 @@ x_52 = lean_ctor_get(x_48, 0); lean_inc(x_52); lean_dec_ref(x_48); x_53 = lean_array_uget(x_1, x_3); -x_54 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__49___lam__0___boxed), 10, 2); +x_54 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__50___lam__0___boxed), 10, 2); lean_closure_set(x_54, 0, x_52); lean_closure_set(x_54, 1, x_53); lean_inc(x_10); @@ -42234,7 +42317,7 @@ x_77 = lean_ctor_get(x_72, 0); lean_inc(x_77); lean_dec_ref(x_72); x_78 = lean_array_uget(x_1, x_3); -x_79 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__49___lam__0___boxed), 10, 2); +x_79 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__50___lam__0___boxed), 10, 2); lean_closure_set(x_79, 0, x_77); lean_closure_set(x_79, 1, x_78); lean_inc(x_10); @@ -42391,7 +42474,7 @@ x_107 = lean_ctor_get(x_101, 0); lean_inc(x_107); lean_dec_ref(x_101); x_108 = lean_array_uget(x_1, x_3); -x_109 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__49___lam__0___boxed), 10, 2); +x_109 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__50___lam__0___boxed), 10, 2); lean_closure_set(x_109, 0, x_107); lean_closure_set(x_109, 1, x_108); lean_inc(x_10); @@ -42467,29 +42550,29 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_10 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_4); lean_dec_ref(x_3); lean_dec(x_2); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__11(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_10 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__11(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_4); lean_dec_ref(x_3); lean_dec(x_2); return x_10; } } -static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__7() { +static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__7() { _start: { lean_object* x_1; @@ -42497,16 +42580,16 @@ x_1 = lean_mk_string_unchecked("failed to transform matcher, type error when con return x_1; } } -static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__8() { +static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__7; +x_1 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__7; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__46(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__47(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { uint8_t x_12; @@ -42582,7 +42665,7 @@ return x_25; } } } -static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___boxed__const__1() { +static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___boxed__const__1() { _start: { size_t x_1; lean_object* x_2; @@ -42591,7 +42674,7 @@ x_2 = lean_box_usize(x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___closed__0() { +static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___closed__0() { _start: { lean_object* x_1; lean_object* x_2; @@ -42600,7 +42683,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_10; @@ -42608,18 +42691,18 @@ x_10 = l_Lean_Meta_isProof(x_1, x_5, x_6, x_7, x_8); return x_10; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_10 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_4); lean_dec_ref(x_3); lean_dec(x_2); return x_10; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_11; @@ -42627,18 +42710,18 @@ x_11 = l_Lean_Meta_mkEqHEq(x_1, x_2, x_6, x_7, x_8, x_9); return x_11; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; -x_11 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___lam__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___lam__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); return x_11; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___lam__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___lam__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_11; @@ -42646,11 +42729,11 @@ x_11 = l_Lean_mkArrow(x_1, x_2, x_8, x_9); return x_11; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___lam__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___lam__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; -x_11 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___lam__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___lam__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_7); lean_dec_ref(x_6); lean_dec(x_5); @@ -42659,7 +42742,7 @@ lean_dec(x_3); return x_11; } } -static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___closed__0() { +static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___closed__0() { _start: { lean_object* x_1; lean_object* x_2; @@ -42669,7 +42752,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47(uint8_t x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48(uint8_t x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_13; lean_object* x_14; uint8_t x_19; @@ -42886,7 +42969,7 @@ lean_dec(x_27); lean_dec(x_25); x_71 = lean_array_uget(x_2, x_4); lean_inc(x_71); -x_72 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___lam__0___boxed), 9, 1); +x_72 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___lam__0___boxed), 9, 1); lean_closure_set(x_72, 0, x_71); lean_inc(x_11); lean_inc_ref(x_10); @@ -42908,7 +42991,7 @@ lean_object* x_76; lean_object* x_77; lean_object* x_78; x_76 = lean_array_fget(x_46, x_47); lean_dec(x_47); lean_dec_ref(x_46); -x_77 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___lam__1___boxed), 10, 2); +x_77 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___lam__1___boxed), 10, 2); lean_closure_set(x_77, 0, x_76); lean_closure_set(x_77, 1, x_71); lean_inc(x_11); @@ -42924,7 +43007,7 @@ x_79 = lean_ctor_get(x_78, 0); lean_inc(x_79); lean_dec_ref(x_78); lean_inc(x_79); -x_80 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___lam__2___boxed), 10, 2); +x_80 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___lam__2___boxed), 10, 2); lean_closure_set(x_80, 0, x_79); lean_closure_set(x_80, 1, x_31); lean_inc(x_11); @@ -42945,7 +43028,7 @@ x_84 = lean_box(x_83); x_85 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_85, 0, x_84); x_86 = lean_array_push(x_28, x_85); -x_87 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___closed__0; +x_87 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___closed__0; x_88 = lean_array_push(x_30, x_87); x_89 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_89, 0, x_88); @@ -43158,7 +43241,7 @@ lean_dec(x_27); lean_dec(x_25); x_120 = lean_array_uget(x_2, x_4); lean_inc(x_120); -x_121 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___lam__0___boxed), 9, 1); +x_121 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___lam__0___boxed), 9, 1); lean_closure_set(x_121, 0, x_120); lean_inc(x_11); lean_inc_ref(x_10); @@ -43180,7 +43263,7 @@ lean_object* x_125; lean_object* x_126; lean_object* x_127; x_125 = lean_array_fget(x_46, x_47); lean_dec(x_47); lean_dec_ref(x_46); -x_126 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___lam__1___boxed), 10, 2); +x_126 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___lam__1___boxed), 10, 2); lean_closure_set(x_126, 0, x_125); lean_closure_set(x_126, 1, x_120); lean_inc(x_11); @@ -43196,7 +43279,7 @@ x_128 = lean_ctor_get(x_127, 0); lean_inc(x_128); lean_dec_ref(x_127); lean_inc(x_128); -x_129 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___lam__2___boxed), 10, 2); +x_129 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___lam__2___boxed), 10, 2); lean_closure_set(x_129, 0, x_128); lean_closure_set(x_129, 1, x_31); lean_inc(x_11); @@ -43217,7 +43300,7 @@ x_133 = lean_box(x_132); x_134 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_134, 0, x_133); x_135 = lean_array_push(x_28, x_134); -x_136 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___closed__0; +x_136 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___closed__0; x_137 = lean_array_push(x_30, x_136); x_138 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_138, 0, x_137); @@ -43507,7 +43590,7 @@ lean_dec(x_27); lean_dec(x_25); x_182 = lean_array_uget(x_2, x_4); lean_inc(x_182); -x_183 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___lam__0___boxed), 9, 1); +x_183 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___lam__0___boxed), 9, 1); lean_closure_set(x_183, 0, x_182); lean_inc(x_11); lean_inc_ref(x_10); @@ -43529,7 +43612,7 @@ lean_object* x_187; lean_object* x_188; lean_object* x_189; x_187 = lean_array_fget(x_158, x_159); lean_dec(x_159); lean_dec_ref(x_158); -x_188 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___lam__1___boxed), 10, 2); +x_188 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___lam__1___boxed), 10, 2); lean_closure_set(x_188, 0, x_187); lean_closure_set(x_188, 1, x_182); lean_inc(x_11); @@ -43545,7 +43628,7 @@ x_190 = lean_ctor_get(x_189, 0); lean_inc(x_190); lean_dec_ref(x_189); lean_inc(x_190); -x_191 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___lam__2___boxed), 10, 2); +x_191 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___lam__2___boxed), 10, 2); lean_closure_set(x_191, 0, x_190); lean_closure_set(x_191, 1, x_31); lean_inc(x_11); @@ -43566,7 +43649,7 @@ x_195 = lean_box(x_194); x_196 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_196, 0, x_195); x_197 = lean_array_push(x_28, x_196); -x_198 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___closed__0; +x_198 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___closed__0; x_199 = lean_array_push(x_30, x_198); x_200 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_200, 0, x_199); @@ -43766,14 +43849,14 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { uint8_t x_14; uint8_t x_15; uint8_t x_16; lean_object* x_17; x_14 = lean_unbox(x_3); x_15 = lean_unbox(x_4); x_16 = lean_unbox(x_5); -x_17 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__14(x_1, x_2, x_14, x_15, x_16, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_17 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__14(x_1, x_2, x_14, x_15, x_16, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_12); lean_dec_ref(x_11); lean_dec(x_10); @@ -43785,18 +43868,18 @@ lean_dec_ref(x_1); return x_17; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__12(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_10 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__12(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_4); lean_dec_ref(x_3); lean_dec(x_2); return x_10; } } -static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___closed__1() { +static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___closed__1() { _start: { lean_object* x_1; @@ -43804,16 +43887,16 @@ x_1 = lean_mk_string_unchecked("unexpected matcher application, motive must be l return x_1; } } -static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___closed__2() { +static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___closed__1; +x_1 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___closed__3() { +static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___closed__3() { _start: { lean_object* x_1; @@ -43821,16 +43904,16 @@ x_1 = lean_mk_string_unchecked(" arguments", 10, 10); return x_1; } } -static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___closed__4() { +static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___closed__3; +x_1 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___closed__3; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, size_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, size_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { lean_object* x_16; lean_object* x_152; lean_object* x_153; uint8_t x_154; @@ -43845,7 +43928,7 @@ lean_dec_ref(x_7); lean_dec_ref(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_155 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___closed__2; +x_155 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___closed__2; x_156 = l_Nat_reprFast(x_153); x_157 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_157, 0, x_156); @@ -43853,7 +43936,7 @@ x_158 = l_Lean_MessageData_ofFormat(x_157); x_159 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_159, 0, x_155); lean_ctor_set(x_159, 1, x_158); -x_160 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___closed__4; +x_160 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___closed__4; x_161 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_161, 0, x_159); lean_ctor_set(x_161, 1, x_160); @@ -43901,7 +43984,7 @@ x_19 = lean_ctor_get(x_2, 4); lean_inc_ref(x_19); lean_dec_ref(x_2); x_20 = lean_unsigned_to_nat(0u); -x_21 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___closed__0; +x_21 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___closed__0; x_22 = lean_array_get_size(x_3); x_23 = l_Array_toSubarray___redArg(x_3, x_20, x_22); x_24 = lean_array_get_size(x_19); @@ -43924,7 +44007,7 @@ lean_inc_ref(x_13); lean_inc(x_12); lean_inc_ref(x_11); lean_inc_ref(x_9); -x_31 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47(x_4, x_7, x_30, x_5, x_29, x_9, x_10, x_11, x_12, x_13, x_14); +x_31 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48(x_4, x_7, x_30, x_5, x_29, x_9, x_10, x_11, x_12, x_13, x_14); if (lean_obj_tag(x_31) == 0) { lean_object* x_32; lean_object* x_33; uint8_t x_34; @@ -43960,7 +44043,7 @@ x_46 = lean_box(x_43); x_47 = lean_box(x_44); x_48 = lean_box(x_45); lean_inc(x_42); -x_49 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__14___boxed), 13, 5); +x_49 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__14___boxed), 13, 5); lean_closure_set(x_49, 0, x_7); lean_closure_set(x_49, 1, x_42); lean_closure_set(x_49, 2, x_46); @@ -43978,7 +44061,7 @@ lean_object* x_51; lean_object* x_52; lean_object* x_53; x_51 = lean_ctor_get(x_50, 0); lean_inc(x_51); lean_dec_ref(x_50); -x_52 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__12___boxed), 9, 1); +x_52 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__12___boxed), 9, 1); lean_closure_set(x_52, 0, x_42); x_53 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_52, x_9, x_10, x_11, x_12, x_13, x_14); lean_dec(x_10); @@ -44086,7 +44169,7 @@ x_70 = lean_box(x_67); x_71 = lean_box(x_68); x_72 = lean_box(x_69); lean_inc(x_66); -x_73 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__14___boxed), 13, 5); +x_73 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__14___boxed), 13, 5); lean_closure_set(x_73, 0, x_7); lean_closure_set(x_73, 1, x_66); lean_closure_set(x_73, 2, x_70); @@ -44104,7 +44187,7 @@ lean_object* x_75; lean_object* x_76; lean_object* x_77; x_75 = lean_ctor_get(x_74, 0); lean_inc(x_75); lean_dec_ref(x_74); -x_76 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__12___boxed), 9, 1); +x_76 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__12___boxed), 9, 1); lean_closure_set(x_76, 0, x_66); x_77 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_76, x_9, x_10, x_11, x_12, x_13, x_14); lean_dec(x_10); @@ -44220,7 +44303,7 @@ x_96 = lean_box(x_93); x_97 = lean_box(x_94); x_98 = lean_box(x_95); lean_inc(x_91); -x_99 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__14___boxed), 13, 5); +x_99 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__14___boxed), 13, 5); lean_closure_set(x_99, 0, x_7); lean_closure_set(x_99, 1, x_91); lean_closure_set(x_99, 2, x_96); @@ -44238,7 +44321,7 @@ lean_object* x_101; lean_object* x_102; lean_object* x_103; x_101 = lean_ctor_get(x_100, 0); lean_inc(x_101); lean_dec_ref(x_100); -x_102 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__12___boxed), 9, 1); +x_102 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__12___boxed), 9, 1); lean_closure_set(x_102, 0, x_91); x_103 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_102, x_9, x_10, x_11, x_12, x_13, x_14); lean_dec(x_10); @@ -44370,7 +44453,7 @@ x_125 = lean_box(x_122); x_126 = lean_box(x_123); x_127 = lean_box(x_124); lean_inc(x_120); -x_128 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__14___boxed), 13, 5); +x_128 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__14___boxed), 13, 5); lean_closure_set(x_128, 0, x_7); lean_closure_set(x_128, 1, x_120); lean_closure_set(x_128, 2, x_125); @@ -44388,7 +44471,7 @@ lean_object* x_130; lean_object* x_131; lean_object* x_132; x_130 = lean_ctor_get(x_129, 0); lean_inc(x_130); lean_dec_ref(x_129); -x_131 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__12___boxed), 9, 1); +x_131 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__12___boxed), 9, 1); lean_closure_set(x_131, 0, x_120); x_132 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_131, x_9, x_10, x_11, x_12, x_13, x_14); lean_dec(x_10); @@ -44546,23 +44629,23 @@ return x_150; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { uint8_t x_16; size_t x_17; lean_object* x_18; x_16 = lean_unbox(x_4); x_17 = lean_unbox_usize(x_5); lean_dec(x_5); -x_18 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13(x_1, x_2, x_3, x_16, x_17, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +x_18 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13(x_1, x_2, x_3, x_16, x_17, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); lean_dec_ref(x_6); return x_18; } } -LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__57(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__58(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { lean_object* x_9; lean_object* x_10; -x_9 = lean_alloc_closure((void*)(l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__57___lam__0___boxed), 8, 0); +x_9 = lean_alloc_closure((void*)(l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__58___lam__0___boxed), 8, 0); x_10 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_9, x_2, x_3, x_4, x_5, x_6, x_7); if (lean_obj_tag(x_10) == 0) { @@ -44610,7 +44693,7 @@ return x_19; } } } -static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__9() { +static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__9() { _start: { lean_object* x_1; @@ -44618,16 +44701,16 @@ x_1 = lean_mk_string_unchecked("matcher ", 8, 8); return x_1; } } -static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__10() { +static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__9; +x_1 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__9; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__11() { +static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__11() { _start: { lean_object* x_1; @@ -44635,20 +44718,20 @@ x_1 = lean_mk_string_unchecked(" has no MatchInfo found", 23, 23); return x_1; } } -static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__12() { +static lean_object* _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__11; +x_1 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__11; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37(lean_object* x_1, uint8_t x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38(lean_object* x_1, uint8_t x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { lean_object* x_15; lean_object* x_16; -x_15 = lean_alloc_closure((void*)(l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__57___lam__0___boxed), 8, 0); +x_15 = lean_alloc_closure((void*)(l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__58___lam__0___boxed), 8, 0); lean_inc(x_13); lean_inc_ref(x_12); lean_inc(x_11); @@ -44657,7 +44740,7 @@ lean_inc_ref(x_8); x_16 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_15, x_8, x_9, x_10, x_11, x_12, x_13); if (lean_obj_tag(x_16) == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; uint8_t x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; uint8_t x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; uint8_t x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; uint8_t x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; uint8_t x_322; lean_object* x_323; lean_object* x_324; size_t x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; uint8_t x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; uint8_t x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; uint8_t x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; uint8_t x_318; lean_object* x_319; size_t x_320; lean_object* x_321; lean_object* x_322; uint8_t x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); lean_dec_ref(x_16); @@ -44673,7 +44756,7 @@ x_24 = lean_ctor_get(x_1, 6); x_25 = lean_ctor_get(x_1, 7); lean_inc_ref(x_25); lean_inc(x_19); -x_255 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__1___boxed), 9, 1); +x_255 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__1___boxed), 9, 1); lean_closure_set(x_255, 0, x_19); lean_inc(x_19); x_318 = l_Lean_isCasesOnRecursor(x_17, x_19); @@ -44686,7 +44769,7 @@ lean_inc(x_11); lean_inc_ref(x_10); lean_inc_ref(x_8); lean_inc(x_19); -x_491 = l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__57(x_19, x_8, x_9, x_10, x_11, x_12, x_13); +x_491 = l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__58(x_19, x_8, x_9, x_10, x_11, x_12, x_13); if (lean_obj_tag(x_491) == 0) { lean_object* x_492; @@ -44704,12 +44787,12 @@ lean_dec_ref(x_6); lean_dec_ref(x_5); lean_dec_ref(x_4); lean_dec_ref(x_1); -x_493 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__10; +x_493 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__10; x_494 = l_Lean_MessageData_ofName(x_19); x_495 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_495, 0, x_493); lean_ctor_set(x_495, 1, x_494); -x_496 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__12; +x_496 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__12; x_497 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_497, 0, x_495); lean_ctor_set(x_497, 1, x_496); @@ -44803,9 +44886,9 @@ goto block_490; { lean_object* x_45; lean_object* x_46; lean_object* x_47; x_45 = lean_array_get_size(x_24); -x_46 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__2___boxed), 10, 2); +x_46 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__2___boxed), 10, 2); lean_closure_set(x_46, 0, x_45); -lean_closure_set(x_46, 1, x_27); +lean_closure_set(x_46, 1, x_26); lean_inc(x_43); lean_inc_ref(x_42); lean_inc(x_41); @@ -44818,19 +44901,19 @@ lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean x_48 = lean_ctor_get(x_47, 0); lean_inc(x_48); lean_dec_ref(x_47); -lean_inc(x_34); +lean_inc(x_29); lean_inc_ref(x_24); -x_49 = l_Array_toSubarray___redArg(x_24, x_34, x_45); +x_49 = l_Array_toSubarray___redArg(x_24, x_29, x_45); x_50 = l_Lean_Meta_MatcherApp_altNumParams(x_1); x_51 = lean_array_get_size(x_50); -lean_inc(x_34); -x_52 = l_Array_toSubarray___redArg(x_50, x_34, x_51); +lean_inc(x_29); +x_52 = l_Array_toSubarray___redArg(x_50, x_29, x_51); x_53 = lean_array_get_size(x_48); -lean_inc(x_34); -x_54 = l_Array_toSubarray___redArg(x_48, x_34, x_53); +lean_inc(x_29); +x_54 = l_Array_toSubarray___redArg(x_48, x_29, x_53); x_55 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_55, 0, x_54); -lean_ctor_set(x_55, 1, x_31); +lean_ctor_set(x_55, 1, x_33); x_56 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_56, 0, x_52); lean_ctor_set(x_56, 1, x_55); @@ -44843,7 +44926,7 @@ lean_inc(x_41); lean_inc_ref(x_40); lean_inc(x_39); lean_inc_ref(x_38); -x_58 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg(x_45, x_32, x_6, x_26, x_36, x_34, x_57, x_38, x_39, x_40, x_41, x_42, x_43); +x_58 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg(x_45, x_34, x_6, x_27, x_32, x_29, x_57, x_38, x_39, x_40, x_41, x_42, x_43); if (lean_obj_tag(x_58) == 0) { lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; @@ -44876,14 +44959,14 @@ x_67 = lean_ctor_get(x_18, 4); lean_dec(x_67); x_68 = l_Array_append___redArg(x_37, x_66); lean_dec(x_66); -lean_ctor_set(x_18, 4, x_30); +lean_ctor_set(x_18, 4, x_35); x_69 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_69, 0, x_18); lean_ctor_set(x_69, 1, x_19); -lean_ctor_set(x_69, 2, x_29); -lean_ctor_set(x_69, 3, x_35); -lean_ctor_set(x_69, 4, x_28); -lean_ctor_set(x_69, 5, x_33); +lean_ctor_set(x_69, 2, x_31); +lean_ctor_set(x_69, 3, x_28); +lean_ctor_set(x_69, 4, x_36); +lean_ctor_set(x_69, 5, x_30); lean_ctor_set(x_69, 6, x_62); lean_ctor_set(x_69, 7, x_68); lean_ctor_set(x_63, 0, x_69); @@ -44911,15 +44994,15 @@ lean_ctor_set(x_77, 0, x_71); lean_ctor_set(x_77, 1, x_72); lean_ctor_set(x_77, 2, x_73); lean_ctor_set(x_77, 3, x_74); -lean_ctor_set(x_77, 4, x_30); +lean_ctor_set(x_77, 4, x_35); lean_ctor_set(x_77, 5, x_75); x_78 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_78, 0, x_77); lean_ctor_set(x_78, 1, x_19); -lean_ctor_set(x_78, 2, x_29); -lean_ctor_set(x_78, 3, x_35); -lean_ctor_set(x_78, 4, x_28); -lean_ctor_set(x_78, 5, x_33); +lean_ctor_set(x_78, 2, x_31); +lean_ctor_set(x_78, 3, x_28); +lean_ctor_set(x_78, 4, x_36); +lean_ctor_set(x_78, 5, x_30); lean_ctor_set(x_78, 6, x_62); lean_ctor_set(x_78, 7, x_76); lean_ctor_set(x_63, 0, x_78); @@ -44965,15 +45048,15 @@ lean_ctor_set(x_87, 0, x_80); lean_ctor_set(x_87, 1, x_81); lean_ctor_set(x_87, 2, x_82); lean_ctor_set(x_87, 3, x_83); -lean_ctor_set(x_87, 4, x_30); +lean_ctor_set(x_87, 4, x_35); lean_ctor_set(x_87, 5, x_84); x_88 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_88, 0, x_87); lean_ctor_set(x_88, 1, x_19); -lean_ctor_set(x_88, 2, x_29); -lean_ctor_set(x_88, 3, x_35); -lean_ctor_set(x_88, 4, x_28); -lean_ctor_set(x_88, 5, x_33); +lean_ctor_set(x_88, 2, x_31); +lean_ctor_set(x_88, 3, x_28); +lean_ctor_set(x_88, 4, x_36); +lean_ctor_set(x_88, 5, x_30); lean_ctor_set(x_88, 6, x_62); lean_ctor_set(x_88, 7, x_86); x_89 = lean_alloc_ctor(0, 1, 0); @@ -44986,10 +45069,10 @@ else uint8_t x_90; lean_dec(x_62); lean_dec_ref(x_37); +lean_dec_ref(x_36); lean_dec_ref(x_35); -lean_dec_ref(x_33); +lean_dec_ref(x_31); lean_dec_ref(x_30); -lean_dec_ref(x_29); lean_dec_ref(x_28); lean_dec(x_19); lean_dec_ref(x_18); @@ -45020,10 +45103,10 @@ lean_dec_ref(x_40); lean_dec(x_39); lean_dec_ref(x_38); lean_dec_ref(x_37); +lean_dec_ref(x_36); lean_dec_ref(x_35); -lean_dec_ref(x_33); +lean_dec_ref(x_31); lean_dec_ref(x_30); -lean_dec_ref(x_29); lean_dec_ref(x_28); lean_dec_ref(x_25); lean_dec(x_19); @@ -45056,13 +45139,13 @@ lean_dec_ref(x_40); lean_dec(x_39); lean_dec_ref(x_38); lean_dec_ref(x_37); -lean_dec(x_36); +lean_dec_ref(x_36); lean_dec_ref(x_35); -lean_dec(x_34); lean_dec_ref(x_33); +lean_dec(x_32); lean_dec_ref(x_31); lean_dec_ref(x_30); -lean_dec_ref(x_29); +lean_dec(x_29); lean_dec_ref(x_28); lean_dec_ref(x_25); lean_dec(x_19); @@ -45090,23 +45173,23 @@ return x_98; block_142: { lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; -lean_inc_ref(x_100); -x_117 = lean_array_to_list(x_100); +lean_inc_ref(x_110); +x_117 = lean_array_to_list(x_110); lean_inc(x_19); x_118 = l_Lean_mkConst(x_19, x_117); -x_119 = l_Lean_mkAppN(x_118, x_114); -lean_inc_ref(x_109); -x_120 = l_Lean_Expr_app___override(x_119, x_109); -x_121 = l_Lean_mkAppN(x_120, x_110); +x_119 = l_Lean_mkAppN(x_118, x_103); +lean_inc_ref(x_116); +x_120 = l_Lean_Expr_app___override(x_119, x_116); +x_121 = l_Lean_mkAppN(x_120, x_111); lean_inc_ref(x_121); -x_122 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__0___boxed), 9, 1); +x_122 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__0___boxed), 9, 1); lean_closure_set(x_122, 0, x_121); -lean_inc(x_112); -lean_inc_ref(x_113); -lean_inc(x_116); -lean_inc_ref(x_111); -lean_inc_ref(x_107); -x_123 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_122, x_107, x_102, x_111, x_116, x_113, x_112); +lean_inc(x_102); +lean_inc_ref(x_106); +lean_inc(x_113); +lean_inc_ref(x_101); +lean_inc_ref(x_100); +x_123 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_122, x_100, x_107, x_101, x_113, x_106, x_102); if (lean_obj_tag(x_123) == 0) { lean_object* x_124; uint8_t x_125; uint8_t x_126; @@ -45119,52 +45202,52 @@ lean_dec(x_124); if (x_126 == 0) { lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; -x_127 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__1; +x_127 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__1; lean_inc_ref(x_121); x_128 = l_Lean_indentExpr(x_121); x_129 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_129, 0, x_127); lean_ctor_set(x_129, 1, x_128); -lean_inc(x_112); -lean_inc_ref(x_113); -lean_inc(x_116); -lean_inc_ref(x_111); -lean_inc_ref(x_107); -x_130 = l_Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52(x_129, x_107, x_102, x_111, x_116, x_113, x_112); +lean_inc(x_102); +lean_inc_ref(x_106); +lean_inc(x_113); +lean_inc_ref(x_101); +lean_inc_ref(x_100); +x_130 = l_Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53(x_129, x_100, x_107, x_101, x_113, x_106, x_102); if (lean_obj_tag(x_130) == 0) { lean_object* x_131; lean_object* x_132; lean_dec_ref(x_130); lean_inc_ref(x_121); -x_131 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__3___boxed), 9, 1); +x_131 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__3___boxed), 9, 1); lean_closure_set(x_131, 0, x_121); -lean_inc(x_112); -lean_inc_ref(x_113); -lean_inc(x_116); -lean_inc_ref(x_111); -lean_inc_ref(x_107); -x_132 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_131, x_107, x_102, x_111, x_116, x_113, x_112); +lean_inc(x_102); +lean_inc_ref(x_106); +lean_inc(x_113); +lean_inc_ref(x_101); +lean_inc_ref(x_100); +x_132 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_131, x_100, x_107, x_101, x_113, x_106, x_102); if (lean_obj_tag(x_132) == 0) { lean_dec_ref(x_132); -x_26 = x_125; -x_27 = x_121; -x_28 = x_109; -x_29 = x_100; -x_30 = x_101; -x_31 = x_104; -x_32 = x_105; -x_33 = x_110; -x_34 = x_106; +x_26 = x_121; +x_27 = x_125; +x_28 = x_103; +x_29 = x_109; +x_30 = x_111; +x_31 = x_110; +x_32 = x_112; +x_33 = x_104; +x_34 = x_105; x_35 = x_114; -x_36 = x_108; -x_37 = x_103; -x_38 = x_107; -x_39 = x_102; -x_40 = x_111; -x_41 = x_116; -x_42 = x_113; -x_43 = x_112; +x_36 = x_116; +x_37 = x_115; +x_38 = x_100; +x_39 = x_107; +x_40 = x_101; +x_41 = x_113; +x_42 = x_106; +x_43 = x_102; x_44 = lean_box(0); goto block_99; } @@ -45172,16 +45255,16 @@ else { uint8_t x_133; lean_dec_ref(x_121); -lean_dec(x_116); +lean_dec_ref(x_116); +lean_dec_ref(x_115); lean_dec_ref(x_114); -lean_dec_ref(x_113); +lean_dec(x_113); lean_dec(x_112); lean_dec_ref(x_111); lean_dec_ref(x_110); -lean_dec_ref(x_109); -lean_dec(x_108); -lean_dec_ref(x_107); -lean_dec(x_106); +lean_dec(x_109); +lean_dec(x_107); +lean_dec_ref(x_106); lean_dec_ref(x_104); lean_dec_ref(x_103); lean_dec(x_102); @@ -45214,16 +45297,16 @@ else { uint8_t x_136; lean_dec_ref(x_121); -lean_dec(x_116); +lean_dec_ref(x_116); +lean_dec_ref(x_115); lean_dec_ref(x_114); -lean_dec_ref(x_113); +lean_dec(x_113); lean_dec(x_112); lean_dec_ref(x_111); lean_dec_ref(x_110); -lean_dec_ref(x_109); -lean_dec(x_108); -lean_dec_ref(x_107); -lean_dec(x_106); +lean_dec(x_109); +lean_dec(x_107); +lean_dec_ref(x_106); lean_dec_ref(x_104); lean_dec_ref(x_103); lean_dec(x_102); @@ -45254,24 +45337,24 @@ return x_138; } else { -x_26 = x_125; -x_27 = x_121; -x_28 = x_109; -x_29 = x_100; -x_30 = x_101; -x_31 = x_104; -x_32 = x_105; -x_33 = x_110; -x_34 = x_106; +x_26 = x_121; +x_27 = x_125; +x_28 = x_103; +x_29 = x_109; +x_30 = x_111; +x_31 = x_110; +x_32 = x_112; +x_33 = x_104; +x_34 = x_105; x_35 = x_114; -x_36 = x_108; -x_37 = x_103; -x_38 = x_107; -x_39 = x_102; -x_40 = x_111; -x_41 = x_116; -x_42 = x_113; -x_43 = x_112; +x_36 = x_116; +x_37 = x_115; +x_38 = x_100; +x_39 = x_107; +x_40 = x_101; +x_41 = x_113; +x_42 = x_106; +x_43 = x_102; x_44 = lean_box(0); goto block_99; } @@ -45280,16 +45363,16 @@ else { uint8_t x_139; lean_dec_ref(x_121); -lean_dec(x_116); +lean_dec_ref(x_116); +lean_dec_ref(x_115); lean_dec_ref(x_114); -lean_dec_ref(x_113); +lean_dec(x_113); lean_dec(x_112); lean_dec_ref(x_111); lean_dec_ref(x_110); -lean_dec_ref(x_109); -lean_dec(x_108); -lean_dec_ref(x_107); -lean_dec(x_106); +lean_dec(x_109); +lean_dec(x_107); +lean_dec_ref(x_106); lean_dec_ref(x_104); lean_dec_ref(x_103); lean_dec(x_102); @@ -45344,37 +45427,37 @@ lean_inc(x_171); x_172 = lean_ctor_get(x_18, 5); lean_inc_ref(x_172); lean_dec_ref(x_18); -x_173 = !lean_is_exclusive(x_154); +x_173 = !lean_is_exclusive(x_149); if (x_173 == 0) { lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; -x_174 = lean_ctor_get(x_154, 2); -x_175 = lean_ctor_get(x_154, 5); +x_174 = lean_ctor_get(x_149, 2); +x_175 = lean_ctor_get(x_149, 5); lean_dec(x_175); -x_176 = lean_ctor_get(x_154, 4); +x_176 = lean_ctor_get(x_149, 4); lean_dec(x_176); -x_177 = lean_ctor_get(x_154, 3); +x_177 = lean_ctor_get(x_149, 3); lean_dec(x_177); -x_178 = lean_ctor_get(x_154, 1); +x_178 = lean_ctor_get(x_149, 1); lean_dec(x_178); -x_179 = lean_ctor_get(x_154, 0); +x_179 = lean_ctor_get(x_149, 0); lean_dec(x_179); -lean_inc(x_150); -lean_inc(x_148); -x_180 = l_Array_toSubarray___redArg(x_24, x_148, x_150); +lean_inc(x_153); +lean_inc(x_151); +x_180 = l_Array_toSubarray___redArg(x_24, x_151, x_153); x_181 = lean_array_get_size(x_170); -lean_inc(x_148); +lean_inc(x_151); lean_inc_ref(x_170); -x_182 = l_Array_toSubarray___redArg(x_170, x_148, x_181); +x_182 = l_Array_toSubarray___redArg(x_170, x_151, x_181); x_183 = lean_array_get_size(x_174); -lean_inc(x_148); -x_184 = l_Array_toSubarray___redArg(x_174, x_148, x_183); -x_185 = lean_array_get_size(x_149); -lean_inc(x_148); -x_186 = l_Array_toSubarray___redArg(x_149, x_148, x_185); +lean_inc(x_151); +x_184 = l_Array_toSubarray___redArg(x_174, x_151, x_183); +x_185 = lean_array_get_size(x_150); +lean_inc(x_151); +x_186 = l_Array_toSubarray___redArg(x_150, x_151, x_185); x_187 = lean_array_get_size(x_167); -lean_inc(x_148); -x_188 = l_Array_toSubarray___redArg(x_167, x_148, x_187); +lean_inc(x_151); +x_188 = l_Array_toSubarray___redArg(x_167, x_151, x_187); x_189 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_189, 0, x_188); lean_ctor_set(x_189, 1, x_146); @@ -45396,10 +45479,10 @@ lean_inc(x_162); lean_inc_ref(x_161); lean_inc(x_160); lean_inc_ref(x_159); -lean_inc(x_148); -x_194 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg(x_150, x_148, x_6, x_147, x_2, x_152, x_155, x_148, x_193, x_159, x_160, x_161, x_162, x_163, x_164); -lean_dec(x_148); -lean_dec(x_150); +lean_inc(x_151); +x_194 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg(x_153, x_151, x_6, x_147, x_2, x_154, x_148, x_151, x_193, x_159, x_160, x_161, x_162, x_163, x_164); +lean_dec(x_151); +lean_dec(x_153); if (lean_obj_tag(x_194) == 0) { lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; @@ -45432,19 +45515,19 @@ lean_object* x_203; lean_object* x_204; lean_object* x_205; x_203 = lean_ctor_get(x_201, 0); x_204 = l_Array_append___redArg(x_158, x_203); lean_dec(x_203); -lean_ctor_set(x_154, 5, x_172); -lean_ctor_set(x_154, 4, x_144); -lean_ctor_set(x_154, 3, x_171); -lean_ctor_set(x_154, 2, x_170); -lean_ctor_set(x_154, 1, x_169); -lean_ctor_set(x_154, 0, x_168); +lean_ctor_set(x_149, 5, x_172); +lean_ctor_set(x_149, 4, x_156); +lean_ctor_set(x_149, 3, x_171); +lean_ctor_set(x_149, 2, x_170); +lean_ctor_set(x_149, 1, x_169); +lean_ctor_set(x_149, 0, x_168); x_205 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_205, 0, x_154); -lean_ctor_set(x_205, 1, x_151); -lean_ctor_set(x_205, 2, x_145); -lean_ctor_set(x_205, 3, x_157); -lean_ctor_set(x_205, 4, x_153); -lean_ctor_set(x_205, 5, x_156); +lean_ctor_set(x_205, 0, x_149); +lean_ctor_set(x_205, 1, x_144); +lean_ctor_set(x_205, 2, x_155); +lean_ctor_set(x_205, 3, x_145); +lean_ctor_set(x_205, 4, x_157); +lean_ctor_set(x_205, 5, x_152); lean_ctor_set(x_205, 6, x_200); lean_ctor_set(x_205, 7, x_204); lean_ctor_set(x_201, 0, x_205); @@ -45458,19 +45541,19 @@ lean_inc(x_206); lean_dec(x_201); x_207 = l_Array_append___redArg(x_158, x_206); lean_dec(x_206); -lean_ctor_set(x_154, 5, x_172); -lean_ctor_set(x_154, 4, x_144); -lean_ctor_set(x_154, 3, x_171); -lean_ctor_set(x_154, 2, x_170); -lean_ctor_set(x_154, 1, x_169); -lean_ctor_set(x_154, 0, x_168); +lean_ctor_set(x_149, 5, x_172); +lean_ctor_set(x_149, 4, x_156); +lean_ctor_set(x_149, 3, x_171); +lean_ctor_set(x_149, 2, x_170); +lean_ctor_set(x_149, 1, x_169); +lean_ctor_set(x_149, 0, x_168); x_208 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_208, 0, x_154); -lean_ctor_set(x_208, 1, x_151); -lean_ctor_set(x_208, 2, x_145); -lean_ctor_set(x_208, 3, x_157); -lean_ctor_set(x_208, 4, x_153); -lean_ctor_set(x_208, 5, x_156); +lean_ctor_set(x_208, 0, x_149); +lean_ctor_set(x_208, 1, x_144); +lean_ctor_set(x_208, 2, x_155); +lean_ctor_set(x_208, 3, x_145); +lean_ctor_set(x_208, 4, x_157); +lean_ctor_set(x_208, 5, x_152); lean_ctor_set(x_208, 6, x_200); lean_ctor_set(x_208, 7, x_207); x_209 = lean_alloc_ctor(0, 1, 0); @@ -45482,7 +45565,7 @@ else { uint8_t x_210; lean_dec(x_200); -lean_free_object(x_154); +lean_free_object(x_149); lean_dec_ref(x_172); lean_dec(x_171); lean_dec_ref(x_170); @@ -45491,10 +45574,10 @@ lean_dec(x_168); lean_dec_ref(x_158); lean_dec_ref(x_157); lean_dec_ref(x_156); -lean_dec_ref(x_153); -lean_dec(x_151); +lean_dec_ref(x_155); +lean_dec_ref(x_152); lean_dec_ref(x_145); -lean_dec_ref(x_144); +lean_dec(x_144); x_210 = !lean_is_exclusive(x_201); if (x_210 == 0) { @@ -45515,7 +45598,7 @@ return x_212; else { uint8_t x_213; -lean_free_object(x_154); +lean_free_object(x_149); lean_dec_ref(x_172); lean_dec(x_171); lean_dec_ref(x_170); @@ -45530,10 +45613,10 @@ lean_dec_ref(x_159); lean_dec_ref(x_158); lean_dec_ref(x_157); lean_dec_ref(x_156); -lean_dec_ref(x_153); -lean_dec(x_151); +lean_dec_ref(x_155); +lean_dec_ref(x_152); lean_dec_ref(x_145); -lean_dec_ref(x_144); +lean_dec(x_144); lean_dec_ref(x_25); lean_dec_ref(x_7); x_213 = !lean_is_exclusive(x_194); @@ -45556,25 +45639,25 @@ return x_215; else { lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; -x_216 = lean_ctor_get(x_154, 2); +x_216 = lean_ctor_get(x_149, 2); lean_inc(x_216); -lean_dec(x_154); -lean_inc(x_150); -lean_inc(x_148); -x_217 = l_Array_toSubarray___redArg(x_24, x_148, x_150); +lean_dec(x_149); +lean_inc(x_153); +lean_inc(x_151); +x_217 = l_Array_toSubarray___redArg(x_24, x_151, x_153); x_218 = lean_array_get_size(x_170); -lean_inc(x_148); +lean_inc(x_151); lean_inc_ref(x_170); -x_219 = l_Array_toSubarray___redArg(x_170, x_148, x_218); +x_219 = l_Array_toSubarray___redArg(x_170, x_151, x_218); x_220 = lean_array_get_size(x_216); -lean_inc(x_148); -x_221 = l_Array_toSubarray___redArg(x_216, x_148, x_220); -x_222 = lean_array_get_size(x_149); -lean_inc(x_148); -x_223 = l_Array_toSubarray___redArg(x_149, x_148, x_222); +lean_inc(x_151); +x_221 = l_Array_toSubarray___redArg(x_216, x_151, x_220); +x_222 = lean_array_get_size(x_150); +lean_inc(x_151); +x_223 = l_Array_toSubarray___redArg(x_150, x_151, x_222); x_224 = lean_array_get_size(x_167); -lean_inc(x_148); -x_225 = l_Array_toSubarray___redArg(x_167, x_148, x_224); +lean_inc(x_151); +x_225 = l_Array_toSubarray___redArg(x_167, x_151, x_224); x_226 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_226, 0, x_225); lean_ctor_set(x_226, 1, x_146); @@ -45596,10 +45679,10 @@ lean_inc(x_162); lean_inc_ref(x_161); lean_inc(x_160); lean_inc_ref(x_159); -lean_inc(x_148); -x_231 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg(x_150, x_148, x_6, x_147, x_2, x_152, x_155, x_148, x_230, x_159, x_160, x_161, x_162, x_163, x_164); -lean_dec(x_148); -lean_dec(x_150); +lean_inc(x_151); +x_231 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg(x_153, x_151, x_6, x_147, x_2, x_154, x_148, x_151, x_230, x_159, x_160, x_161, x_162, x_163, x_164); +lean_dec(x_151); +lean_dec(x_153); if (lean_obj_tag(x_231) == 0) { lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; @@ -45641,15 +45724,15 @@ lean_ctor_set(x_242, 0, x_168); lean_ctor_set(x_242, 1, x_169); lean_ctor_set(x_242, 2, x_170); lean_ctor_set(x_242, 3, x_171); -lean_ctor_set(x_242, 4, x_144); +lean_ctor_set(x_242, 4, x_156); lean_ctor_set(x_242, 5, x_172); x_243 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_243, 0, x_242); -lean_ctor_set(x_243, 1, x_151); -lean_ctor_set(x_243, 2, x_145); -lean_ctor_set(x_243, 3, x_157); -lean_ctor_set(x_243, 4, x_153); -lean_ctor_set(x_243, 5, x_156); +lean_ctor_set(x_243, 1, x_144); +lean_ctor_set(x_243, 2, x_155); +lean_ctor_set(x_243, 3, x_145); +lean_ctor_set(x_243, 4, x_157); +lean_ctor_set(x_243, 5, x_152); lean_ctor_set(x_243, 6, x_237); lean_ctor_set(x_243, 7, x_241); if (lean_is_scalar(x_240)) { @@ -45672,10 +45755,10 @@ lean_dec(x_168); lean_dec_ref(x_158); lean_dec_ref(x_157); lean_dec_ref(x_156); -lean_dec_ref(x_153); -lean_dec(x_151); +lean_dec_ref(x_155); +lean_dec_ref(x_152); lean_dec_ref(x_145); -lean_dec_ref(x_144); +lean_dec(x_144); x_245 = lean_ctor_get(x_238, 0); lean_inc(x_245); if (lean_is_exclusive(x_238)) { @@ -45711,10 +45794,10 @@ lean_dec_ref(x_159); lean_dec_ref(x_158); lean_dec_ref(x_157); lean_dec_ref(x_156); -lean_dec_ref(x_153); -lean_dec(x_151); +lean_dec_ref(x_155); +lean_dec_ref(x_152); lean_dec_ref(x_145); -lean_dec_ref(x_144); +lean_dec(x_144); lean_dec_ref(x_25); lean_dec_ref(x_7); x_248 = lean_ctor_get(x_231, 0); @@ -45748,17 +45831,17 @@ lean_dec_ref(x_159); lean_dec_ref(x_158); lean_dec_ref(x_157); lean_dec_ref(x_156); -lean_dec(x_155); -lean_dec_ref(x_154); -lean_dec_ref(x_153); -lean_dec(x_152); +lean_dec_ref(x_155); +lean_dec(x_154); +lean_dec(x_153); +lean_dec_ref(x_152); lean_dec(x_151); -lean_dec(x_150); +lean_dec_ref(x_150); lean_dec_ref(x_149); lean_dec(x_148); lean_dec_ref(x_146); lean_dec_ref(x_145); -lean_dec_ref(x_144); +lean_dec(x_144); lean_dec_ref(x_25); lean_dec_ref(x_24); lean_dec_ref(x_18); @@ -45785,7 +45868,7 @@ return x_253; { lean_object* x_276; lean_object* x_277; lean_object* x_278; x_276 = lean_array_get_size(x_24); -x_277 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__5___boxed), 10, 2); +x_277 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__5___boxed), 10, 2); lean_closure_set(x_277, 0, x_276); lean_closure_set(x_277, 1, x_256); lean_inc(x_274); @@ -45818,13 +45901,13 @@ x_283 = lean_ctor_get(x_281, 2); lean_inc_ref(x_283); lean_dec(x_281); lean_inc(x_282); -x_284 = l_Lean_mkConst(x_282, x_262); -x_285 = l_Lean_mkAppN(x_284, x_266); -lean_inc_ref(x_257); -x_286 = l_Lean_Expr_app___override(x_285, x_257); -x_287 = l_Lean_mkAppN(x_286, x_265); +x_284 = l_Lean_mkConst(x_282, x_265); +x_285 = l_Lean_mkAppN(x_284, x_258); +lean_inc_ref(x_267); +x_286 = l_Lean_Expr_app___override(x_285, x_267); +x_287 = l_Lean_mkAppN(x_286, x_261); lean_inc_ref(x_287); -x_288 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__6___boxed), 9, 1); +x_288 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__6___boxed), 9, 1); lean_closure_set(x_288, 0, x_287); lean_inc(x_274); lean_inc_ref(x_273); @@ -45839,7 +45922,7 @@ x_290 = lean_ctor_get(x_289, 0); lean_inc(x_290); lean_dec_ref(x_289); lean_inc_ref(x_287); -x_291 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__7___boxed), 10, 2); +x_291 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__7___boxed), 10, 2); lean_closure_set(x_291, 0, x_276); lean_closure_set(x_291, 1, x_287); x_292 = lean_unbox(x_290); @@ -45848,24 +45931,24 @@ if (x_292 == 0) { lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_inc_ref(x_287); -x_293 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__8___boxed), 9, 1); +x_293 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__8___boxed), 9, 1); lean_closure_set(x_293, 0, x_287); lean_inc(x_270); lean_inc_ref(x_269); -x_294 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__9___boxed), 8, 3); +x_294 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__9___boxed), 8, 3); lean_closure_set(x_294, 0, x_293); lean_closure_set(x_294, 1, x_269); lean_closure_set(x_294, 2, x_270); -x_295 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__3; +x_295 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__3; x_296 = l_Lean_indentExpr(x_287); x_297 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_297, 0, x_295); lean_ctor_set(x_297, 1, x_296); -x_298 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__5; +x_298 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__5; x_299 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_299, 0, x_297); lean_ctor_set(x_299, 1, x_298); -x_300 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__10), 2, 1); +x_300 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__10), 2, 1); lean_closure_set(x_300, 0, x_299); lean_inc(x_274); lean_inc_ref(x_273); @@ -45878,20 +45961,20 @@ if (lean_obj_tag(x_301) == 0) { lean_dec_ref(x_301); x_143 = x_291; -x_144 = x_259; +x_144 = x_282; x_145 = x_258; -x_146 = x_261; +x_146 = x_262; x_147 = x_263; x_148 = x_264; -x_149 = x_279; -x_150 = x_276; -x_151 = x_282; -x_152 = x_267; -x_153 = x_257; -x_154 = x_283; -x_155 = x_260; -x_156 = x_265; -x_157 = x_266; +x_149 = x_283; +x_150 = x_279; +x_151 = x_257; +x_152 = x_261; +x_153 = x_276; +x_154 = x_260; +x_155 = x_259; +x_156 = x_266; +x_157 = x_267; x_158 = x_268; x_159 = x_269; x_160 = x_270; @@ -45916,15 +45999,15 @@ lean_dec_ref(x_271); lean_dec(x_270); lean_dec_ref(x_269); lean_dec_ref(x_268); -lean_dec(x_267); +lean_dec_ref(x_267); lean_dec_ref(x_266); -lean_dec_ref(x_265); lean_dec(x_264); +lean_dec_ref(x_262); lean_dec_ref(x_261); lean_dec(x_260); lean_dec_ref(x_259); lean_dec_ref(x_258); -lean_dec_ref(x_257); +lean_dec(x_257); lean_dec_ref(x_25); lean_dec_ref(x_24); lean_dec_ref(x_18); @@ -45962,15 +46045,15 @@ lean_dec_ref(x_271); lean_dec(x_270); lean_dec_ref(x_269); lean_dec_ref(x_268); -lean_dec(x_267); +lean_dec_ref(x_267); lean_dec_ref(x_266); -lean_dec_ref(x_265); lean_dec(x_264); +lean_dec_ref(x_262); lean_dec_ref(x_261); lean_dec(x_260); lean_dec_ref(x_259); lean_dec_ref(x_258); -lean_dec_ref(x_257); +lean_dec(x_257); lean_dec_ref(x_25); lean_dec_ref(x_24); lean_dec_ref(x_18); @@ -45997,20 +46080,20 @@ else { lean_dec_ref(x_287); x_143 = x_291; -x_144 = x_259; +x_144 = x_282; x_145 = x_258; -x_146 = x_261; +x_146 = x_262; x_147 = x_263; x_148 = x_264; -x_149 = x_279; -x_150 = x_276; -x_151 = x_282; -x_152 = x_267; -x_153 = x_257; -x_154 = x_283; -x_155 = x_260; -x_156 = x_265; -x_157 = x_266; +x_149 = x_283; +x_150 = x_279; +x_151 = x_257; +x_152 = x_261; +x_153 = x_276; +x_154 = x_260; +x_155 = x_259; +x_156 = x_266; +x_157 = x_267; x_158 = x_268; x_159 = x_269; x_160 = x_270; @@ -46036,15 +46119,15 @@ lean_dec_ref(x_271); lean_dec(x_270); lean_dec_ref(x_269); lean_dec_ref(x_268); -lean_dec(x_267); +lean_dec_ref(x_267); lean_dec_ref(x_266); -lean_dec_ref(x_265); lean_dec(x_264); +lean_dec_ref(x_262); lean_dec_ref(x_261); lean_dec(x_260); lean_dec_ref(x_259); lean_dec_ref(x_258); -lean_dec_ref(x_257); +lean_dec(x_257); lean_dec_ref(x_25); lean_dec_ref(x_24); lean_dec_ref(x_18); @@ -46078,16 +46161,16 @@ lean_dec_ref(x_271); lean_dec(x_270); lean_dec_ref(x_269); lean_dec_ref(x_268); -lean_dec(x_267); +lean_dec_ref(x_267); lean_dec_ref(x_266); -lean_dec_ref(x_265); +lean_dec(x_265); lean_dec(x_264); -lean_dec(x_262); +lean_dec_ref(x_262); lean_dec_ref(x_261); lean_dec(x_260); lean_dec_ref(x_259); lean_dec_ref(x_258); -lean_dec_ref(x_257); +lean_dec(x_257); lean_dec_ref(x_25); lean_dec_ref(x_24); lean_dec_ref(x_18); @@ -46120,16 +46203,16 @@ lean_dec_ref(x_271); lean_dec(x_270); lean_dec_ref(x_269); lean_dec_ref(x_268); -lean_dec(x_267); +lean_dec_ref(x_267); lean_dec_ref(x_266); -lean_dec_ref(x_265); +lean_dec(x_265); lean_dec(x_264); -lean_dec(x_262); +lean_dec_ref(x_262); lean_dec_ref(x_261); lean_dec(x_260); lean_dec_ref(x_259); lean_dec_ref(x_258); -lean_dec_ref(x_257); +lean_dec(x_257); lean_dec_ref(x_255); lean_dec_ref(x_25); lean_dec_ref(x_24); @@ -46158,12 +46241,12 @@ return x_316; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; size_t x_343; lean_object* x_344; x_335 = lean_unsigned_to_nat(0u); x_336 = l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__28_spec__33___redArg___closed__0; -x_337 = l_Array_reverse___redArg(x_326); +x_337 = l_Array_reverse___redArg(x_322); x_338 = lean_array_get_size(x_337); x_339 = l_Array_toSubarray___redArg(x_337, x_335, x_338); -lean_inc_ref(x_323); -x_340 = l_Array_reverse___redArg(x_323); -x_341 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__6; +lean_inc_ref(x_321); +x_340 = l_Array_reverse___redArg(x_321); +x_341 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__6; x_342 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_342, 0, x_339); lean_ctor_set(x_342, 1, x_341); @@ -46173,7 +46256,7 @@ lean_inc_ref(x_332); lean_inc(x_331); lean_inc_ref(x_330); lean_inc_ref(x_328); -x_344 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__49(x_340, x_343, x_325, x_342, x_328, x_329, x_330, x_331, x_332, x_333); +x_344 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__50(x_340, x_343, x_320, x_342, x_328, x_329, x_330, x_331, x_332, x_333); lean_dec_ref(x_340); if (lean_obj_tag(x_344) == 0) { @@ -46192,30 +46275,30 @@ if (x_2 == 0) { lean_object* x_349; lean_object* x_350; lean_free_object(x_345); -lean_dec(x_321); +lean_dec(x_324); lean_dec_ref(x_255); x_349 = lean_ctor_get(x_347, 0); lean_inc(x_349); x_350 = lean_ctor_get(x_347, 1); lean_inc(x_350); lean_dec(x_347); -x_100 = x_327; -x_101 = x_320; -x_102 = x_329; -x_103 = x_350; +x_100 = x_328; +x_101 = x_330; +x_102 = x_333; +x_103 = x_319; x_104 = x_336; -x_105 = x_322; -x_106 = x_335; -x_107 = x_328; -x_108 = x_349; -x_109 = x_319; -x_110 = x_323; -x_111 = x_330; -x_112 = x_333; -x_113 = x_332; -x_114 = x_324; -x_115 = lean_box(0); -x_116 = x_331; +x_105 = x_323; +x_106 = x_332; +x_107 = x_329; +x_108 = lean_box(0); +x_109 = x_335; +x_110 = x_327; +x_111 = x_321; +x_112 = x_349; +x_113 = x_331; +x_114 = x_325; +x_115 = x_350; +x_116 = x_326; goto block_142; } else @@ -46235,12 +46318,12 @@ lean_inc_ref(x_327); x_354 = lean_array_to_list(x_327); lean_inc(x_354); x_355 = l_Lean_mkConst(x_19, x_354); -x_356 = l_Lean_mkAppN(x_355, x_324); -lean_inc_ref(x_319); -x_357 = l_Lean_Expr_app___override(x_356, x_319); -x_358 = l_Lean_mkAppN(x_357, x_323); +x_356 = l_Lean_mkAppN(x_355, x_319); +lean_inc_ref(x_326); +x_357 = l_Lean_Expr_app___override(x_356, x_326); +x_358 = l_Lean_mkAppN(x_357, x_321); lean_inc_ref(x_358); -x_359 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__4___boxed), 9, 1); +x_359 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__4___boxed), 9, 1); lean_closure_set(x_359, 0, x_358); lean_inc(x_333); lean_inc_ref(x_332); @@ -46260,25 +46343,25 @@ if (x_362 == 0) { lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_inc_ref(x_358); -x_363 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__11___boxed), 9, 1); +x_363 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__11___boxed), 9, 1); lean_closure_set(x_363, 0, x_358); lean_inc(x_329); lean_inc_ref(x_328); -x_364 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__9___boxed), 8, 3); +x_364 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__9___boxed), 8, 3); lean_closure_set(x_364, 0, x_363); lean_closure_set(x_364, 1, x_328); lean_closure_set(x_364, 2, x_329); -x_365 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__8; +x_365 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__8; lean_inc_ref(x_358); x_366 = l_Lean_indentExpr(x_358); lean_ctor_set_tag(x_347, 7); lean_ctor_set(x_347, 1, x_366); lean_ctor_set(x_347, 0, x_365); -x_367 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__5; +x_367 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__5; lean_ctor_set_tag(x_345, 7); lean_ctor_set(x_345, 1, x_367); lean_ctor_set(x_345, 0, x_347); -x_368 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__10), 2, 1); +x_368 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__10), 2, 1); lean_closure_set(x_368, 0, x_345); lean_inc(x_333); lean_inc_ref(x_332); @@ -46291,17 +46374,17 @@ if (lean_obj_tag(x_369) == 0) { lean_dec_ref(x_369); x_256 = x_358; -x_257 = x_319; -x_258 = x_327; -x_259 = x_320; -x_260 = x_321; -x_261 = x_336; -x_262 = x_354; -x_263 = x_322; -x_264 = x_335; -x_265 = x_323; -x_266 = x_324; -x_267 = x_352; +x_257 = x_335; +x_258 = x_319; +x_259 = x_327; +x_260 = x_352; +x_261 = x_321; +x_262 = x_336; +x_263 = x_323; +x_264 = x_324; +x_265 = x_354; +x_266 = x_325; +x_267 = x_326; x_268 = x_353; x_269 = x_328; x_270 = x_329; @@ -46326,10 +46409,10 @@ lean_dec_ref(x_330); lean_dec(x_329); lean_dec_ref(x_328); lean_dec_ref(x_327); -lean_dec_ref(x_324); -lean_dec_ref(x_323); -lean_dec(x_321); -lean_dec_ref(x_320); +lean_dec_ref(x_326); +lean_dec_ref(x_325); +lean_dec(x_324); +lean_dec_ref(x_321); lean_dec_ref(x_319); lean_dec_ref(x_255); lean_dec_ref(x_25); @@ -46369,10 +46452,10 @@ lean_dec_ref(x_330); lean_dec(x_329); lean_dec_ref(x_328); lean_dec_ref(x_327); -lean_dec_ref(x_324); -lean_dec_ref(x_323); -lean_dec(x_321); -lean_dec_ref(x_320); +lean_dec_ref(x_326); +lean_dec_ref(x_325); +lean_dec(x_324); +lean_dec_ref(x_321); lean_dec_ref(x_319); lean_dec_ref(x_255); lean_dec_ref(x_25); @@ -46402,17 +46485,17 @@ else lean_free_object(x_347); lean_free_object(x_345); x_256 = x_358; -x_257 = x_319; -x_258 = x_327; -x_259 = x_320; -x_260 = x_321; -x_261 = x_336; -x_262 = x_354; -x_263 = x_322; -x_264 = x_335; -x_265 = x_323; -x_266 = x_324; -x_267 = x_352; +x_257 = x_335; +x_258 = x_319; +x_259 = x_327; +x_260 = x_352; +x_261 = x_321; +x_262 = x_336; +x_263 = x_323; +x_264 = x_324; +x_265 = x_354; +x_266 = x_325; +x_267 = x_326; x_268 = x_353; x_269 = x_328; x_270 = x_329; @@ -46440,10 +46523,10 @@ lean_dec_ref(x_330); lean_dec(x_329); lean_dec_ref(x_328); lean_dec_ref(x_327); -lean_dec_ref(x_324); -lean_dec_ref(x_323); -lean_dec(x_321); -lean_dec_ref(x_320); +lean_dec_ref(x_326); +lean_dec_ref(x_325); +lean_dec(x_324); +lean_dec_ref(x_321); lean_dec_ref(x_319); lean_dec_ref(x_255); lean_dec_ref(x_25); @@ -46480,12 +46563,12 @@ lean_inc_ref(x_327); x_381 = lean_array_to_list(x_327); lean_inc(x_381); x_382 = l_Lean_mkConst(x_19, x_381); -x_383 = l_Lean_mkAppN(x_382, x_324); -lean_inc_ref(x_319); -x_384 = l_Lean_Expr_app___override(x_383, x_319); -x_385 = l_Lean_mkAppN(x_384, x_323); +x_383 = l_Lean_mkAppN(x_382, x_319); +lean_inc_ref(x_326); +x_384 = l_Lean_Expr_app___override(x_383, x_326); +x_385 = l_Lean_mkAppN(x_384, x_321); lean_inc_ref(x_385); -x_386 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__4___boxed), 9, 1); +x_386 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__4___boxed), 9, 1); lean_closure_set(x_386, 0, x_385); lean_inc(x_333); lean_inc_ref(x_332); @@ -46505,25 +46588,25 @@ if (x_389 == 0) { lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_inc_ref(x_385); -x_390 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__11___boxed), 9, 1); +x_390 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__11___boxed), 9, 1); lean_closure_set(x_390, 0, x_385); lean_inc(x_329); lean_inc_ref(x_328); -x_391 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__9___boxed), 8, 3); +x_391 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__9___boxed), 8, 3); lean_closure_set(x_391, 0, x_390); lean_closure_set(x_391, 1, x_328); lean_closure_set(x_391, 2, x_329); -x_392 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__8; +x_392 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__8; lean_inc_ref(x_385); x_393 = l_Lean_indentExpr(x_385); x_394 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_394, 0, x_392); lean_ctor_set(x_394, 1, x_393); -x_395 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__5; +x_395 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__5; lean_ctor_set_tag(x_345, 7); lean_ctor_set(x_345, 1, x_395); lean_ctor_set(x_345, 0, x_394); -x_396 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__10), 2, 1); +x_396 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__10), 2, 1); lean_closure_set(x_396, 0, x_345); lean_inc(x_333); lean_inc_ref(x_332); @@ -46536,17 +46619,17 @@ if (lean_obj_tag(x_397) == 0) { lean_dec_ref(x_397); x_256 = x_385; -x_257 = x_319; -x_258 = x_327; -x_259 = x_320; -x_260 = x_321; -x_261 = x_336; -x_262 = x_381; -x_263 = x_322; -x_264 = x_335; -x_265 = x_323; -x_266 = x_324; -x_267 = x_379; +x_257 = x_335; +x_258 = x_319; +x_259 = x_327; +x_260 = x_379; +x_261 = x_321; +x_262 = x_336; +x_263 = x_323; +x_264 = x_324; +x_265 = x_381; +x_266 = x_325; +x_267 = x_326; x_268 = x_380; x_269 = x_328; x_270 = x_329; @@ -46571,10 +46654,10 @@ lean_dec_ref(x_330); lean_dec(x_329); lean_dec_ref(x_328); lean_dec_ref(x_327); -lean_dec_ref(x_324); -lean_dec_ref(x_323); -lean_dec(x_321); -lean_dec_ref(x_320); +lean_dec_ref(x_326); +lean_dec_ref(x_325); +lean_dec(x_324); +lean_dec_ref(x_321); lean_dec_ref(x_319); lean_dec_ref(x_255); lean_dec_ref(x_25); @@ -46615,10 +46698,10 @@ lean_dec_ref(x_330); lean_dec(x_329); lean_dec_ref(x_328); lean_dec_ref(x_327); -lean_dec_ref(x_324); -lean_dec_ref(x_323); -lean_dec(x_321); -lean_dec_ref(x_320); +lean_dec_ref(x_326); +lean_dec_ref(x_325); +lean_dec(x_324); +lean_dec_ref(x_321); lean_dec_ref(x_319); lean_dec_ref(x_255); lean_dec_ref(x_25); @@ -46648,17 +46731,17 @@ else { lean_free_object(x_345); x_256 = x_385; -x_257 = x_319; -x_258 = x_327; -x_259 = x_320; -x_260 = x_321; -x_261 = x_336; -x_262 = x_381; -x_263 = x_322; -x_264 = x_335; -x_265 = x_323; -x_266 = x_324; -x_267 = x_379; +x_257 = x_335; +x_258 = x_319; +x_259 = x_327; +x_260 = x_379; +x_261 = x_321; +x_262 = x_336; +x_263 = x_323; +x_264 = x_324; +x_265 = x_381; +x_266 = x_325; +x_267 = x_326; x_268 = x_380; x_269 = x_328; x_270 = x_329; @@ -46685,10 +46768,10 @@ lean_dec_ref(x_330); lean_dec(x_329); lean_dec_ref(x_328); lean_dec_ref(x_327); -lean_dec_ref(x_324); -lean_dec_ref(x_323); -lean_dec(x_321); -lean_dec_ref(x_320); +lean_dec_ref(x_326); +lean_dec_ref(x_325); +lean_dec(x_324); +lean_dec_ref(x_321); lean_dec_ref(x_319); lean_dec_ref(x_255); lean_dec_ref(x_25); @@ -46719,30 +46802,30 @@ else { lean_object* x_407; lean_object* x_408; lean_free_object(x_345); -lean_dec(x_321); +lean_dec(x_324); lean_dec_ref(x_255); x_407 = lean_ctor_get(x_347, 0); lean_inc(x_407); x_408 = lean_ctor_get(x_347, 1); lean_inc(x_408); lean_dec(x_347); -x_100 = x_327; -x_101 = x_320; -x_102 = x_329; -x_103 = x_408; +x_100 = x_328; +x_101 = x_330; +x_102 = x_333; +x_103 = x_319; x_104 = x_336; -x_105 = x_322; -x_106 = x_335; -x_107 = x_328; -x_108 = x_407; -x_109 = x_319; -x_110 = x_323; -x_111 = x_330; -x_112 = x_333; -x_113 = x_332; -x_114 = x_324; -x_115 = lean_box(0); -x_116 = x_331; +x_105 = x_323; +x_106 = x_332; +x_107 = x_329; +x_108 = lean_box(0); +x_109 = x_335; +x_110 = x_327; +x_111 = x_321; +x_112 = x_407; +x_113 = x_331; +x_114 = x_325; +x_115 = x_408; +x_116 = x_326; goto block_142; } } @@ -46756,30 +46839,30 @@ lean_dec(x_345); if (x_2 == 0) { lean_object* x_410; lean_object* x_411; -lean_dec(x_321); +lean_dec(x_324); lean_dec_ref(x_255); x_410 = lean_ctor_get(x_409, 0); lean_inc(x_410); x_411 = lean_ctor_get(x_409, 1); lean_inc(x_411); lean_dec(x_409); -x_100 = x_327; -x_101 = x_320; -x_102 = x_329; -x_103 = x_411; +x_100 = x_328; +x_101 = x_330; +x_102 = x_333; +x_103 = x_319; x_104 = x_336; -x_105 = x_322; -x_106 = x_335; -x_107 = x_328; -x_108 = x_410; -x_109 = x_319; -x_110 = x_323; -x_111 = x_330; -x_112 = x_333; -x_113 = x_332; -x_114 = x_324; -x_115 = lean_box(0); -x_116 = x_331; +x_105 = x_323; +x_106 = x_332; +x_107 = x_329; +x_108 = lean_box(0); +x_109 = x_335; +x_110 = x_327; +x_111 = x_321; +x_112 = x_410; +x_113 = x_331; +x_114 = x_325; +x_115 = x_411; +x_116 = x_326; goto block_142; } else @@ -46805,12 +46888,12 @@ lean_inc_ref(x_327); x_415 = lean_array_to_list(x_327); lean_inc(x_415); x_416 = l_Lean_mkConst(x_19, x_415); -x_417 = l_Lean_mkAppN(x_416, x_324); -lean_inc_ref(x_319); -x_418 = l_Lean_Expr_app___override(x_417, x_319); -x_419 = l_Lean_mkAppN(x_418, x_323); +x_417 = l_Lean_mkAppN(x_416, x_319); +lean_inc_ref(x_326); +x_418 = l_Lean_Expr_app___override(x_417, x_326); +x_419 = l_Lean_mkAppN(x_418, x_321); lean_inc_ref(x_419); -x_420 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__4___boxed), 9, 1); +x_420 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__4___boxed), 9, 1); lean_closure_set(x_420, 0, x_419); lean_inc(x_333); lean_inc_ref(x_332); @@ -46830,15 +46913,15 @@ if (x_423 == 0) { lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_inc_ref(x_419); -x_424 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__11___boxed), 9, 1); +x_424 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__11___boxed), 9, 1); lean_closure_set(x_424, 0, x_419); lean_inc(x_329); lean_inc_ref(x_328); -x_425 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__9___boxed), 8, 3); +x_425 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__9___boxed), 8, 3); lean_closure_set(x_425, 0, x_424); lean_closure_set(x_425, 1, x_328); lean_closure_set(x_425, 2, x_329); -x_426 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__8; +x_426 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__8; lean_inc_ref(x_419); x_427 = l_Lean_indentExpr(x_419); if (lean_is_scalar(x_414)) { @@ -46849,11 +46932,11 @@ if (lean_is_scalar(x_414)) { } lean_ctor_set(x_428, 0, x_426); lean_ctor_set(x_428, 1, x_427); -x_429 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__5; +x_429 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__5; x_430 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_430, 0, x_428); lean_ctor_set(x_430, 1, x_429); -x_431 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__10), 2, 1); +x_431 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__10), 2, 1); lean_closure_set(x_431, 0, x_430); lean_inc(x_333); lean_inc_ref(x_332); @@ -46866,17 +46949,17 @@ if (lean_obj_tag(x_432) == 0) { lean_dec_ref(x_432); x_256 = x_419; -x_257 = x_319; -x_258 = x_327; -x_259 = x_320; -x_260 = x_321; -x_261 = x_336; -x_262 = x_415; -x_263 = x_322; -x_264 = x_335; -x_265 = x_323; -x_266 = x_324; -x_267 = x_412; +x_257 = x_335; +x_258 = x_319; +x_259 = x_327; +x_260 = x_412; +x_261 = x_321; +x_262 = x_336; +x_263 = x_323; +x_264 = x_324; +x_265 = x_415; +x_266 = x_325; +x_267 = x_326; x_268 = x_413; x_269 = x_328; x_270 = x_329; @@ -46901,10 +46984,10 @@ lean_dec_ref(x_330); lean_dec(x_329); lean_dec_ref(x_328); lean_dec_ref(x_327); -lean_dec_ref(x_324); -lean_dec_ref(x_323); -lean_dec(x_321); -lean_dec_ref(x_320); +lean_dec_ref(x_326); +lean_dec_ref(x_325); +lean_dec(x_324); +lean_dec_ref(x_321); lean_dec_ref(x_319); lean_dec_ref(x_255); lean_dec_ref(x_25); @@ -46945,10 +47028,10 @@ lean_dec_ref(x_330); lean_dec(x_329); lean_dec_ref(x_328); lean_dec_ref(x_327); -lean_dec_ref(x_324); -lean_dec_ref(x_323); -lean_dec(x_321); -lean_dec_ref(x_320); +lean_dec_ref(x_326); +lean_dec_ref(x_325); +lean_dec(x_324); +lean_dec_ref(x_321); lean_dec_ref(x_319); lean_dec_ref(x_255); lean_dec_ref(x_25); @@ -46978,17 +47061,17 @@ else { lean_dec(x_414); x_256 = x_419; -x_257 = x_319; -x_258 = x_327; -x_259 = x_320; -x_260 = x_321; -x_261 = x_336; -x_262 = x_415; -x_263 = x_322; -x_264 = x_335; -x_265 = x_323; -x_266 = x_324; -x_267 = x_412; +x_257 = x_335; +x_258 = x_319; +x_259 = x_327; +x_260 = x_412; +x_261 = x_321; +x_262 = x_336; +x_263 = x_323; +x_264 = x_324; +x_265 = x_415; +x_266 = x_325; +x_267 = x_326; x_268 = x_413; x_269 = x_328; x_270 = x_329; @@ -47015,10 +47098,10 @@ lean_dec_ref(x_330); lean_dec(x_329); lean_dec_ref(x_328); lean_dec_ref(x_327); -lean_dec_ref(x_324); -lean_dec_ref(x_323); -lean_dec(x_321); -lean_dec_ref(x_320); +lean_dec_ref(x_326); +lean_dec_ref(x_325); +lean_dec(x_324); +lean_dec_ref(x_321); lean_dec_ref(x_319); lean_dec_ref(x_255); lean_dec_ref(x_25); @@ -47047,30 +47130,30 @@ return x_441; else { lean_object* x_442; lean_object* x_443; -lean_dec(x_321); +lean_dec(x_324); lean_dec_ref(x_255); x_442 = lean_ctor_get(x_409, 0); lean_inc(x_442); x_443 = lean_ctor_get(x_409, 1); lean_inc(x_443); lean_dec(x_409); -x_100 = x_327; -x_101 = x_320; -x_102 = x_329; -x_103 = x_443; +x_100 = x_328; +x_101 = x_330; +x_102 = x_333; +x_103 = x_319; x_104 = x_336; -x_105 = x_322; -x_106 = x_335; -x_107 = x_328; -x_108 = x_442; -x_109 = x_319; -x_110 = x_323; -x_111 = x_330; -x_112 = x_333; -x_113 = x_332; -x_114 = x_324; -x_115 = lean_box(0); -x_116 = x_331; +x_105 = x_323; +x_106 = x_332; +x_107 = x_329; +x_108 = lean_box(0); +x_109 = x_335; +x_110 = x_327; +x_111 = x_321; +x_112 = x_442; +x_113 = x_331; +x_114 = x_325; +x_115 = x_443; +x_116 = x_326; goto block_142; } } @@ -47086,10 +47169,10 @@ lean_dec_ref(x_330); lean_dec(x_329); lean_dec_ref(x_328); lean_dec_ref(x_327); -lean_dec_ref(x_324); -lean_dec_ref(x_323); -lean_dec(x_321); -lean_dec_ref(x_320); +lean_dec_ref(x_326); +lean_dec_ref(x_325); +lean_dec(x_324); +lean_dec_ref(x_321); lean_dec_ref(x_319); lean_dec_ref(x_255); lean_dec_ref(x_25); @@ -47128,7 +47211,7 @@ lean_inc(x_450); lean_inc_ref(x_449); lean_inc_ref(x_21); lean_inc_ref(x_4); -x_458 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__46(x_4, x_456, x_457, x_21, x_449, x_450, x_451, x_452, x_453, x_454); +x_458 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__47(x_4, x_456, x_457, x_21, x_449, x_450, x_451, x_452, x_453, x_454); if (lean_obj_tag(x_458) == 0) { lean_object* x_459; size_t x_460; lean_object* x_461; @@ -47143,7 +47226,7 @@ lean_inc_ref(x_451); lean_inc(x_450); lean_inc_ref(x_449); lean_inc_ref(x_23); -x_461 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__46(x_4, x_460, x_457, x_23, x_449, x_450, x_451, x_452, x_453, x_454); +x_461 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__47(x_4, x_460, x_457, x_23, x_449, x_450, x_451, x_452, x_453, x_454); if (lean_obj_tag(x_461) == 0) { lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; uint8_t x_466; lean_object* x_467; @@ -47151,11 +47234,11 @@ x_462 = lean_ctor_get(x_461, 0); lean_inc(x_462); lean_dec_ref(x_461); x_463 = lean_box(x_3); -x_464 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___boxed__const__1; +x_464 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___boxed__const__1; lean_inc_ref(x_23); lean_inc(x_462); lean_inc_ref(x_18); -x_465 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___boxed), 15, 6); +x_465 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___boxed), 15, 6); lean_closure_set(x_465, 0, x_5); lean_closure_set(x_465, 1, x_18); lean_closure_set(x_465, 2, x_462); @@ -47170,7 +47253,7 @@ lean_inc_ref(x_451); lean_inc(x_450); lean_inc_ref(x_449); lean_inc_ref(x_22); -x_467 = l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__48___redArg(x_22, x_465, x_466, x_449, x_450, x_451, x_452, x_453, x_454); +x_467 = l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__49___redArg(x_22, x_465, x_466, x_449, x_450, x_451, x_452, x_453, x_454); if (lean_obj_tag(x_467) == 0) { lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; @@ -47193,14 +47276,14 @@ x_474 = lean_ctor_get(x_470, 1); lean_inc(x_474); lean_dec(x_470); lean_inc_ref(x_20); -x_319 = x_472; -x_320 = x_474; -x_321 = x_448; -x_322 = x_466; -x_323 = x_462; -x_324 = x_459; -x_325 = x_457; -x_326 = x_473; +x_319 = x_459; +x_320 = x_457; +x_321 = x_462; +x_322 = x_473; +x_323 = x_466; +x_324 = x_448; +x_325 = x_474; +x_326 = x_472; x_327 = x_20; x_328 = x_449; x_329 = x_450; @@ -47229,14 +47312,14 @@ lean_dec(x_470); x_479 = lean_ctor_get(x_471, 0); lean_inc_ref(x_20); x_480 = lean_array_set(x_20, x_479, x_476); -x_319 = x_475; -x_320 = x_478; -x_321 = x_448; -x_322 = x_466; -x_323 = x_462; -x_324 = x_459; -x_325 = x_457; -x_326 = x_477; +x_319 = x_459; +x_320 = x_457; +x_321 = x_462; +x_322 = x_477; +x_323 = x_466; +x_324 = x_448; +x_325 = x_478; +x_326 = x_475; x_327 = x_480; x_328 = x_449; x_329 = x_450; @@ -47389,7 +47472,7 @@ return x_510; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { switch (lean_obj_tag(x_1)) { @@ -47400,7 +47483,7 @@ x_12 = lean_ctor_get(x_1, 0); lean_inc_ref(x_12); lean_dec_ref(x_1); lean_inc_ref(x_2); -x_13 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__0___boxed), 9, 1); +x_13 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__0___boxed), 9, 1); lean_closure_set(x_13, 0, x_2); lean_inc(x_10); lean_inc_ref(x_9); @@ -47430,7 +47513,7 @@ lean_dec_ref(x_12); if (x_4 == 0) { lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_25 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59___lam__1___closed__1; +x_25 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__1___closed__1; x_26 = lean_unsigned_to_nat(0u); x_27 = l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00Lean_Elab_Tactic_Do_ProofMode_mRevertForallN___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__15_spec__24_spec__28_spec__33___redArg___closed__0; lean_inc_ref(x_3); @@ -47440,15 +47523,17 @@ lean_inc(x_8); lean_inc_ref(x_7); lean_inc(x_6); lean_inc_ref(x_5); -x_28 = lean_apply_10(x_3, x_25, x_26, x_27, x_5, x_6, x_7, x_8, x_9, x_10, lean_box(0)); +lean_inc_ref(x_2); +x_28 = lean_apply_11(x_3, x_25, x_2, x_26, x_27, x_5, x_6, x_7, x_8, x_9, x_10, lean_box(0)); if (lean_obj_tag(x_28) == 0) { lean_object* x_29; lean_object* x_30; lean_object* x_31; x_29 = lean_ctor_get(x_28, 0); lean_inc(x_29); lean_dec_ref(x_28); -x_30 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62___lam__1___closed__1; -x_31 = lean_apply_10(x_3, x_30, x_16, x_27, x_5, x_6, x_7, x_8, x_9, x_10, lean_box(0)); +x_30 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__63___lam__1___closed__1; +lean_inc_ref(x_2); +x_31 = lean_apply_11(x_3, x_30, x_2, x_16, x_27, x_5, x_6, x_7, x_8, x_9, x_10, lean_box(0)); if (lean_obj_tag(x_31) == 0) { uint8_t x_32; @@ -47457,7 +47542,7 @@ if (x_32 == 0) { lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; x_33 = lean_ctor_get(x_31, 0); -x_34 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__1; +x_34 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__1; x_35 = lean_box(0); x_36 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_36, 0, x_15); @@ -47473,7 +47558,7 @@ lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean x_39 = lean_ctor_get(x_31, 0); lean_inc(x_39); lean_dec(x_31); -x_40 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__1; +x_40 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__1; x_41 = lean_box(0); x_42 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_42, 0, x_15); @@ -47537,8 +47622,9 @@ lean_inc(x_6); lean_inc_ref(x_5); lean_inc_ref(x_20); lean_inc(x_48); +lean_inc_ref(x_2); lean_inc_ref(x_3); -x_51 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59(x_3, x_16, x_4, x_48, x_49, x_20, x_50, x_5, x_6, x_7, x_8, x_9, x_10); +x_51 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60(x_3, x_2, x_16, x_4, x_48, x_49, x_20, x_50, x_5, x_6, x_7, x_8, x_9, x_10); if (lean_obj_tag(x_51) == 0) { lean_object* x_52; lean_object* x_53; lean_object* x_54; @@ -47547,7 +47633,8 @@ lean_inc(x_52); lean_dec_ref(x_51); lean_inc_ref(x_20); x_53 = l_Lean_mkNot(x_20); -x_54 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60(x_3, x_16, x_4, x_48, x_49, x_53, x_50, x_5, x_6, x_7, x_8, x_9, x_10); +lean_inc_ref(x_2); +x_54 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__61(x_3, x_2, x_16, x_4, x_48, x_49, x_53, x_50, x_5, x_6, x_7, x_8, x_9, x_10); if (lean_obj_tag(x_54) == 0) { uint8_t x_55; @@ -47556,7 +47643,7 @@ if (x_55 == 0) { lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; x_56 = lean_ctor_get(x_54, 0); -x_57 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__3; +x_57 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__3; x_58 = lean_box(0); x_59 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_59, 0, x_15); @@ -47572,7 +47659,7 @@ lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean x_62 = lean_ctor_get(x_54, 0); lean_inc(x_62); lean_dec(x_54); -x_63 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__3; +x_63 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__3; x_64 = lean_box(0); x_65 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_65, 0, x_15); @@ -47679,7 +47766,7 @@ x_75 = lean_ctor_get(x_1, 0); lean_inc_ref(x_75); lean_dec_ref(x_1); lean_inc_ref(x_2); -x_76 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__0___boxed), 9, 1); +x_76 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__0___boxed), 9, 1); lean_closure_set(x_76, 0, x_2); lean_inc(x_10); lean_inc_ref(x_9); @@ -47722,8 +47809,9 @@ lean_inc(x_6); lean_inc_ref(x_5); lean_inc_ref(x_86); lean_inc(x_81); +lean_inc_ref(x_2); lean_inc_ref(x_3); -x_89 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__61(x_3, x_81, x_87, x_86, x_88, x_5, x_6, x_7, x_8, x_9, x_10); +x_89 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62(x_3, x_2, x_81, x_87, x_86, x_88, x_5, x_6, x_7, x_8, x_9, x_10); if (lean_obj_tag(x_89) == 0) { lean_object* x_90; lean_object* x_91; lean_object* x_92; @@ -47732,7 +47820,8 @@ lean_inc(x_90); lean_dec_ref(x_89); lean_inc_ref(x_86); x_91 = l_Lean_mkNot(x_86); -x_92 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62(x_82, x_3, x_81, x_87, x_91, x_88, x_5, x_6, x_7, x_8, x_9, x_10); +lean_inc_ref(x_2); +x_92 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__63(x_82, x_3, x_2, x_81, x_87, x_91, x_88, x_5, x_6, x_7, x_8, x_9, x_10); if (lean_obj_tag(x_92) == 0) { uint8_t x_93; @@ -47748,7 +47837,7 @@ x_97 = lean_nat_sub(x_96, x_82); lean_dec(x_96); x_98 = l_Lean_Expr_getRevArg_x21(x_75, x_97); lean_dec_ref(x_75); -x_99 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__3; +x_99 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__3; x_100 = lean_box(0); x_101 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_101, 0, x_78); @@ -47771,7 +47860,7 @@ x_107 = lean_nat_sub(x_106, x_82); lean_dec(x_106); x_108 = l_Lean_Expr_getRevArg_x21(x_75, x_107); lean_dec_ref(x_75); -x_109 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__3; +x_109 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__3; x_110 = lean_box(0); x_111 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_111, 0, x_78); @@ -47873,65 +47962,102 @@ return x_120; } default: { -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_121; lean_object* x_122; size_t x_123; size_t x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; x_121 = lean_ctor_get(x_1, 0); lean_inc_ref(x_121); lean_dec_ref(x_1); -x_122 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__4___boxed), 12, 1); -lean_closure_set(x_122, 0, x_3); -x_123 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__1___boxed), 10, 1); -lean_closure_set(x_123, 0, x_2); -x_124 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__2___boxed), 8, 0); -x_125 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__3___boxed), 8, 0); -x_126 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37(x_121, x_4, x_4, x_124, x_123, x_122, x_125, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_126) == 0) +x_122 = lean_ctor_get(x_121, 5); +x_123 = lean_array_size(x_122); +x_124 = 0; +lean_inc_ref(x_122); +x_125 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37(x_123, x_124, x_122); +lean_inc_ref(x_122); +x_126 = l_Array_mask___redArg(x_125, x_122); +x_127 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__4___boxed), 10, 2); +lean_closure_set(x_127, 0, x_2); +lean_closure_set(x_127, 1, x_126); +lean_inc(x_10); +lean_inc_ref(x_9); +lean_inc(x_8); +lean_inc_ref(x_7); +lean_inc_ref(x_5); +x_128 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_127, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_128) == 0) { -uint8_t x_127; -x_127 = !lean_is_exclusive(x_126); -if (x_127 == 0) +lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; +x_129 = lean_ctor_get(x_128, 0); +lean_inc(x_129); +lean_dec_ref(x_128); +x_130 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__1___boxed), 12, 1); +lean_closure_set(x_130, 0, x_3); +x_131 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__2___boxed), 8, 0); +x_132 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__3___boxed), 8, 0); +x_133 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___lam__5___boxed), 11, 2); +lean_closure_set(x_133, 0, x_125); +lean_closure_set(x_133, 1, x_129); +x_134 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38(x_121, x_4, x_4, x_131, x_133, x_130, x_132, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_134) == 0) { -lean_object* x_128; lean_object* x_129; -x_128 = lean_ctor_get(x_126, 0); -x_129 = l_Lean_Meta_MatcherApp_toExpr(x_128); -lean_ctor_set(x_126, 0, x_129); -return x_126; +uint8_t x_135; +x_135 = !lean_is_exclusive(x_134); +if (x_135 == 0) +{ +lean_object* x_136; lean_object* x_137; +x_136 = lean_ctor_get(x_134, 0); +x_137 = l_Lean_Meta_MatcherApp_toExpr(x_136); +lean_ctor_set(x_134, 0, x_137); +return x_134; } else { -lean_object* x_130; lean_object* x_131; lean_object* x_132; -x_130 = lean_ctor_get(x_126, 0); -lean_inc(x_130); -lean_dec(x_126); -x_131 = l_Lean_Meta_MatcherApp_toExpr(x_130); -x_132 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_132, 0, x_131); -return x_132; +lean_object* x_138; lean_object* x_139; lean_object* x_140; +x_138 = lean_ctor_get(x_134, 0); +lean_inc(x_138); +lean_dec(x_134); +x_139 = l_Lean_Meta_MatcherApp_toExpr(x_138); +x_140 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_140, 0, x_139); +return x_140; } } else { -uint8_t x_133; -x_133 = !lean_is_exclusive(x_126); -if (x_133 == 0) +uint8_t x_141; +x_141 = !lean_is_exclusive(x_134); +if (x_141 == 0) { -return x_126; +return x_134; } else { -lean_object* x_134; lean_object* x_135; -x_134 = lean_ctor_get(x_126, 0); -lean_inc(x_134); -lean_dec(x_126); -x_135 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_135, 0, x_134); -return x_135; +lean_object* x_142; lean_object* x_143; +x_142 = lean_ctor_get(x_134, 0); +lean_inc(x_142); +lean_dec(x_134); +x_143 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_143, 0, x_142); +return x_143; } } } +else +{ +lean_dec_ref(x_125); +lean_dec_ref(x_121); +lean_dec(x_10); +lean_dec_ref(x_9); +lean_dec(x_8); +lean_dec_ref(x_7); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec_ref(x_3); +return x_128; +} +} } } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; @@ -47941,8 +48067,8 @@ x_16 = lean_unsigned_to_nat(4u); x_17 = lean_unsigned_to_nat(0u); x_18 = l_Array_extract___redArg(x_15, x_17, x_16); lean_dec_ref(x_15); -x_19 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__0___closed__0; -x_20 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__0___closed__1; +x_19 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__0___closed__0; +x_20 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__0___closed__1; x_21 = l_Lean_Name_mkStr4(x_8, x_9, x_19, x_20); x_22 = l_Lean_mkConst(x_21, x_10); x_23 = l_Lean_mkAppN(x_22, x_18); @@ -47958,23 +48084,23 @@ lean_ctor_set(x_25, 3, x_23); return x_25; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { lean_object* x_14; -x_14 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_14 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_7); return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { uint8_t x_14; uint8_t x_15; uint8_t x_16; lean_object* x_17; x_14 = lean_unbox(x_3); x_15 = lean_unbox(x_4); x_16 = lean_unbox(x_5); -x_17 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__2(x_1, x_2, x_14, x_15, x_16, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_17 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__2(x_1, x_2, x_14, x_15, x_16, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_12); lean_dec_ref(x_11); lean_dec(x_10); @@ -47986,7 +48112,7 @@ lean_dec_ref(x_1); return x_17; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, uint8_t x_8, uint8_t x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, uint8_t x_8, uint8_t x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { lean_object* x_17; lean_object* x_18; lean_object* x_19; @@ -48037,7 +48163,7 @@ x_31 = l_Lean_Elab_Tactic_Do_ProofMode_MGoal_toExpr(x_30); x_32 = lean_box(x_8); x_33 = lean_box(x_9); x_34 = lean_box(x_22); -x_35 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__2___boxed), 13, 5); +x_35 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__2___boxed), 13, 5); lean_closure_set(x_35, 0, x_21); lean_closure_set(x_35, 1, x_31); lean_closure_set(x_35, 2, x_32); @@ -48155,21 +48281,21 @@ return x_51; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { uint8_t x_17; uint8_t x_18; lean_object* x_19; x_17 = lean_unbox(x_8); x_18 = lean_unbox(x_9); -x_19 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_17, x_18, x_10, x_11, x_12, x_13, x_14, x_15); +x_19 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_17, x_18, x_10, x_11, x_12, x_13, x_14, x_15); return x_19; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, uint8_t x_17, uint8_t x_18, lean_object* x_19, lean_object* x_20, uint8_t x_21, lean_object* x_22, uint8_t x_23, lean_object* x_24, lean_object* x_25, lean_object* x_26, lean_object* x_27, lean_object* x_28, lean_object* x_29) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, uint8_t x_17, uint8_t x_18, lean_object* x_19, lean_object* x_20, uint8_t x_21, lean_object* x_22, uint8_t x_23, lean_object* x_24, lean_object* x_25, lean_object* x_26, lean_object* x_27, lean_object* x_28, lean_object* x_29) { _start: { lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_31 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__0___boxed), 13, 12); +x_31 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__0___boxed), 13, 12); lean_closure_set(x_31, 0, x_3); lean_closure_set(x_31, 1, x_4); lean_closure_set(x_31, 2, x_5); @@ -48184,7 +48310,7 @@ lean_closure_set(x_31, 10, x_13); lean_closure_set(x_31, 11, x_14); x_32 = lean_box(x_17); x_33 = lean_box(x_18); -x_34 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__1___boxed), 16, 10); +x_34 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__1___boxed), 16, 10); lean_closure_set(x_34, 0, x_1); lean_closure_set(x_34, 1, x_2); lean_closure_set(x_34, 2, x_31); @@ -48226,7 +48352,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tact { lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_inc_ref(x_23); -x_32 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___lam__4___boxed), 13, 3); +x_32 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___lam__4___boxed), 14, 3); lean_closure_set(x_32, 0, x_1); lean_closure_set(x_32, 1, x_2); lean_closure_set(x_32, 2, x_23); @@ -48237,7 +48363,7 @@ lean_inc(x_28); lean_inc_ref(x_27); lean_inc(x_26); lean_inc_ref(x_25); -x_34 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18(x_4, x_33, x_32, x_5, x_25, x_26, x_27, x_28, x_29, x_30); +x_34 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18(x_4, x_33, x_32, x_5, x_25, x_26, x_27, x_28, x_29, x_30); if (lean_obj_tag(x_34) == 0) { lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; @@ -48261,7 +48387,7 @@ lean_inc(x_39); lean_dec_ref(x_38); x_40 = 0; x_41 = 0; -x_42 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40(x_6, x_23, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_5, x_21, x_22, x_39, x_40, x_35, x_41, x_25, x_26, x_27, x_28, x_29, x_30); +x_42 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41(x_6, x_23, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_5, x_21, x_22, x_39, x_40, x_35, x_41, x_25, x_26, x_27, x_28, x_29, x_30); return x_42; } else @@ -48484,14 +48610,14 @@ lean_dec_ref(x_3); return x_17; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__42___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__43___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { uint8_t x_14; uint8_t x_15; uint8_t x_16; lean_object* x_17; x_14 = lean_unbox(x_3); x_15 = lean_unbox(x_4); x_16 = lean_unbox(x_5); -x_17 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__42___lam__0(x_1, x_2, x_14, x_15, x_16, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_17 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__43___lam__0(x_1, x_2, x_14, x_15, x_16, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_12); lean_dec_ref(x_11); lean_dec(x_10); @@ -48503,7 +48629,7 @@ lean_dec_ref(x_1); return x_17; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__42___lam__1(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__43___lam__1(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; @@ -48517,7 +48643,7 @@ x_16 = 1; x_17 = lean_box(x_15); x_18 = lean_box(x_2); x_19 = lean_box(x_16); -x_20 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__42___lam__0___boxed), 13, 5); +x_20 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__43___lam__0___boxed), 13, 5); lean_closure_set(x_20, 0, x_12); lean_closure_set(x_20, 1, x_14); lean_closure_set(x_20, 2, x_17); @@ -48527,22 +48653,22 @@ x_21 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_20, x_3, x_4, x_6, x_7, x_8, x return x_21; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__42___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__43___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { uint8_t x_11; lean_object* x_12; x_11 = lean_unbox(x_2); -x_12 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__42___lam__1(x_1, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_12 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__43___lam__1(x_1, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_4); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__42(lean_object* x_1, uint8_t x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__43(lean_object* x_1, uint8_t x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { lean_object* x_14; lean_object* x_15; lean_object* x_16; x_14 = lean_box(x_2); -x_15 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__42___lam__1___boxed), 10, 4); +x_15 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__43___lam__1___boxed), 10, 4); lean_closure_set(x_15, 0, x_1); lean_closure_set(x_15, 1, x_14); lean_closure_set(x_15, 2, x_7); @@ -48573,177 +48699,696 @@ return x_19; } } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; -x_11 = l_Lean_Elab_Tactic_Do_rwIfOrMatcher(x_1, x_2, x_6, x_7, x_8, x_9); +x_11 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__1(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Meta_mkFreshExprMVar(x_1, x_2, x_3, x_7, x_8, x_9, x_10); +return x_12; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_2); +x_13 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__1(x_1, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_10); +lean_dec_ref(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +return x_13; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_isExprDefEq(x_1, x_2, x_6, x_7, x_8, x_9); return x_11; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; -x_11 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_5); lean_dec_ref(x_4); lean_dec(x_3); return x_11; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Meta_instantiateMVarsIfMVarApp___redArg(x_1, x_6); +return x_10; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_8); +lean_dec_ref(x_7); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_11; -x_11 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_Lean_Elab_Tactic_Do_rwIfOrMatcher(x_1, x_2, x_6, x_7, x_8, x_9); return x_11; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_17; lean_object* x_18; -lean_inc(x_8); -x_17 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__1___boxed), 10, 2); -lean_closure_set(x_17, 0, x_8); -lean_closure_set(x_17, 1, x_1); -lean_inc(x_15); -lean_inc_ref(x_14); -lean_inc(x_13); -lean_inc_ref(x_12); -lean_inc_ref(x_10); -x_18 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_17, x_10, x_11, x_12, x_13, x_14, x_15); -if (lean_obj_tag(x_18) == 0) +lean_object* x_11; +x_11 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec(x_3); +return x_11; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__0() { +_start: { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -lean_dec_ref(x_18); -x_20 = lean_ctor_get(x_19, 0); -lean_inc_ref(x_20); -x_21 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_ProofMode_MGoal_withNewProg(x_2, x_20); -x_22 = l_Lean_Name_append(x_3, x_7); -x_23 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___boxed), 10, 3); -lean_closure_set(x_23, 0, x_4); -lean_closure_set(x_23, 1, x_21); -lean_closure_set(x_23, 2, x_22); -lean_inc(x_15); -lean_inc_ref(x_14); -lean_inc(x_13); -lean_inc_ref(x_12); -lean_inc(x_11); -lean_inc_ref(x_10); -x_24 = lean_apply_10(x_5, x_8, x_9, x_23, x_10, x_11, x_12, x_13, x_14, x_15, lean_box(0)); -if (lean_obj_tag(x_24) == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("`rwMatcher` failed to rewrite ", 30, 30); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__1() { +_start: { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -lean_dec_ref(x_24); -x_26 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__6___boxed), 10, 2); -lean_closure_set(x_26, 0, x_6); -lean_closure_set(x_26, 1, x_19); -lean_inc(x_15); -lean_inc_ref(x_14); -lean_inc(x_13); -lean_inc_ref(x_12); -lean_inc_ref(x_10); -x_27 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_26, x_10, x_11, x_12, x_13, x_14, x_15); -if (lean_obj_tag(x_27) == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__0; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__2() { +_start: { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -lean_dec_ref(x_27); -x_29 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__8___boxed), 10, 2); -lean_closure_set(x_29, 0, x_28); -lean_closure_set(x_29, 1, x_25); -x_30 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_29, x_10, x_11, x_12, x_13, x_14, x_15); -lean_dec(x_11); -return x_30; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("\nCheck the output of `trace.Elab.Tactic.Do.vcgen.split` for more info and submit a bug report.", 94, 94); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("The alternative type ", 21, 21); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" returned by `splitWith` does not match ", 40, 40); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__6; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(". This is a bug in `mvcgen`.", 28, 28); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__8; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_40; +x_40 = l_Lean_Elab_Tactic_Do_burnOne___redArg(x_12); +if (lean_obj_tag(x_40) == 0) +{ +lean_object* x_41; uint8_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +lean_dec_ref(x_40); +x_41 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_41, 0, x_1); +x_42 = 0; +x_43 = lean_box(0); +x_44 = lean_box(x_42); +x_45 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__1___boxed), 11, 3); +lean_closure_set(x_45, 0, x_41); +lean_closure_set(x_45, 1, x_44); +lean_closure_set(x_45, 2, x_43); +lean_inc(x_16); +lean_inc_ref(x_15); +lean_inc(x_14); +lean_inc_ref(x_13); +lean_inc_ref(x_11); +x_46 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_45, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_46) == 0) +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_47 = lean_ctor_get(x_46, 0); +lean_inc(x_47); +lean_dec_ref(x_46); +lean_inc(x_47); +lean_inc_ref(x_2); +x_48 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_ProofMode_MGoal_withNewProg(x_2, x_47); +x_49 = l_Lean_Elab_Tactic_Do_ProofMode_MGoal_toExpr(x_48); +lean_inc_ref(x_8); +lean_inc_ref(x_49); +x_50 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__3___boxed), 10, 2); +lean_closure_set(x_50, 0, x_49); +lean_closure_set(x_50, 1, x_8); +lean_inc(x_16); +lean_inc_ref(x_15); +lean_inc(x_14); +lean_inc_ref(x_13); +lean_inc_ref(x_11); +x_51 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_50, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_51) == 0) +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_82; +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +lean_dec_ref(x_51); +x_53 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__5___boxed), 9, 1); +lean_closure_set(x_53, 0, x_47); +x_82 = lean_unbox(x_52); +lean_dec(x_52); +if (x_82 == 0) +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_83 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__5; +x_84 = l_Lean_MessageData_ofExpr(x_8); +x_85 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_85, 0, x_83); +lean_ctor_set(x_85, 1, x_84); +x_86 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__7; +x_87 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_87, 0, x_85); +lean_ctor_set(x_87, 1, x_86); +x_88 = l_Lean_MessageData_ofExpr(x_49); +x_89 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_89, 0, x_87); +lean_ctor_set(x_89, 1, x_88); +x_90 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__9; +x_91 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_91, 0, x_89); +lean_ctor_set(x_91, 1, x_90); +lean_inc(x_16); +lean_inc_ref(x_15); +lean_inc(x_14); +lean_inc_ref(x_13); +lean_inc_ref(x_11); +x_92 = l_Lean_throwError___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__0___redArg(x_91, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_92) == 0) +{ +lean_dec_ref(x_92); +x_54 = x_11; +x_55 = x_12; +x_56 = x_13; +x_57 = x_14; +x_58 = x_15; +x_59 = x_16; +x_60 = lean_box(0); +goto block_81; } else { -uint8_t x_31; -lean_dec(x_25); -lean_dec(x_15); -lean_dec_ref(x_14); -lean_dec(x_13); -lean_dec_ref(x_12); -lean_dec(x_11); +uint8_t x_93; +lean_dec_ref(x_53); +lean_dec(x_16); +lean_dec_ref(x_15); +lean_dec(x_14); +lean_dec_ref(x_13); +lean_dec(x_12); +lean_dec_ref(x_11); lean_dec_ref(x_10); -x_31 = !lean_is_exclusive(x_27); -if (x_31 == 0) +lean_dec(x_9); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec_ref(x_5); +lean_dec_ref(x_4); +lean_dec(x_3); +lean_dec_ref(x_2); +x_93 = !lean_is_exclusive(x_92); +if (x_93 == 0) { -return x_27; +return x_92; } else { -lean_object* x_32; lean_object* x_33; -x_32 = lean_ctor_get(x_27, 0); -lean_inc(x_32); -lean_dec(x_27); -x_33 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_33, 0, x_32); -return x_33; +lean_object* x_94; lean_object* x_95; +x_94 = lean_ctor_get(x_92, 0); +lean_inc(x_94); +lean_dec(x_92); +x_95 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_95, 0, x_94); +return x_95; } } } else { -lean_dec(x_19); -lean_dec(x_15); -lean_dec_ref(x_14); -lean_dec(x_13); -lean_dec_ref(x_12); -lean_dec(x_11); +lean_dec_ref(x_49); +lean_dec_ref(x_8); +x_54 = x_11; +x_55 = x_12; +x_56 = x_13; +x_57 = x_14; +x_58 = x_15; +x_59 = x_16; +x_60 = lean_box(0); +goto block_81; +} +block_81: +{ +lean_object* x_61; +lean_inc(x_59); +lean_inc_ref(x_58); +lean_inc(x_57); +lean_inc_ref(x_56); +lean_inc_ref(x_54); +x_61 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_53, x_54, x_55, x_56, x_57, x_58, x_59); +if (lean_obj_tag(x_61) == 0) +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_61, 0); +lean_inc(x_62); +lean_dec_ref(x_61); +lean_inc(x_62); +lean_inc(x_9); +x_63 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__4___boxed), 10, 2); +lean_closure_set(x_63, 0, x_9); +lean_closure_set(x_63, 1, x_62); +lean_inc(x_59); +lean_inc_ref(x_58); +lean_inc(x_57); +lean_inc_ref(x_56); +lean_inc_ref(x_54); +x_64 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_63, x_54, x_55, x_56, x_57, x_58, x_59); +if (lean_obj_tag(x_64) == 0) +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; +x_65 = lean_ctor_get(x_64, 0); +lean_inc(x_65); +lean_dec_ref(x_64); +x_66 = lean_ctor_get(x_65, 0); +lean_inc_ref(x_66); +x_67 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__6___boxed), 10, 2); +lean_closure_set(x_67, 0, x_6); +lean_closure_set(x_67, 1, x_65); +x_68 = lean_expr_eqv(x_66, x_62); +if (x_68 == 0) +{ +lean_dec(x_62); +x_18 = x_67; +x_19 = x_66; +x_20 = x_54; +x_21 = x_55; +x_22 = x_56; +x_23 = x_57; +x_24 = x_58; +x_25 = x_59; +x_26 = lean_box(0); +goto block_39; +} +else +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_69 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__1; +x_70 = l_Lean_indentExpr(x_62); +x_71 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_71, 0, x_69); +lean_ctor_set(x_71, 1, x_70); +x_72 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__3; +x_73 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_73, 0, x_71); +lean_ctor_set(x_73, 1, x_72); +lean_inc(x_59); +lean_inc_ref(x_58); +lean_inc(x_57); +lean_inc_ref(x_56); +lean_inc_ref(x_54); +x_74 = l_Lean_throwError___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__0___redArg(x_73, x_54, x_55, x_56, x_57, x_58, x_59); +if (lean_obj_tag(x_74) == 0) +{ +lean_dec_ref(x_74); +x_18 = x_67; +x_19 = x_66; +x_20 = x_54; +x_21 = x_55; +x_22 = x_56; +x_23 = x_57; +x_24 = x_58; +x_25 = x_59; +x_26 = lean_box(0); +goto block_39; +} +else +{ +uint8_t x_75; +lean_dec_ref(x_67); +lean_dec_ref(x_66); +lean_dec(x_59); +lean_dec_ref(x_58); +lean_dec(x_57); +lean_dec_ref(x_56); +lean_dec(x_55); +lean_dec_ref(x_54); +lean_dec_ref(x_10); +lean_dec(x_9); +lean_dec(x_7); +lean_dec_ref(x_5); +lean_dec_ref(x_4); +lean_dec(x_3); +lean_dec_ref(x_2); +x_75 = !lean_is_exclusive(x_74); +if (x_75 == 0) +{ +return x_74; +} +else +{ +lean_object* x_76; lean_object* x_77; +x_76 = lean_ctor_get(x_74, 0); +lean_inc(x_76); +lean_dec(x_74); +x_77 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_77, 0, x_76); +return x_77; +} +} +} +} +else +{ +uint8_t x_78; +lean_dec(x_62); +lean_dec(x_59); +lean_dec_ref(x_58); +lean_dec(x_57); +lean_dec_ref(x_56); +lean_dec(x_55); +lean_dec_ref(x_54); lean_dec_ref(x_10); +lean_dec(x_9); +lean_dec(x_7); lean_dec_ref(x_6); -return x_24; +lean_dec_ref(x_5); +lean_dec_ref(x_4); +lean_dec(x_3); +lean_dec_ref(x_2); +x_78 = !lean_is_exclusive(x_64); +if (x_78 == 0) +{ +return x_64; +} +else +{ +lean_object* x_79; lean_object* x_80; +x_79 = lean_ctor_get(x_64, 0); +lean_inc(x_79); +lean_dec(x_64); +x_80 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_80, 0, x_79); +return x_80; +} } } else { -uint8_t x_34; -lean_dec(x_15); -lean_dec_ref(x_14); -lean_dec(x_13); -lean_dec_ref(x_12); -lean_dec(x_11); +lean_dec(x_59); +lean_dec_ref(x_58); +lean_dec(x_57); +lean_dec_ref(x_56); +lean_dec(x_55); +lean_dec_ref(x_54); lean_dec_ref(x_10); -lean_dec_ref(x_9); -lean_dec(x_8); +lean_dec(x_9); lean_dec(x_7); lean_dec_ref(x_6); lean_dec_ref(x_5); lean_dec_ref(x_4); lean_dec(x_3); lean_dec_ref(x_2); -x_34 = !lean_is_exclusive(x_18); -if (x_34 == 0) +return x_61; +} +} +} +else { -return x_18; +uint8_t x_96; +lean_dec_ref(x_49); +lean_dec(x_47); +lean_dec(x_16); +lean_dec_ref(x_15); +lean_dec(x_14); +lean_dec_ref(x_13); +lean_dec(x_12); +lean_dec_ref(x_11); +lean_dec_ref(x_10); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec_ref(x_5); +lean_dec_ref(x_4); +lean_dec(x_3); +lean_dec_ref(x_2); +x_96 = !lean_is_exclusive(x_51); +if (x_96 == 0) +{ +return x_51; } else { -lean_object* x_35; lean_object* x_36; -x_35 = lean_ctor_get(x_18, 0); -lean_inc(x_35); -lean_dec(x_18); -x_36 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_36, 0, x_35); -return x_36; +lean_object* x_97; lean_object* x_98; +x_97 = lean_ctor_get(x_51, 0); +lean_inc(x_97); +lean_dec(x_51); +x_98 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_98, 0, x_97); +return x_98; } } } +else +{ +lean_dec(x_16); +lean_dec_ref(x_15); +lean_dec(x_14); +lean_dec_ref(x_13); +lean_dec(x_12); +lean_dec_ref(x_11); +lean_dec_ref(x_10); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec_ref(x_5); +lean_dec_ref(x_4); +lean_dec(x_3); +lean_dec_ref(x_2); +return x_46; +} } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +else +{ +uint8_t x_99; +lean_dec(x_16); +lean_dec_ref(x_15); +lean_dec(x_14); +lean_dec_ref(x_13); +lean_dec(x_12); +lean_dec_ref(x_11); +lean_dec_ref(x_10); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec_ref(x_5); +lean_dec_ref(x_4); +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_99 = !lean_is_exclusive(x_40); +if (x_99 == 0) +{ +return x_40; +} +else +{ +lean_object* x_100; lean_object* x_101; +x_100 = lean_ctor_get(x_40, 0); +lean_inc(x_100); +lean_dec(x_40); +x_101 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_101, 0, x_100); +return x_101; +} +} +block_39: +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_27 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_ProofMode_MGoal_withNewProg(x_2, x_19); +x_28 = l_Lean_Name_append(x_3, x_7); +x_29 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___boxed), 10, 3); +lean_closure_set(x_29, 0, x_4); +lean_closure_set(x_29, 1, x_27); +lean_closure_set(x_29, 2, x_28); +lean_inc(x_25); +lean_inc_ref(x_24); +lean_inc(x_23); +lean_inc_ref(x_22); +lean_inc(x_21); +lean_inc_ref(x_20); +x_30 = lean_apply_10(x_5, x_9, x_10, x_29, x_20, x_21, x_22, x_23, x_24, x_25, lean_box(0)); +if (lean_obj_tag(x_30) == 0) +{ +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +lean_dec_ref(x_30); +lean_inc(x_25); +lean_inc_ref(x_24); +lean_inc(x_23); +lean_inc_ref(x_22); +lean_inc_ref(x_20); +x_32 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_18, x_20, x_21, x_22, x_23, x_24, x_25); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +lean_dec_ref(x_32); +x_34 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__8___boxed), 10, 2); +lean_closure_set(x_34, 0, x_33); +lean_closure_set(x_34, 1, x_31); +x_35 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_34, x_20, x_21, x_22, x_23, x_24, x_25); +lean_dec(x_21); +return x_35; +} +else +{ +uint8_t x_36; +lean_dec(x_31); +lean_dec(x_25); +lean_dec_ref(x_24); +lean_dec(x_23); +lean_dec_ref(x_22); +lean_dec(x_21); +lean_dec_ref(x_20); +x_36 = !lean_is_exclusive(x_32); +if (x_36 == 0) +{ +return x_32; +} +else +{ +lean_object* x_37; lean_object* x_38; +x_37 = lean_ctor_get(x_32, 0); +lean_inc(x_37); +lean_dec(x_32); +x_38 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_38, 0, x_37); +return x_38; +} +} +} +else +{ +lean_dec(x_25); +lean_dec_ref(x_24); +lean_dec(x_23); +lean_dec_ref(x_22); +lean_dec(x_21); +lean_dec_ref(x_20); +lean_dec_ref(x_18); +return x_30; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; _start: { -lean_object* x_17; -x_17 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -return x_17; +lean_object* x_18; +x_18 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +return x_18; } } LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { @@ -48987,7 +49632,7 @@ lean_inc_ref(x_6); x_50 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_49, x_6, x_7, x_8, x_9, x_10, x_11); if (lean_obj_tag(x_50) == 0) { -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; x_51 = lean_ctor_get(x_50, 0); lean_inc(x_51); lean_dec_ref(x_50); @@ -48999,117 +49644,117 @@ lean_inc_ref(x_53); lean_dec_ref(x_43); if (lean_obj_tag(x_51) == 1) { -lean_object* x_76; lean_object* x_77; +lean_object* x_75; lean_object* x_76; lean_dec_ref(x_48); lean_dec_ref(x_5); lean_dec_ref(x_3); -x_76 = lean_ctor_get(x_51, 0); -lean_inc(x_76); +x_75 = lean_ctor_get(x_51, 0); +lean_inc(x_75); lean_dec_ref(x_51); -x_77 = l_Lean_Elab_Tactic_Do_burnOne___redArg(x_7); -if (lean_obj_tag(x_77) == 0) +x_76 = l_Lean_Elab_Tactic_Do_burnOne___redArg(x_7); +if (lean_obj_tag(x_76) == 0) { -lean_object* x_78; lean_object* x_79; lean_object* x_80; -lean_dec_ref(x_77); -x_78 = lean_ctor_get(x_76, 0); -lean_inc_ref(x_78); +lean_object* x_77; lean_object* x_78; lean_object* x_79; +lean_dec_ref(x_76); +x_77 = lean_ctor_get(x_75, 0); +lean_inc_ref(x_77); lean_inc_ref(x_2); -x_79 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_ProofMode_MGoal_withNewProg(x_2, x_78); +x_78 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_ProofMode_MGoal_withNewProg(x_2, x_77); lean_inc(x_11); lean_inc_ref(x_10); lean_inc(x_9); lean_inc_ref(x_8); lean_inc(x_7); lean_inc_ref(x_6); -x_80 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp(x_1, x_79, x_4, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_80) == 0) +x_79 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp(x_1, x_78, x_4, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_79) == 0) { -lean_object* x_81; lean_object* x_82; uint8_t x_83; lean_object* x_84; uint8_t x_85; lean_object* x_86; -x_81 = lean_ctor_get(x_80, 0); -lean_inc(x_81); -lean_dec_ref(x_80); -x_82 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__9___closed__1; -x_83 = 0; -x_84 = l_Lean_Expr_app___override(x_53, x_52); -x_85 = 0; +lean_object* x_80; lean_object* x_81; uint8_t x_82; lean_object* x_83; uint8_t x_84; lean_object* x_85; +x_80 = lean_ctor_get(x_79, 0); +lean_inc(x_80); +lean_dec_ref(x_79); +x_81 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__9___closed__1; +x_82 = 0; +x_83 = l_Lean_Expr_app___override(x_53, x_52); +x_84 = 0; lean_inc(x_11); lean_inc_ref(x_10); lean_inc(x_9); lean_inc_ref(x_8); lean_inc(x_7); lean_inc_ref(x_6); -x_86 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__42(x_2, x_47, x_82, x_83, x_84, x_85, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_86) == 0) +x_85 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__43(x_2, x_47, x_81, x_82, x_83, x_84, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_85) == 0) { -lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_87 = lean_ctor_get(x_86, 0); -lean_inc(x_87); -lean_dec_ref(x_86); -x_88 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__0___boxed), 10, 2); -lean_closure_set(x_88, 0, x_87); -lean_closure_set(x_88, 1, x_76); +lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_86 = lean_ctor_get(x_85, 0); +lean_inc(x_86); +lean_dec_ref(x_85); +x_87 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__0___boxed), 10, 2); +lean_closure_set(x_87, 0, x_86); +lean_closure_set(x_87, 1, x_75); lean_inc(x_11); lean_inc_ref(x_10); lean_inc(x_9); lean_inc_ref(x_8); lean_inc_ref(x_6); -x_89 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_88, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_89) == 0) +x_88 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_87, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_88) == 0) { -lean_object* x_90; lean_object* x_91; lean_object* x_92; -x_90 = lean_ctor_get(x_89, 0); -lean_inc(x_90); -lean_dec_ref(x_89); -x_91 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__8___boxed), 10, 2); -lean_closure_set(x_91, 0, x_90); -lean_closure_set(x_91, 1, x_81); -x_92 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_91, x_6, x_7, x_8, x_9, x_10, x_11); +lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_89 = lean_ctor_get(x_88, 0); +lean_inc(x_89); +lean_dec_ref(x_88); +x_90 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__8___boxed), 10, 2); +lean_closure_set(x_90, 0, x_89); +lean_closure_set(x_90, 1, x_80); +x_91 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_90, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_7); -return x_92; +return x_91; } else { -uint8_t x_93; -lean_dec(x_81); +uint8_t x_92; +lean_dec(x_80); lean_dec(x_11); lean_dec_ref(x_10); lean_dec(x_9); lean_dec_ref(x_8); lean_dec(x_7); lean_dec_ref(x_6); -x_93 = !lean_is_exclusive(x_89); -if (x_93 == 0) +x_92 = !lean_is_exclusive(x_88); +if (x_92 == 0) { -return x_89; +return x_88; } else { -lean_object* x_94; lean_object* x_95; -x_94 = lean_ctor_get(x_89, 0); -lean_inc(x_94); -lean_dec(x_89); -x_95 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_95, 0, x_94); -return x_95; +lean_object* x_93; lean_object* x_94; +x_93 = lean_ctor_get(x_88, 0); +lean_inc(x_93); +lean_dec(x_88); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_93); +return x_94; } } } else { -lean_dec(x_81); -lean_dec(x_76); +lean_dec(x_80); +lean_dec(x_75); lean_dec(x_11); lean_dec_ref(x_10); lean_dec(x_9); lean_dec_ref(x_8); lean_dec(x_7); lean_dec_ref(x_6); -return x_86; +return x_85; } } else { -lean_dec(x_76); +lean_dec(x_75); lean_dec_ref(x_53); lean_dec_ref(x_52); lean_dec(x_11); @@ -49119,13 +49764,13 @@ lean_dec_ref(x_8); lean_dec(x_7); lean_dec_ref(x_6); lean_dec_ref(x_2); -return x_80; +return x_79; } } else { -uint8_t x_96; -lean_dec(x_76); +uint8_t x_95; +lean_dec(x_75); lean_dec_ref(x_53); lean_dec_ref(x_52); lean_dec(x_11); @@ -49137,67 +49782,67 @@ lean_dec_ref(x_6); lean_dec(x_4); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_96 = !lean_is_exclusive(x_77); -if (x_96 == 0) +x_95 = !lean_is_exclusive(x_76); +if (x_95 == 0) { -return x_77; +return x_76; } else { -lean_object* x_97; lean_object* x_98; -x_97 = lean_ctor_get(x_77, 0); -lean_inc(x_97); -lean_dec(x_77); -x_98 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_98, 0, x_97); -return x_98; +lean_object* x_96; lean_object* x_97; +x_96 = lean_ctor_get(x_76, 0); +lean_inc(x_96); +lean_dec(x_76); +x_97 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_97, 0, x_96); +return x_97; } } } else { -lean_object* x_99; lean_object* x_100; +lean_object* x_98; lean_object* x_99; lean_dec(x_51); lean_inc_ref(x_48); -x_99 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__2___boxed), 9, 1); -lean_closure_set(x_99, 0, x_48); +x_98 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__2___boxed), 9, 1); +lean_closure_set(x_98, 0, x_48); lean_inc(x_11); lean_inc_ref(x_10); lean_inc(x_9); lean_inc_ref(x_8); lean_inc_ref(x_6); -x_100 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_99, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_100) == 0) +x_99 = l_Lean_Elab_Tactic_Do_liftSimpM___redArg(x_98, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_99) == 0) { -lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_101 = lean_ctor_get(x_100, 0); -lean_inc(x_101); -lean_dec_ref(x_100); -switch (lean_obj_tag(x_101)) { +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_100 = lean_ctor_get(x_99, 0); +lean_inc(x_100); +lean_dec_ref(x_99); +switch (lean_obj_tag(x_100)) { case 0: { -lean_object* x_124; lean_object* x_125; +lean_object* x_123; lean_object* x_124; lean_dec_ref(x_53); lean_dec_ref(x_52); lean_dec_ref(x_48); lean_dec_ref(x_5); lean_dec_ref(x_3); -x_124 = lean_ctor_get(x_101, 0); -lean_inc_ref(x_124); -lean_dec_ref(x_101); -x_125 = l_Lean_Elab_Tactic_Do_burnOne___redArg(x_7); -if (lean_obj_tag(x_125) == 0) +x_123 = lean_ctor_get(x_100, 0); +lean_inc_ref(x_123); +lean_dec_ref(x_100); +x_124 = l_Lean_Elab_Tactic_Do_burnOne___redArg(x_7); +if (lean_obj_tag(x_124) == 0) { -lean_object* x_126; lean_object* x_127; -lean_dec_ref(x_125); -x_126 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_ProofMode_MGoal_withNewProg(x_2, x_124); -x_127 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp(x_1, x_126, x_4, x_6, x_7, x_8, x_9, x_10, x_11); -return x_127; +lean_object* x_125; lean_object* x_126; +lean_dec_ref(x_124); +x_125 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_ProofMode_MGoal_withNewProg(x_2, x_123); +x_126 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp(x_1, x_125, x_4, x_6, x_7, x_8, x_9, x_10, x_11); +return x_126; } else { -uint8_t x_128; -lean_dec_ref(x_124); +uint8_t x_127; +lean_dec_ref(x_123); lean_dec(x_11); lean_dec_ref(x_10); lean_dec(x_9); @@ -49207,115 +49852,112 @@ lean_dec_ref(x_6); lean_dec(x_4); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_128 = !lean_is_exclusive(x_125); -if (x_128 == 0) +x_127 = !lean_is_exclusive(x_124); +if (x_127 == 0) { -return x_125; +return x_124; } else { -lean_object* x_129; lean_object* x_130; -x_129 = lean_ctor_get(x_125, 0); -lean_inc(x_129); -lean_dec(x_125); -x_130 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_130, 0, x_129); -return x_130; +lean_object* x_128; lean_object* x_129; +x_128 = lean_ctor_get(x_124, 0); +lean_inc(x_128); +lean_dec(x_124); +x_129 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_129, 0, x_128); +return x_129; } } } case 1: { -lean_dec_ref(x_101); -x_102 = x_48; -x_103 = x_6; -x_104 = x_7; -x_105 = x_8; -x_106 = x_9; -x_107 = x_10; -x_108 = x_11; -goto block_123; +lean_dec_ref(x_100); +x_101 = x_48; +x_102 = x_6; +x_103 = x_7; +x_104 = x_8; +x_105 = x_9; +x_106 = x_10; +x_107 = x_11; +goto block_122; } default: { -lean_dec(x_101); -x_102 = x_48; -x_103 = x_6; -x_104 = x_7; -x_105 = x_8; -x_106 = x_9; -x_107 = x_10; -x_108 = x_11; -goto block_123; +lean_dec(x_100); +x_101 = x_48; +x_102 = x_6; +x_103 = x_7; +x_104 = x_8; +x_105 = x_9; +x_106 = x_10; +x_107 = x_11; +goto block_122; } } -block_123: +block_122: { -lean_object* x_109; lean_object* x_110; -x_109 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__17___closed__0; -lean_inc(x_108); -lean_inc_ref(x_107); -lean_inc(x_106); -lean_inc_ref(x_105); -lean_inc_ref(x_103); -x_110 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__3(x_109, x_103, x_104, x_105, x_106, x_107, x_108); -if (lean_obj_tag(x_110) == 0) +lean_object* x_108; lean_object* x_109; +x_108 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__17___closed__0; +lean_inc(x_107); +lean_inc_ref(x_106); +lean_inc(x_105); +lean_inc_ref(x_104); +lean_inc_ref(x_102); +x_109 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__3(x_108, x_102, x_103, x_104, x_105, x_106, x_107); +if (lean_obj_tag(x_109) == 0) { -lean_object* x_111; uint8_t x_112; -x_111 = lean_ctor_get(x_110, 0); -lean_inc(x_111); -lean_dec_ref(x_110); -x_112 = lean_unbox(x_111); -lean_dec(x_111); -if (x_112 == 0) +lean_object* x_110; uint8_t x_111; +x_110 = lean_ctor_get(x_109, 0); +lean_inc(x_110); +lean_dec_ref(x_109); +x_111 = lean_unbox(x_110); +lean_dec(x_110); +if (x_111 == 0) { +lean_dec_ref(x_101); x_54 = x_102; x_55 = x_103; x_56 = x_104; x_57 = x_105; x_58 = x_106; x_59 = x_107; -x_60 = x_108; -x_61 = lean_box(0); -goto block_75; +x_60 = lean_box(0); +goto block_74; } else { -lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; -x_113 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___closed__1; +lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; +x_112 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___closed__1; +x_113 = l_Lean_MessageData_ofExpr(x_101); +x_114 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_114, 0, x_112); +lean_ctor_set(x_114, 1, x_113); +lean_inc(x_107); +lean_inc_ref(x_106); +lean_inc(x_105); +lean_inc_ref(x_104); lean_inc_ref(x_102); -x_114 = l_Lean_MessageData_ofExpr(x_102); -x_115 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_115, 0, x_113); -lean_ctor_set(x_115, 1, x_114); -lean_inc(x_108); -lean_inc_ref(x_107); -lean_inc(x_106); -lean_inc_ref(x_105); -lean_inc_ref(x_103); -x_116 = l_Lean_addTrace___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__5(x_109, x_115, x_103, x_104, x_105, x_106, x_107, x_108); -if (lean_obj_tag(x_116) == 0) +x_115 = l_Lean_addTrace___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__5(x_108, x_114, x_102, x_103, x_104, x_105, x_106, x_107); +if (lean_obj_tag(x_115) == 0) { -lean_dec_ref(x_116); +lean_dec_ref(x_115); x_54 = x_102; x_55 = x_103; x_56 = x_104; x_57 = x_105; x_58 = x_106; x_59 = x_107; -x_60 = x_108; -x_61 = lean_box(0); -goto block_75; +x_60 = lean_box(0); +goto block_74; } else { -uint8_t x_117; -lean_dec(x_108); -lean_dec_ref(x_107); -lean_dec(x_106); -lean_dec_ref(x_105); -lean_dec(x_104); -lean_dec_ref(x_103); +uint8_t x_116; +lean_dec(x_107); +lean_dec_ref(x_106); +lean_dec(x_105); +lean_dec_ref(x_104); +lean_dec(x_103); lean_dec_ref(x_102); lean_dec_ref(x_53); lean_dec_ref(x_52); @@ -49324,34 +49966,34 @@ lean_dec(x_4); lean_dec_ref(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_117 = !lean_is_exclusive(x_116); -if (x_117 == 0) +x_116 = !lean_is_exclusive(x_115); +if (x_116 == 0) { -return x_116; +return x_115; } else { -lean_object* x_118; lean_object* x_119; -x_118 = lean_ctor_get(x_116, 0); -lean_inc(x_118); -lean_dec(x_116); -x_119 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_119, 0, x_118); -return x_119; +lean_object* x_117; lean_object* x_118; +x_117 = lean_ctor_get(x_115, 0); +lean_inc(x_117); +lean_dec(x_115); +x_118 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_118, 0, x_117); +return x_118; } } } } else { -uint8_t x_120; -lean_dec(x_108); -lean_dec_ref(x_107); -lean_dec(x_106); -lean_dec_ref(x_105); -lean_dec(x_104); -lean_dec_ref(x_103); +uint8_t x_119; +lean_dec(x_107); +lean_dec_ref(x_106); +lean_dec(x_105); +lean_dec_ref(x_104); +lean_dec(x_103); lean_dec_ref(x_102); +lean_dec_ref(x_101); lean_dec_ref(x_53); lean_dec_ref(x_52); lean_dec_ref(x_5); @@ -49359,27 +50001,27 @@ lean_dec(x_4); lean_dec_ref(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_120 = !lean_is_exclusive(x_110); -if (x_120 == 0) +x_119 = !lean_is_exclusive(x_109); +if (x_119 == 0) { -return x_110; +return x_109; } else { -lean_object* x_121; lean_object* x_122; -x_121 = lean_ctor_get(x_110, 0); -lean_inc(x_121); -lean_dec(x_110); -x_122 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_122, 0, x_121); -return x_122; +lean_object* x_120; lean_object* x_121; +x_120 = lean_ctor_get(x_109, 0); +lean_inc(x_120); +lean_dec(x_109); +x_121 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_121, 0, x_120); +return x_121; } } } } else { -uint8_t x_131; +uint8_t x_130; lean_dec_ref(x_53); lean_dec_ref(x_52); lean_dec_ref(x_48); @@ -49394,87 +50036,87 @@ lean_dec(x_4); lean_dec_ref(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_131 = !lean_is_exclusive(x_100); -if (x_131 == 0) +x_130 = !lean_is_exclusive(x_99); +if (x_130 == 0) { -return x_100; +return x_99; } else { -lean_object* x_132; lean_object* x_133; -x_132 = lean_ctor_get(x_100, 0); -lean_inc(x_132); -lean_dec(x_100); -x_133 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_133, 0, x_132); -return x_133; +lean_object* x_131; lean_object* x_132; +x_131 = lean_ctor_get(x_99, 0); +lean_inc(x_131); +lean_dec(x_99); +x_132 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_132, 0, x_131); +return x_132; } } } -block_75: +block_74: { -lean_object* x_62; -x_62 = l_Lean_Elab_Tactic_Do_burnOne___redArg(x_56); -if (lean_obj_tag(x_62) == 0) +lean_object* x_61; +x_61 = l_Lean_Elab_Tactic_Do_burnOne___redArg(x_55); +if (lean_obj_tag(x_61) == 0) { -lean_object* x_63; uint8_t x_64; lean_object* x_65; uint8_t x_66; lean_object* x_67; -lean_dec_ref(x_62); -x_63 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__9___closed__1; -x_64 = 0; -x_65 = l_Lean_Expr_app___override(x_53, x_52); -x_66 = 0; -lean_inc(x_60); -lean_inc_ref(x_59); -lean_inc(x_58); -lean_inc_ref(x_57); -lean_inc(x_56); -lean_inc_ref(x_55); +lean_object* x_62; uint8_t x_63; lean_object* x_64; uint8_t x_65; lean_object* x_66; +lean_dec_ref(x_61); +x_62 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__9___closed__1; +x_63 = 0; +x_64 = l_Lean_Expr_app___override(x_53, x_52); +x_65 = 0; +lean_inc(x_59); +lean_inc_ref(x_58); +lean_inc(x_57); +lean_inc_ref(x_56); +lean_inc(x_55); +lean_inc_ref(x_54); +lean_inc_ref(x_64); lean_inc_ref(x_2); -x_67 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__42(x_2, x_47, x_63, x_64, x_65, x_66, x_55, x_56, x_57, x_58, x_59, x_60); -if (lean_obj_tag(x_67) == 0) +x_66 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__43(x_2, x_47, x_62, x_63, x_64, x_65, x_54, x_55, x_56, x_57, x_58, x_59); +if (lean_obj_tag(x_66) == 0) { -lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_68 = lean_ctor_get(x_67, 0); -lean_inc(x_68); -lean_dec_ref(x_67); +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +lean_dec_ref(x_66); lean_inc_ref(x_2); -x_69 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__5___boxed), 16, 6); -lean_closure_set(x_69, 0, x_54); -lean_closure_set(x_69, 1, x_2); -lean_closure_set(x_69, 2, x_4); -lean_closure_set(x_69, 3, x_1); -lean_closure_set(x_69, 4, x_5); -lean_closure_set(x_69, 5, x_68); -x_70 = l_Lean_Elab_Tactic_Do_ProofMode_MGoal_toExpr(x_2); -x_71 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18(x_3, x_70, x_69, x_47, x_55, x_56, x_57, x_58, x_59, x_60); -return x_71; +x_68 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___boxed), 17, 6); +lean_closure_set(x_68, 0, x_64); +lean_closure_set(x_68, 1, x_2); +lean_closure_set(x_68, 2, x_4); +lean_closure_set(x_68, 3, x_1); +lean_closure_set(x_68, 4, x_5); +lean_closure_set(x_68, 5, x_67); +x_69 = l_Lean_Elab_Tactic_Do_ProofMode_MGoal_toExpr(x_2); +x_70 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18(x_3, x_69, x_68, x_47, x_54, x_55, x_56, x_57, x_58, x_59); +return x_70; } else { -lean_dec(x_60); -lean_dec_ref(x_59); -lean_dec(x_58); -lean_dec_ref(x_57); -lean_dec(x_56); -lean_dec_ref(x_55); +lean_dec_ref(x_64); +lean_dec(x_59); +lean_dec_ref(x_58); +lean_dec(x_57); +lean_dec_ref(x_56); +lean_dec(x_55); lean_dec_ref(x_54); lean_dec_ref(x_5); lean_dec(x_4); lean_dec_ref(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -return x_67; +return x_66; } } else { -uint8_t x_72; -lean_dec(x_60); -lean_dec_ref(x_59); -lean_dec(x_58); -lean_dec_ref(x_57); -lean_dec(x_56); -lean_dec_ref(x_55); +uint8_t x_71; +lean_dec(x_59); +lean_dec_ref(x_58); +lean_dec(x_57); +lean_dec_ref(x_56); +lean_dec(x_55); lean_dec_ref(x_54); lean_dec_ref(x_53); lean_dec_ref(x_52); @@ -49483,27 +50125,27 @@ lean_dec(x_4); lean_dec_ref(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_72 = !lean_is_exclusive(x_62); -if (x_72 == 0) +x_71 = !lean_is_exclusive(x_61); +if (x_71 == 0) { -return x_62; +return x_61; } else { -lean_object* x_73; lean_object* x_74; -x_73 = lean_ctor_get(x_62, 0); -lean_inc(x_73); -lean_dec(x_62); -x_74 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_74, 0, x_73); -return x_74; +lean_object* x_72; lean_object* x_73; +x_72 = lean_ctor_get(x_61, 0); +lean_inc(x_72); +lean_dec(x_61); +x_73 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_73, 0, x_72); +return x_73; } } } } else { -uint8_t x_134; +uint8_t x_133; lean_dec_ref(x_48); lean_dec_ref(x_43); lean_dec_ref(x_37); @@ -49518,20 +50160,20 @@ lean_dec(x_4); lean_dec_ref(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_134 = !lean_is_exclusive(x_50); -if (x_134 == 0) +x_133 = !lean_is_exclusive(x_50); +if (x_133 == 0) { return x_50; } else { -lean_object* x_135; lean_object* x_136; -x_135 = lean_ctor_get(x_50, 0); -lean_inc(x_135); +lean_object* x_134; lean_object* x_135; +x_134 = lean_ctor_get(x_50, 0); +lean_inc(x_134); lean_dec(x_50); -x_136 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_136, 0, x_135); -return x_136; +x_135 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_135, 0, x_134); +return x_135; } } } @@ -52407,8 +53049,8 @@ x_36 = l_Lean_MessageData_ofList(x_35); x_37 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_37, 0, x_29); lean_ctor_set(x_37, 1, x_36); -x_38 = l_Lean_addTrace___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__5(x_8, x_37, x_24, x_26, x_22, x_20, x_25, x_23); -lean_dec(x_26); +x_38 = l_Lean_addTrace___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__5(x_8, x_37, x_25, x_21, x_22, x_24, x_20, x_23); +lean_dec(x_21); if (lean_obj_tag(x_38) == 0) { uint8_t x_39; @@ -52418,7 +53060,7 @@ if (x_39 == 0) lean_object* x_40; x_40 = lean_ctor_get(x_38, 0); lean_dec(x_40); -lean_ctor_set(x_38, 0, x_21); +lean_ctor_set(x_38, 0, x_26); return x_38; } else @@ -52426,14 +53068,14 @@ else lean_object* x_41; lean_dec(x_38); x_41 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_41, 0, x_21); +lean_ctor_set(x_41, 0, x_26); return x_41; } } else { uint8_t x_42; -lean_dec_ref(x_21); +lean_dec_ref(x_26); x_42 = !lean_is_exclusive(x_38); if (x_42 == 0) { @@ -52460,12 +53102,12 @@ x_54 = lean_ctor_get(x_53, 0); lean_inc(x_54); lean_dec_ref(x_53); x_20 = x_46; -x_21 = x_48; -x_22 = x_47; +x_21 = x_47; +x_22 = x_48; x_23 = x_49; x_24 = x_50; -x_25 = x_52; -x_26 = x_51; +x_25 = x_51; +x_26 = x_52; x_27 = x_54; x_28 = lean_box(0); goto block_45; @@ -52474,12 +53116,12 @@ else { uint8_t x_55; lean_dec_ref(x_52); -lean_dec(x_51); -lean_dec_ref(x_50); +lean_dec_ref(x_51); +lean_dec(x_50); lean_dec(x_49); lean_dec_ref(x_48); -lean_dec_ref(x_47); -lean_dec(x_46); +lean_dec(x_47); +lean_dec_ref(x_46); lean_dec(x_8); x_55 = !lean_is_exclusive(x_53); if (x_55 == 0) @@ -52513,24 +53155,24 @@ lean_inc(x_69); lean_dec(x_67); lean_inc(x_69); x_70 = lean_array_to_list(x_69); -lean_inc(x_61); -lean_inc_ref(x_64); -lean_inc(x_59); -lean_inc_ref(x_60); +lean_inc(x_62); +lean_inc_ref(x_59); lean_inc(x_63); -lean_inc_ref(x_62); -x_71 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars(x_4, x_70, x_62, x_63, x_60, x_59, x_64, x_61); +lean_inc_ref(x_61); +lean_inc(x_60); +lean_inc_ref(x_64); +x_71 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars(x_4, x_70, x_64, x_60, x_61, x_63, x_59, x_62); if (lean_obj_tag(x_71) == 0) { lean_object* x_72; lean_dec_ref(x_71); -lean_inc(x_61); +lean_inc(x_62); +lean_inc_ref(x_59); +lean_inc(x_63); +lean_inc_ref(x_61); lean_inc_ref(x_64); -lean_inc(x_59); -lean_inc_ref(x_60); -lean_inc_ref(x_62); lean_inc(x_8); -x_72 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__3(x_8, x_62, x_63, x_60, x_59, x_64, x_61); +x_72 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__3(x_8, x_64, x_60, x_61, x_63, x_59, x_62); if (lean_obj_tag(x_72) == 0) { uint8_t x_73; @@ -52546,10 +53188,10 @@ if (x_75 == 0) lean_dec(x_69); lean_dec_ref(x_64); lean_dec(x_63); -lean_dec_ref(x_62); -lean_dec(x_61); -lean_dec_ref(x_60); -lean_dec(x_59); +lean_dec(x_62); +lean_dec_ref(x_61); +lean_dec(x_60); +lean_dec_ref(x_59); lean_dec(x_8); lean_ctor_set(x_72, 0, x_68); return x_72; @@ -52566,12 +53208,12 @@ if (x_79 == 0) { lean_dec(x_69); x_20 = x_59; -x_21 = x_68; -x_22 = x_60; -x_23 = x_61; -x_24 = x_62; +x_21 = x_60; +x_22 = x_61; +x_23 = x_62; +x_24 = x_63; x_25 = x_64; -x_26 = x_63; +x_26 = x_68; x_27 = x_78; x_28 = lean_box(0); goto block_45; @@ -52586,12 +53228,12 @@ if (x_79 == 0) { lean_dec(x_69); x_20 = x_59; -x_21 = x_68; -x_22 = x_60; -x_23 = x_61; -x_24 = x_62; +x_21 = x_60; +x_22 = x_61; +x_23 = x_62; +x_24 = x_63; x_25 = x_64; -x_26 = x_63; +x_26 = x_68; x_27 = x_78; x_28 = lean_box(0); goto block_45; @@ -52601,20 +53243,20 @@ else size_t x_81; size_t x_82; lean_object* x_83; x_81 = 0; x_82 = lean_usize_of_nat(x_77); -lean_inc(x_61); +lean_inc(x_62); +lean_inc_ref(x_59); +lean_inc(x_63); +lean_inc_ref(x_61); lean_inc_ref(x_64); -lean_inc(x_59); -lean_inc_ref(x_60); -lean_inc_ref(x_62); -x_83 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__9___redArg(x_69, x_81, x_82, x_78, x_62, x_63, x_60, x_59, x_64, x_61); +x_83 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__9___redArg(x_69, x_81, x_82, x_78, x_64, x_60, x_61, x_63, x_59, x_62); lean_dec(x_69); x_46 = x_59; x_47 = x_60; -x_48 = x_68; -x_49 = x_61; -x_50 = x_62; -x_51 = x_63; -x_52 = x_64; +x_48 = x_61; +x_49 = x_62; +x_50 = x_63; +x_51 = x_64; +x_52 = x_68; x_53 = x_83; goto block_58; } @@ -52624,20 +53266,20 @@ else size_t x_84; size_t x_85; lean_object* x_86; x_84 = 0; x_85 = lean_usize_of_nat(x_77); -lean_inc(x_61); +lean_inc(x_62); +lean_inc_ref(x_59); +lean_inc(x_63); +lean_inc_ref(x_61); lean_inc_ref(x_64); -lean_inc(x_59); -lean_inc_ref(x_60); -lean_inc_ref(x_62); -x_86 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__9___redArg(x_69, x_84, x_85, x_78, x_62, x_63, x_60, x_59, x_64, x_61); +x_86 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__9___redArg(x_69, x_84, x_85, x_78, x_64, x_60, x_61, x_63, x_59, x_62); lean_dec(x_69); x_46 = x_59; x_47 = x_60; -x_48 = x_68; -x_49 = x_61; -x_50 = x_62; -x_51 = x_63; -x_52 = x_64; +x_48 = x_61; +x_49 = x_62; +x_50 = x_63; +x_51 = x_64; +x_52 = x_68; x_53 = x_86; goto block_58; } @@ -52658,10 +53300,10 @@ lean_object* x_89; lean_dec(x_69); lean_dec_ref(x_64); lean_dec(x_63); -lean_dec_ref(x_62); -lean_dec(x_61); -lean_dec_ref(x_60); -lean_dec(x_59); +lean_dec(x_62); +lean_dec_ref(x_61); +lean_dec(x_60); +lean_dec_ref(x_59); lean_dec(x_8); x_89 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_89, 0, x_68); @@ -52678,12 +53320,12 @@ if (x_93 == 0) { lean_dec(x_69); x_20 = x_59; -x_21 = x_68; -x_22 = x_60; -x_23 = x_61; -x_24 = x_62; +x_21 = x_60; +x_22 = x_61; +x_23 = x_62; +x_24 = x_63; x_25 = x_64; -x_26 = x_63; +x_26 = x_68; x_27 = x_92; x_28 = lean_box(0); goto block_45; @@ -52698,12 +53340,12 @@ if (x_93 == 0) { lean_dec(x_69); x_20 = x_59; -x_21 = x_68; -x_22 = x_60; -x_23 = x_61; -x_24 = x_62; +x_21 = x_60; +x_22 = x_61; +x_23 = x_62; +x_24 = x_63; x_25 = x_64; -x_26 = x_63; +x_26 = x_68; x_27 = x_92; x_28 = lean_box(0); goto block_45; @@ -52713,20 +53355,20 @@ else size_t x_95; size_t x_96; lean_object* x_97; x_95 = 0; x_96 = lean_usize_of_nat(x_91); -lean_inc(x_61); +lean_inc(x_62); +lean_inc_ref(x_59); +lean_inc(x_63); +lean_inc_ref(x_61); lean_inc_ref(x_64); -lean_inc(x_59); -lean_inc_ref(x_60); -lean_inc_ref(x_62); -x_97 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__9___redArg(x_69, x_95, x_96, x_92, x_62, x_63, x_60, x_59, x_64, x_61); +x_97 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__9___redArg(x_69, x_95, x_96, x_92, x_64, x_60, x_61, x_63, x_59, x_62); lean_dec(x_69); x_46 = x_59; x_47 = x_60; -x_48 = x_68; -x_49 = x_61; -x_50 = x_62; -x_51 = x_63; -x_52 = x_64; +x_48 = x_61; +x_49 = x_62; +x_50 = x_63; +x_51 = x_64; +x_52 = x_68; x_53 = x_97; goto block_58; } @@ -52736,20 +53378,20 @@ else size_t x_98; size_t x_99; lean_object* x_100; x_98 = 0; x_99 = lean_usize_of_nat(x_91); -lean_inc(x_61); +lean_inc(x_62); +lean_inc_ref(x_59); +lean_inc(x_63); +lean_inc_ref(x_61); lean_inc_ref(x_64); -lean_inc(x_59); -lean_inc_ref(x_60); -lean_inc_ref(x_62); -x_100 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__9___redArg(x_69, x_98, x_99, x_92, x_62, x_63, x_60, x_59, x_64, x_61); +x_100 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__9___redArg(x_69, x_98, x_99, x_92, x_64, x_60, x_61, x_63, x_59, x_62); lean_dec(x_69); x_46 = x_59; x_47 = x_60; -x_48 = x_68; -x_49 = x_61; -x_50 = x_62; -x_51 = x_63; -x_52 = x_64; +x_48 = x_61; +x_49 = x_62; +x_50 = x_63; +x_51 = x_64; +x_52 = x_68; x_53 = x_100; goto block_58; } @@ -52764,10 +53406,10 @@ lean_dec(x_69); lean_dec(x_68); lean_dec_ref(x_64); lean_dec(x_63); -lean_dec_ref(x_62); -lean_dec(x_61); -lean_dec_ref(x_60); -lean_dec(x_59); +lean_dec(x_62); +lean_dec_ref(x_61); +lean_dec(x_60); +lean_dec_ref(x_59); lean_dec(x_8); x_101 = !lean_is_exclusive(x_72); if (x_101 == 0) @@ -52793,10 +53435,10 @@ lean_dec(x_69); lean_dec(x_68); lean_dec_ref(x_64); lean_dec(x_63); -lean_dec_ref(x_62); -lean_dec(x_61); -lean_dec_ref(x_60); -lean_dec(x_59); +lean_dec(x_62); +lean_dec_ref(x_61); +lean_dec(x_60); +lean_dec_ref(x_59); lean_dec(x_8); x_104 = !lean_is_exclusive(x_71); if (x_104 == 0) @@ -52820,10 +53462,10 @@ else lean_object* x_107; lean_object* x_108; lean_dec_ref(x_64); lean_dec(x_63); -lean_dec_ref(x_62); -lean_dec(x_61); -lean_dec_ref(x_60); -lean_dec(x_59); +lean_dec(x_62); +lean_dec_ref(x_61); +lean_dec(x_60); +lean_dec_ref(x_59); lean_dec(x_8); lean_dec_ref(x_4); x_107 = lean_ctor_get(x_65, 0); @@ -52846,8 +53488,8 @@ x_59 = x_110; x_60 = x_111; x_61 = x_112; x_62 = x_113; -x_63 = x_115; -x_64 = x_114; +x_63 = x_114; +x_64 = x_115; x_65 = x_117; x_66 = lean_box(0); goto block_109; @@ -52855,12 +53497,12 @@ goto block_109; else { uint8_t x_118; -lean_dec(x_115); -lean_dec_ref(x_114); -lean_dec_ref(x_113); -lean_dec(x_112); -lean_dec_ref(x_111); -lean_dec(x_110); +lean_dec_ref(x_115); +lean_dec(x_114); +lean_dec(x_113); +lean_dec_ref(x_112); +lean_dec(x_111); +lean_dec_ref(x_110); lean_dec(x_8); lean_dec_ref(x_4); x_118 = !lean_is_exclusive(x_116); @@ -52885,13 +53527,13 @@ return x_120; if (x_131 == 0) { lean_object* x_132; -lean_inc(x_127); -lean_inc_ref(x_129); -lean_inc(x_124); -lean_inc_ref(x_126); -lean_inc_ref(x_128); +lean_inc(x_128); +lean_inc_ref(x_124); +lean_inc(x_129); +lean_inc_ref(x_127); +lean_inc_ref(x_130); lean_inc(x_8); -x_132 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__3(x_8, x_128, x_130, x_126, x_124, x_129, x_127); +x_132 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__3(x_8, x_130, x_126, x_127, x_129, x_124, x_128); if (lean_obj_tag(x_132) == 0) { lean_object* x_133; uint8_t x_134; @@ -52903,18 +53545,18 @@ lean_dec(x_133); if (x_134 == 0) { lean_object* x_135; lean_object* x_136; -lean_dec_ref(x_123); +lean_dec_ref(x_125); lean_dec_ref(x_12); x_135 = lean_box(0); -lean_inc(x_127); -lean_inc_ref(x_129); -lean_inc(x_124); -lean_inc_ref(x_126); -lean_inc(x_130); -lean_inc_ref(x_128); +lean_inc(x_128); +lean_inc_ref(x_124); +lean_inc(x_129); +lean_inc_ref(x_127); +lean_inc(x_126); +lean_inc_ref(x_130); lean_inc_ref(x_4); lean_inc(x_8); -x_136 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__9(x_2, x_3, x_5, x_122, x_8, x_4, x_9, x_10, x_131, x_11, x_135, x_128, x_130, x_126, x_124, x_129, x_127); +x_136 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__9(x_2, x_3, x_5, x_122, x_8, x_4, x_9, x_10, x_131, x_11, x_135, x_130, x_126, x_127, x_129, x_124, x_128); x_110 = x_124; x_111 = x_126; x_112 = x_127; @@ -52936,32 +53578,32 @@ x_140 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVC x_141 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_141, 0, x_139); lean_ctor_set(x_141, 1, x_140); -x_142 = l_Lean_Exception_toMessageData(x_123); +x_142 = l_Lean_Exception_toMessageData(x_125); x_143 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_143, 0, x_141); lean_ctor_set(x_143, 1, x_142); -lean_inc(x_127); -lean_inc_ref(x_129); -lean_inc(x_124); -lean_inc_ref(x_126); -lean_inc_ref(x_128); +lean_inc(x_128); +lean_inc_ref(x_124); +lean_inc(x_129); +lean_inc_ref(x_127); +lean_inc_ref(x_130); lean_inc(x_8); -x_144 = l_Lean_addTrace___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__5(x_8, x_143, x_128, x_130, x_126, x_124, x_129, x_127); +x_144 = l_Lean_addTrace___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__5(x_8, x_143, x_130, x_126, x_127, x_129, x_124, x_128); if (lean_obj_tag(x_144) == 0) { lean_object* x_145; lean_object* x_146; x_145 = lean_ctor_get(x_144, 0); lean_inc(x_145); lean_dec_ref(x_144); -lean_inc(x_127); -lean_inc_ref(x_129); -lean_inc(x_124); -lean_inc_ref(x_126); -lean_inc(x_130); -lean_inc_ref(x_128); +lean_inc(x_128); +lean_inc_ref(x_124); +lean_inc(x_129); +lean_inc_ref(x_127); +lean_inc(x_126); +lean_inc_ref(x_130); lean_inc_ref(x_4); lean_inc(x_8); -x_146 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__9(x_2, x_3, x_5, x_122, x_8, x_4, x_9, x_10, x_131, x_11, x_145, x_128, x_130, x_126, x_124, x_129, x_127); +x_146 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__9(x_2, x_3, x_5, x_122, x_8, x_4, x_9, x_10, x_131, x_11, x_145, x_130, x_126, x_127, x_129, x_124, x_128); x_110 = x_124; x_111 = x_126; x_112 = x_127; @@ -52974,12 +53616,12 @@ goto block_121; else { uint8_t x_147; -lean_dec(x_130); -lean_dec_ref(x_129); -lean_dec_ref(x_128); -lean_dec(x_127); -lean_dec_ref(x_126); -lean_dec(x_124); +lean_dec_ref(x_130); +lean_dec(x_129); +lean_dec(x_128); +lean_dec_ref(x_127); +lean_dec(x_126); +lean_dec_ref(x_124); lean_dec_ref(x_10); lean_dec_ref(x_9); lean_dec(x_8); @@ -53008,13 +53650,13 @@ return x_149; else { uint8_t x_150; -lean_dec(x_130); -lean_dec_ref(x_129); -lean_dec_ref(x_128); -lean_dec(x_127); -lean_dec_ref(x_126); -lean_dec(x_124); -lean_dec_ref(x_123); +lean_dec_ref(x_130); +lean_dec(x_129); +lean_dec(x_128); +lean_dec_ref(x_127); +lean_dec(x_126); +lean_dec_ref(x_125); +lean_dec_ref(x_124); lean_dec_ref(x_12); lean_dec_ref(x_10); lean_dec_ref(x_9); @@ -53043,12 +53685,12 @@ return x_152; else { lean_object* x_153; -lean_dec(x_130); -lean_dec_ref(x_129); -lean_dec_ref(x_128); -lean_dec(x_127); -lean_dec_ref(x_126); -lean_dec(x_124); +lean_dec_ref(x_130); +lean_dec(x_129); +lean_dec(x_128); +lean_dec_ref(x_127); +lean_dec(x_126); +lean_dec_ref(x_124); lean_dec_ref(x_12); lean_dec_ref(x_10); lean_dec_ref(x_9); @@ -53058,7 +53700,7 @@ lean_dec_ref(x_4); lean_dec_ref(x_3); lean_dec_ref(x_2); x_153 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_153, 0, x_123); +lean_ctor_set(x_153, 0, x_125); return x_153; } } @@ -53072,28 +53714,28 @@ uint8_t x_165; lean_inc_ref(x_162); x_165 = l_Lean_Exception_isRuntime(x_162); x_122 = x_155; -x_123 = x_162; +x_123 = lean_box(0); x_124 = x_156; -x_125 = lean_box(0); +x_125 = x_162; x_126 = x_157; x_127 = x_158; x_128 = x_159; -x_129 = x_161; -x_130 = x_160; +x_129 = x_160; +x_130 = x_161; x_131 = x_165; goto block_154; } else { x_122 = x_155; -x_123 = x_162; +x_123 = lean_box(0); x_124 = x_156; -x_125 = lean_box(0); +x_125 = x_162; x_126 = x_157; x_127 = x_158; x_128 = x_159; -x_129 = x_161; -x_130 = x_160; +x_129 = x_160; +x_130 = x_161; x_131 = x_164; goto block_154; } @@ -53116,8 +53758,8 @@ x_59 = x_168; x_60 = x_169; x_61 = x_170; x_62 = x_171; -x_63 = x_173; -x_64 = x_172; +x_63 = x_172; +x_64 = x_173; x_65 = x_175; x_66 = lean_box(0); goto block_109; @@ -53133,8 +53775,8 @@ x_156 = x_168; x_157 = x_169; x_158 = x_170; x_159 = x_171; -x_160 = x_173; -x_161 = x_172; +x_160 = x_172; +x_161 = x_173; x_162 = x_176; x_163 = lean_box(0); goto block_166; @@ -53146,28 +53788,28 @@ lean_object* x_189; lean_object* x_190; x_189 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_189, 0, x_180); lean_ctor_set(x_189, 1, x_188); -lean_inc(x_184); -lean_inc_ref(x_186); -lean_inc(x_179); -lean_inc_ref(x_183); -lean_inc_ref(x_185); +lean_inc(x_185); +lean_inc_ref(x_182); +lean_inc(x_186); +lean_inc_ref(x_184); +lean_inc_ref(x_187); lean_inc(x_8); -x_190 = l_Lean_addTrace___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__5(x_8, x_189, x_185, x_187, x_183, x_179, x_186, x_184); +x_190 = l_Lean_addTrace___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__5(x_8, x_189, x_187, x_183, x_184, x_186, x_182, x_185); if (lean_obj_tag(x_190) == 0) { lean_object* x_191; lean_object* x_192; x_191 = lean_ctor_get(x_190, 0); lean_inc(x_191); lean_dec_ref(x_190); -lean_inc(x_184); -lean_inc_ref(x_186); -lean_inc(x_179); -lean_inc_ref(x_183); -lean_inc(x_187); -lean_inc_ref(x_185); -x_192 = lean_apply_8(x_181, x_191, x_185, x_187, x_183, x_179, x_186, x_184, lean_box(0)); +lean_inc(x_185); +lean_inc_ref(x_182); +lean_inc(x_186); +lean_inc_ref(x_184); +lean_inc(x_183); +lean_inc_ref(x_187); +x_192 = lean_apply_8(x_181, x_191, x_187, x_183, x_184, x_186, x_182, x_185, lean_box(0)); x_167 = x_178; -x_168 = x_179; +x_168 = x_182; x_169 = x_183; x_170 = x_184; x_171 = x_185; @@ -53184,12 +53826,12 @@ x_193 = lean_ctor_get(x_190, 0); lean_inc(x_193); lean_dec_ref(x_190); x_155 = x_178; -x_156 = x_179; +x_156 = x_182; x_157 = x_183; x_158 = x_184; x_159 = x_185; -x_160 = x_187; -x_161 = x_186; +x_160 = x_186; +x_161 = x_187; x_162 = x_193; x_163 = lean_box(0); goto block_166; @@ -53283,12 +53925,12 @@ lean_inc_ref(x_195); lean_inc_ref(x_3); x_216 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__14(x_3, x_212, x_215, x_195, x_196, x_197, x_198, x_199, x_200); x_167 = x_202; -x_168 = x_198; -x_169 = x_197; -x_170 = x_200; -x_171 = x_195; -x_172 = x_199; -x_173 = x_196; +x_168 = x_199; +x_169 = x_196; +x_170 = x_197; +x_171 = x_200; +x_172 = x_198; +x_173 = x_195; x_174 = x_216; goto block_177; } @@ -53323,15 +53965,15 @@ x_227 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_227, 0, x_225); lean_ctor_set(x_227, 1, x_226); x_178 = x_202; -x_179 = x_198; +x_179 = lean_box(0); x_180 = x_223; x_181 = x_213; -x_182 = lean_box(0); -x_183 = x_197; -x_184 = x_200; -x_185 = x_195; -x_186 = x_199; -x_187 = x_196; +x_182 = x_199; +x_183 = x_196; +x_184 = x_197; +x_185 = x_200; +x_186 = x_198; +x_187 = x_195; x_188 = x_227; goto block_194; } @@ -53348,15 +53990,15 @@ x_232 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_232, 0, x_229); lean_ctor_set(x_232, 1, x_231); x_178 = x_202; -x_179 = x_198; +x_179 = lean_box(0); x_180 = x_223; x_181 = x_213; -x_182 = lean_box(0); -x_183 = x_197; -x_184 = x_200; -x_185 = x_195; -x_186 = x_199; -x_187 = x_196; +x_182 = x_199; +x_183 = x_196; +x_184 = x_197; +x_185 = x_200; +x_186 = x_198; +x_187 = x_195; x_188 = x_232; goto block_194; } @@ -53382,15 +54024,15 @@ x_241 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_241, 0, x_239); lean_ctor_set(x_241, 1, x_240); x_178 = x_202; -x_179 = x_198; +x_179 = lean_box(0); x_180 = x_223; x_181 = x_213; -x_182 = lean_box(0); -x_183 = x_197; -x_184 = x_200; -x_185 = x_195; -x_186 = x_199; -x_187 = x_196; +x_182 = x_199; +x_183 = x_196; +x_184 = x_197; +x_185 = x_200; +x_186 = x_198; +x_187 = x_195; x_188 = x_241; goto block_194; } @@ -53406,12 +54048,12 @@ x_242 = lean_ctor_get(x_209, 0); lean_inc(x_242); lean_dec_ref(x_209); x_155 = x_202; -x_156 = x_198; -x_157 = x_197; -x_158 = x_200; -x_159 = x_195; -x_160 = x_196; -x_161 = x_199; +x_156 = x_199; +x_157 = x_196; +x_158 = x_197; +x_159 = x_200; +x_160 = x_198; +x_161 = x_195; x_162 = x_242; x_163 = lean_box(0); goto block_166; @@ -53425,12 +54067,12 @@ x_243 = lean_ctor_get(x_207, 0); lean_inc(x_243); lean_dec_ref(x_207); x_155 = x_202; -x_156 = x_198; -x_157 = x_197; -x_158 = x_200; -x_159 = x_195; -x_160 = x_196; -x_161 = x_199; +x_156 = x_199; +x_157 = x_196; +x_158 = x_197; +x_159 = x_200; +x_160 = x_198; +x_161 = x_195; x_162 = x_243; x_163 = lean_box(0); goto block_166; @@ -55607,7 +56249,7 @@ x_14 = l_Lean_MVarId_withContext___at___00__private_Lean_Elab_Tactic_Do_VCGen_0_ return x_14; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { if (lean_obj_tag(x_3) == 0) @@ -55726,12 +56368,12 @@ lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean x_38 = lean_ctor_get(x_37, 0); lean_inc(x_38); lean_dec_ref(x_37); -x_39 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__1; +x_39 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__1; x_40 = l_Lean_MessageData_ofName(x_38); lean_ctor_set_tag(x_3, 7); lean_ctor_set(x_3, 1, x_40); lean_ctor_set(x_3, 0, x_39); -x_41 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__3; +x_41 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__3; x_42 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_42, 0, x_3); lean_ctor_set(x_42, 1, x_41); @@ -55740,7 +56382,7 @@ lean_ctor_set(x_20, 0, x_14); x_43 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_43, 0, x_42); lean_ctor_set(x_43, 1, x_20); -x_44 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__5; +x_44 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__5; x_45 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_45, 0, x_43); lean_ctor_set(x_45, 1, x_44); @@ -55749,7 +56391,7 @@ x_47 = l_Lean_MessageData_ofExpr(x_46); x_48 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_48, 0, x_45); lean_ctor_set(x_48, 1, x_47); -x_49 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__7; +x_49 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__7; x_50 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_50, 0, x_48); lean_ctor_set(x_50, 1, x_49); @@ -55911,12 +56553,12 @@ lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean x_76 = lean_ctor_get(x_75, 0); lean_inc(x_76); lean_dec_ref(x_75); -x_77 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__1; +x_77 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__1; x_78 = l_Lean_MessageData_ofName(x_76); lean_ctor_set_tag(x_3, 7); lean_ctor_set(x_3, 1, x_78); lean_ctor_set(x_3, 0, x_77); -x_79 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__3; +x_79 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__3; x_80 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_80, 0, x_3); lean_ctor_set(x_80, 1, x_79); @@ -55926,7 +56568,7 @@ lean_ctor_set(x_81, 0, x_14); x_82 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_82, 0, x_80); lean_ctor_set(x_82, 1, x_81); -x_83 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__5; +x_83 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__5; x_84 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_84, 0, x_82); lean_ctor_set(x_84, 1, x_83); @@ -55935,7 +56577,7 @@ x_86 = l_Lean_MessageData_ofExpr(x_85); x_87 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_87, 0, x_84); lean_ctor_set(x_87, 1, x_86); -x_88 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__7; +x_88 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__7; x_89 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_89, 0, x_87); lean_ctor_set(x_89, 1, x_88); @@ -56251,12 +56893,12 @@ lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; x_132 = lean_ctor_get(x_131, 0); lean_inc(x_132); lean_dec_ref(x_131); -x_133 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__1; +x_133 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__1; x_134 = l_Lean_MessageData_ofName(x_132); x_135 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_135, 0, x_133); lean_ctor_set(x_135, 1, x_134); -x_136 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__3; +x_136 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__3; x_137 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_137, 0, x_135); lean_ctor_set(x_137, 1, x_136); @@ -56270,7 +56912,7 @@ lean_ctor_set(x_138, 0, x_108); x_139 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_139, 0, x_137); lean_ctor_set(x_139, 1, x_138); -x_140 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__5; +x_140 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__5; x_141 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_141, 0, x_139); lean_ctor_set(x_141, 1, x_140); @@ -56279,7 +56921,7 @@ x_143 = l_Lean_MessageData_ofExpr(x_142); x_144 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_144, 0, x_141); lean_ctor_set(x_144, 1, x_143); -x_145 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__7; +x_145 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__7; x_146 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_146, 0, x_144); lean_ctor_set(x_146, 1, x_145); @@ -56491,11 +57133,11 @@ return x_164; } } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { lean_object* x_14; -x_14 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg(x_1, x_2, x_4, x_5, x_7, x_8, x_9, x_10, x_11, x_12); +x_14 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg(x_1, x_2, x_4, x_5, x_7, x_8, x_9, x_10, x_11, x_12); return x_14; } } @@ -56515,27 +57157,27 @@ x_10 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00__private_Lea return x_10; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51(lean_object* x_1, uint8_t x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52(lean_object* x_1, uint8_t x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { lean_object* x_18; -x_18 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg(x_1, x_2, x_3, x_4, x_5, x_8, x_9, x_11, x_12, x_13, x_14, x_15, x_16); +x_18 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg(x_1, x_2, x_3, x_4, x_5, x_8, x_9, x_11, x_12, x_13, x_14, x_15, x_16); return x_18; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19) { _start: { lean_object* x_21; -x_21 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_11, x_12, x_14, x_15, x_16, x_17, x_18, x_19); +x_21 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_11, x_12, x_14, x_15, x_16, x_17, x_18, x_19); return x_21; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56_spec__64(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57_spec__65(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { _start: { lean_object* x_20; -x_20 = l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56_spec__64___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_10, x_11, x_13, x_14, x_15, x_16, x_17, x_18); +x_20 = l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57_spec__65___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_10, x_11, x_13, x_14, x_15, x_16, x_17, x_18); return x_20; } } @@ -56563,11 +57205,11 @@ x_9 = l___private_Init_Data_Array_QSort_Basic_0__Array_qsort_sort___at___00Lean_ return x_9; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__55(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__56(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_12; -x_12 = l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__55___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_12 = l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__56___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_12; } } @@ -56638,11 +57280,11 @@ x_11 = l_Lean_MVarId_withContext___at___00__private_Lean_Elab_Tactic_Do_VCGen_0_ return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__50___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__51___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; -x_12 = l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__50(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_12 = l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__51(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec_ref(x_3); lean_dec_ref(x_2); return x_12; @@ -56666,12 +57308,12 @@ x_14 = l_Lean_Meta_lambdaBoundedTelescope___at___00__private_Lean_Elab_Tactic_Do return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__48___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__49___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { uint8_t x_12; lean_object* x_13; x_12 = lean_unbox(x_4); -x_13 = l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__48(x_1, x_2, x_3, x_12, x_5, x_6, x_7, x_8, x_9, x_10); +x_13 = l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__49(x_1, x_2, x_3, x_12, x_5, x_6, x_7, x_8, x_9, x_10); return x_13; } } @@ -56770,11 +57412,11 @@ x_15 = l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean return x_15; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { lean_object* x_14; -x_14 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_14 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_3); return x_14; } @@ -56823,7 +57465,7 @@ lean_dec_ref(x_3); return x_10; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___boxed(lean_object** _args) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -56846,12 +57488,12 @@ lean_object* x_17 = _args[16]; uint8_t x_18; uint8_t x_19; lean_object* x_20; x_18 = lean_unbox(x_2); x_19 = lean_unbox(x_4); -x_20 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51(x_1, x_18, x_3, x_19, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +x_20 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52(x_1, x_18, x_3, x_19, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); lean_dec(x_1); return x_20; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___boxed(lean_object** _args) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -56877,14 +57519,14 @@ lean_object* x_20 = _args[19]; uint8_t x_21; uint8_t x_22; lean_object* x_23; x_21 = lean_unbox(x_5); x_22 = lean_unbox(x_6); -x_23 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56(x_1, x_2, x_3, x_4, x_21, x_22, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); +x_23 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57(x_1, x_2, x_3, x_4, x_21, x_22, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); lean_dec_ref(x_3); lean_dec(x_2); lean_dec(x_1); return x_23; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56_spec__64___boxed(lean_object** _args) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57_spec__65___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -56909,7 +57551,7 @@ lean_object* x_19 = _args[18]; uint8_t x_20; uint8_t x_21; lean_object* x_22; x_20 = lean_unbox(x_4); x_21 = lean_unbox(x_5); -x_22 = l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56_spec__64(x_1, x_2, x_3, x_20, x_21, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +x_22 = l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57_spec__65(x_1, x_2, x_3, x_20, x_21, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); lean_dec(x_2); lean_dec(x_1); return x_22; @@ -56978,11 +57620,11 @@ lean_dec(x_2); return x_9; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__55___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__56___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; -x_12 = l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__55(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_12 = l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__56(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_12; } } @@ -57004,11 +57646,11 @@ x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52___boxed(lean_object* x_1, lean_object* x_2, lean_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_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_9 = l_Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_3); return x_9; } @@ -57046,11 +57688,11 @@ x_10 = l_Lean_MVarId_withContext___at___00__private_Lean_Elab_Tactic_Do_VCGen_0_ return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__50___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__51___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__50___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__51___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec_ref(x_2); lean_dec_ref(x_1); return x_11; @@ -57066,56 +57708,56 @@ x_16 = l_Lean_Meta_withLetDecl___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__L return x_16; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { -uint8_t x_15; uint8_t x_16; uint8_t x_17; lean_object* x_18; -x_15 = lean_unbox(x_3); -x_16 = lean_unbox(x_5); -x_17 = lean_unbox(x_7); -x_18 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59(x_1, x_2, x_15, x_4, x_16, x_6, x_17, x_8, x_9, x_10, x_11, x_12, x_13); -return x_18; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -uint8_t x_15; uint8_t x_16; uint8_t x_17; lean_object* x_18; -x_15 = lean_unbox(x_3); -x_16 = lean_unbox(x_5); -x_17 = lean_unbox(x_7); -x_18 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60(x_1, x_2, x_15, x_4, x_16, x_6, x_17, x_8, x_9, x_10, x_11, x_12, x_13); -return x_18; +uint8_t x_16; uint8_t x_17; uint8_t x_18; lean_object* x_19; +x_16 = lean_unbox(x_4); +x_17 = lean_unbox(x_6); +x_18 = lean_unbox(x_8); +x_19 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60(x_1, x_2, x_3, x_16, x_5, x_17, x_7, x_18, x_9, x_10, x_11, x_12, x_13, x_14); +return x_19; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__61___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__61___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { -uint8_t x_13; uint8_t x_14; lean_object* x_15; -x_13 = lean_unbox(x_3); -x_14 = lean_unbox(x_5); -x_15 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__61(x_1, x_2, x_13, x_4, x_14, x_6, x_7, x_8, x_9, x_10, x_11); -return x_15; +uint8_t x_16; uint8_t x_17; uint8_t x_18; lean_object* x_19; +x_16 = lean_unbox(x_4); +x_17 = lean_unbox(x_6); +x_18 = lean_unbox(x_8); +x_19 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__61(x_1, x_2, x_3, x_16, x_5, x_17, x_7, x_18, x_9, x_10, x_11, x_12, x_13, x_14); +return x_19; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { uint8_t x_14; uint8_t x_15; lean_object* x_16; x_14 = lean_unbox(x_4); x_15 = lean_unbox(x_6); -x_16 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62(x_1, x_2, x_3, x_14, x_5, x_15, x_7, x_8, x_9, x_10, x_11, x_12); +x_16 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62(x_1, x_2, x_3, x_14, x_5, x_15, x_7, x_8, x_9, x_10, x_11, x_12); return x_16; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__42___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__63___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +uint8_t x_15; uint8_t x_16; lean_object* x_17; +x_15 = lean_unbox(x_5); +x_16 = lean_unbox(x_7); +x_17 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__63(x_1, x_2, x_3, x_4, x_15, x_6, x_16, x_8, x_9, x_10, x_11, x_12, x_13); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__43___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { uint8_t x_14; uint8_t x_15; uint8_t x_16; lean_object* x_17; x_14 = lean_unbox(x_2); x_15 = lean_unbox(x_4); x_16 = lean_unbox(x_6); -x_17 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__42(x_1, x_14, x_3, x_15, x_5, x_16, x_7, x_8, x_9, x_10, x_11, x_12); +x_17 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit_spec__43(x_1, x_14, x_3, x_15, x_5, x_16, x_7, x_8, x_9, x_10, x_11, x_12); return x_17; } } @@ -57139,15 +57781,15 @@ x_10 = l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs return x_10; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__55___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__56___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; -x_11 = l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__55___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__56___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___boxed(lean_object** _args) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -57185,7 +57827,7 @@ x_31 = lean_unbox(x_17); x_32 = lean_unbox(x_18); x_33 = lean_unbox(x_21); x_34 = lean_unbox(x_23); -x_35 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_31, x_32, x_19, x_20, x_33, x_22, x_34, x_24, x_25, x_26, x_27, x_28, x_29); +x_35 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_31, x_32, x_19, x_20, x_33, x_22, x_34, x_24, x_25, x_26, x_27, x_28, x_29); return x_35; } } @@ -57199,13 +57841,13 @@ x_13 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean return x_13; } } -LEAN_EXPORT lean_object* l_Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { uint8_t x_11; uint8_t x_12; lean_object* x_13; x_11 = lean_unbox(x_2); x_12 = lean_unbox(x_3); -x_13 = l_Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59(x_1, x_11, x_12, x_4, x_5, x_6, x_7, x_8, x_9); +x_13 = l_Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60(x_1, x_11, x_12, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_5); return x_13; } @@ -57228,11 +57870,11 @@ x_12 = l_Lean_Meta_forallTelescope___at___00__private_Lean_Elab_Tactic_Do_VCGen_ return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__57___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__58___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; -x_9 = l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__57(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_9 = l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__58(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_3); return x_9; } @@ -57256,12 +57898,12 @@ x_13 = l_Lean_Meta_lambdaBoundedTelescope___at___00__private_Lean_Elab_Tactic_Do return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__48___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__49___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { uint8_t x_11; lean_object* x_12; x_11 = lean_unbox(x_3); -x_12 = l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__48___redArg(x_1, x_2, x_11, x_4, x_5, x_6, x_7, x_8, x_9); +x_12 = l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__49___redArg(x_1, x_2, x_11, x_4, x_5, x_6, x_7, x_8, x_9); return x_12; } } @@ -57296,6 +57938,18 @@ lean_dec(x_1); return x_9; } } +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___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___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37(x_4, x_5, x_3); +return x_6; +} +} LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at___00Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__6_spec__8_spec__12_spec__34___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { @@ -57358,7 +58012,7 @@ lean_dec(x_4); return x_5; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__46___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__47___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { size_t x_12; size_t x_13; lean_object* x_14; @@ -57366,7 +58020,7 @@ x_12 = lean_unbox_usize(x_2); lean_dec(x_2); x_13 = lean_unbox_usize(x_3); lean_dec(x_3); -x_14 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__46(x_1, x_12, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_14 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__47(x_1, x_12, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_14; } } @@ -57384,7 +58038,7 @@ lean_dec_ref(x_3); return x_17; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__45___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__46___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { size_t x_11; size_t x_12; lean_object* x_13; @@ -57392,7 +58046,7 @@ x_11 = lean_unbox_usize(x_1); lean_dec(x_1); x_12 = lean_unbox_usize(x_2); lean_dec(x_2); -x_13 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__45(x_11, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_13 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__46(x_11, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_5); return x_13; } @@ -57518,15 +58172,15 @@ x_9 = l_Lean_Elab_Tactic_collectFreshMVars___at___00__private_Lean_Elab_Tactic_D return x_9; } } -LEAN_EXPORT lean_object* l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__54___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__55___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__54(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_9 = l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__55(x_1, x_2, x_3, x_4, x_5, x_6, x_7); return x_9; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__49___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__50___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { size_t x_12; size_t x_13; lean_object* x_14; @@ -57534,7 +58188,7 @@ x_12 = lean_unbox_usize(x_2); lean_dec(x_2); x_13 = lean_unbox_usize(x_3); lean_dec(x_3); -x_14 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__49(x_1, x_12, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_14 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__50(x_1, x_12, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_6); lean_dec_ref(x_1); return x_14; @@ -57578,11 +58232,11 @@ lean_dec_ref(x_6); return x_20; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; -x_12 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_12 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_12; } } @@ -57595,24 +58249,24 @@ lean_dec(x_3); return x_13; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { uint8_t x_15; uint8_t x_16; lean_object* x_17; x_15 = lean_unbox(x_2); x_16 = lean_unbox(x_4); -x_17 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__51___redArg(x_1, x_15, x_3, x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_17 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__52___redArg(x_1, x_15, x_3, x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_1); return x_17; } } -LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { uint8_t x_12; uint8_t x_13; lean_object* x_14; x_12 = lean_unbox(x_3); x_13 = lean_unbox(x_4); -x_14 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70(x_1, x_2, x_12, x_13, x_5, x_6, x_7, x_8, x_9, x_10); +x_14 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71(x_1, x_2, x_12, x_13, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_6); lean_dec(x_1); return x_14; @@ -57663,7 +58317,7 @@ x_11 = l_Lean_Elab_Tactic_Do_ProofMode_mIntro___at___00__private_Lean_Elab_Tacti return x_11; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { uint8_t x_13; size_t x_14; size_t x_15; lean_object* x_16; @@ -57672,7 +58326,7 @@ x_14 = lean_unbox_usize(x_3); lean_dec(x_3); x_15 = lean_unbox_usize(x_4); lean_dec(x_4); -x_16 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47(x_13, x_2, x_14, x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_16 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48(x_13, x_2, x_14, x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_7); lean_dec_ref(x_2); return x_16; @@ -57686,39 +58340,39 @@ x_13 = l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePost return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -uint8_t x_12; lean_object* x_13; -x_12 = lean_unbox(x_4); -x_13 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18(x_1, x_2, x_3, x_12, x_5, x_6, x_7, x_8, x_9, x_10); -return x_13; -} -} -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { uint8_t x_17; uint8_t x_18; lean_object* x_19; x_17 = lean_unbox(x_4); x_18 = lean_unbox(x_5); -x_19 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg(x_1, x_2, x_3, x_17, x_18, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_19 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg(x_1, x_2, x_3, x_17, x_18, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); lean_dec(x_2); lean_dec(x_1); return x_19; } } -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56_spec__64___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57_spec__65___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { uint8_t x_17; uint8_t x_18; lean_object* x_19; x_17 = lean_unbox(x_4); x_18 = lean_unbox(x_5); -x_19 = l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56_spec__64___redArg(x_1, x_2, x_3, x_17, x_18, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_19 = l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57_spec__65___redArg(x_1, x_2, x_3, x_17, x_18, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); lean_dec(x_2); lean_dec(x_1); return x_19; } } +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_4); +x_13 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18(x_1, x_2, x_3, x_12, x_5, x_6, x_7, x_8, x_9, x_10); +return x_13; +} +} LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { @@ -57752,13 +58406,13 @@ x_14 = l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { uint8_t x_15; uint8_t x_16; lean_object* x_17; x_15 = lean_unbox(x_2); x_16 = lean_unbox(x_3); -x_17 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37(x_1, x_15, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_17 = l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38(x_1, x_15, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_17; } } @@ -59040,7 +59694,7 @@ return x_11; LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_logErrorAt___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__1_spec__1___redArg(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_10; uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_50; uint8_t x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; lean_object* x_55; uint8_t x_56; lean_object* x_57; lean_object* x_77; uint8_t x_78; lean_object* x_79; lean_object* x_80; uint8_t x_81; lean_object* x_82; uint8_t x_83; lean_object* x_84; lean_object* x_88; uint8_t x_89; lean_object* x_90; uint8_t x_91; lean_object* x_92; lean_object* x_93; uint8_t x_94; uint8_t x_100; uint8_t x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; uint8_t x_106; uint8_t x_107; uint8_t x_109; uint8_t x_125; +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_50; lean_object* x_51; uint8_t x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; uint8_t x_56; lean_object* x_57; lean_object* x_77; uint8_t x_78; lean_object* x_79; lean_object* x_80; uint8_t x_81; lean_object* x_82; uint8_t x_83; lean_object* x_84; lean_object* x_88; lean_object* x_89; lean_object* x_90; uint8_t x_91; lean_object* x_92; uint8_t x_93; uint8_t x_94; uint8_t x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; uint8_t x_105; uint8_t x_106; uint8_t x_107; uint8_t x_109; uint8_t x_125; x_100 = 2; x_125 = l_Lean_instBEqMessageSeverity_beq(x_3, x_100); if (x_125 == 0) @@ -59075,15 +59729,15 @@ lean_ctor_set(x_25, 0, x_21); lean_ctor_set(x_25, 1, x_22); x_26 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_13); +lean_ctor_set(x_26, 1, x_14); x_27 = lean_alloc_ctor(0, 5, 3); -lean_ctor_set(x_27, 0, x_12); -lean_ctor_set(x_27, 1, x_16); -lean_ctor_set(x_27, 2, x_10); -lean_ctor_set(x_27, 3, x_14); +lean_ctor_set(x_27, 0, x_16); +lean_ctor_set(x_27, 1, x_10); +lean_ctor_set(x_27, 2, x_11); +lean_ctor_set(x_27, 3, x_12); lean_ctor_set(x_27, 4, x_26); -lean_ctor_set_uint8(x_27, sizeof(void*)*5, x_11); -lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 1, x_15); +lean_ctor_set_uint8(x_27, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 1, x_13); lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 2, x_4); x_28 = l_Lean_MessageLog_add(x_27, x_24); lean_ctor_set(x_20, 6, x_28); @@ -59120,15 +59774,15 @@ lean_ctor_set(x_41, 0, x_21); lean_ctor_set(x_41, 1, x_22); x_42 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_13); +lean_ctor_set(x_42, 1, x_14); x_43 = lean_alloc_ctor(0, 5, 3); -lean_ctor_set(x_43, 0, x_12); -lean_ctor_set(x_43, 1, x_16); -lean_ctor_set(x_43, 2, x_10); -lean_ctor_set(x_43, 3, x_14); +lean_ctor_set(x_43, 0, x_16); +lean_ctor_set(x_43, 1, x_10); +lean_ctor_set(x_43, 2, x_11); +lean_ctor_set(x_43, 3, x_12); lean_ctor_set(x_43, 4, x_42); -lean_ctor_set_uint8(x_43, sizeof(void*)*5, x_11); -lean_ctor_set_uint8(x_43, sizeof(void*)*5 + 1, x_15); +lean_ctor_set_uint8(x_43, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_43, sizeof(void*)*5 + 1, x_13); lean_ctor_set_uint8(x_43, sizeof(void*)*5 + 2, x_4); x_44 = l_Lean_MessageLog_add(x_43, x_38); x_45 = lean_alloc_ctor(0, 9, 0); @@ -59159,24 +59813,24 @@ if (x_60 == 0) lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; x_61 = lean_ctor_get(x_59, 0); lean_inc_ref(x_53); -x_62 = l_Lean_FileMap_toPosition(x_53, x_52); -lean_dec(x_52); +x_62 = l_Lean_FileMap_toPosition(x_53, x_51); +lean_dec(x_51); x_63 = l_Lean_FileMap_toPosition(x_53, x_57); lean_dec(x_57); x_64 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_64, 0, x_63); x_65 = l_Lean_addTrace___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__5___lam__1___closed__1; -if (x_51 == 0) +if (x_56 == 0) { lean_free_object(x_59); lean_dec_ref(x_50); -x_10 = x_64; -x_11 = x_54; -x_12 = x_55; -x_13 = x_61; -x_14 = x_65; -x_15 = x_56; -x_16 = x_62; +x_10 = x_62; +x_11 = x_64; +x_12 = x_65; +x_13 = x_52; +x_14 = x_61; +x_15 = x_55; +x_16 = x_54; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -59193,7 +59847,7 @@ lean_object* x_67; lean_dec_ref(x_64); lean_dec_ref(x_62); lean_dec(x_61); -lean_dec_ref(x_55); +lean_dec_ref(x_54); lean_dec_ref(x_7); x_67 = lean_box(0); lean_ctor_set(x_59, 0, x_67); @@ -59202,13 +59856,13 @@ return x_59; else { lean_free_object(x_59); -x_10 = x_64; -x_11 = x_54; -x_12 = x_55; -x_13 = x_61; -x_14 = x_65; -x_15 = x_56; -x_16 = x_62; +x_10 = x_62; +x_11 = x_64; +x_12 = x_65; +x_13 = x_52; +x_14 = x_61; +x_15 = x_55; +x_16 = x_54; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -59223,23 +59877,23 @@ x_68 = lean_ctor_get(x_59, 0); lean_inc(x_68); lean_dec(x_59); lean_inc_ref(x_53); -x_69 = l_Lean_FileMap_toPosition(x_53, x_52); -lean_dec(x_52); +x_69 = l_Lean_FileMap_toPosition(x_53, x_51); +lean_dec(x_51); x_70 = l_Lean_FileMap_toPosition(x_53, x_57); lean_dec(x_57); x_71 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_71, 0, x_70); x_72 = l_Lean_addTrace___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__5___lam__1___closed__1; -if (x_51 == 0) +if (x_56 == 0) { lean_dec_ref(x_50); -x_10 = x_71; -x_11 = x_54; -x_12 = x_55; -x_13 = x_68; -x_14 = x_72; -x_15 = x_56; -x_16 = x_69; +x_10 = x_69; +x_11 = x_71; +x_12 = x_72; +x_13 = x_52; +x_14 = x_68; +x_15 = x_55; +x_16 = x_54; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -59256,7 +59910,7 @@ lean_object* x_74; lean_object* x_75; lean_dec_ref(x_71); lean_dec_ref(x_69); lean_dec(x_68); -lean_dec_ref(x_55); +lean_dec_ref(x_54); lean_dec_ref(x_7); x_74 = lean_box(0); x_75 = lean_alloc_ctor(0, 1, 0); @@ -59265,13 +59919,13 @@ return x_75; } else { -x_10 = x_71; -x_11 = x_54; -x_12 = x_55; -x_13 = x_68; -x_14 = x_72; -x_15 = x_56; -x_16 = x_69; +x_10 = x_69; +x_11 = x_71; +x_12 = x_72; +x_13 = x_52; +x_14 = x_68; +x_15 = x_55; +x_16 = x_54; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -59283,17 +59937,17 @@ goto block_49; block_87: { lean_object* x_85; -x_85 = l_Lean_Syntax_getTailPos_x3f(x_80, x_81); -lean_dec(x_80); +x_85 = l_Lean_Syntax_getTailPos_x3f(x_79, x_81); +lean_dec(x_79); if (lean_obj_tag(x_85) == 0) { lean_inc(x_84); x_50 = x_77; -x_51 = x_78; -x_52 = x_84; -x_53 = x_79; -x_54 = x_81; -x_55 = x_82; +x_51 = x_84; +x_52 = x_78; +x_53 = x_80; +x_54 = x_82; +x_55 = x_81; x_56 = x_83; x_57 = x_84; goto block_76; @@ -59305,11 +59959,11 @@ x_86 = lean_ctor_get(x_85, 0); lean_inc(x_86); lean_dec_ref(x_85); x_50 = x_77; -x_51 = x_78; -x_52 = x_84; -x_53 = x_79; -x_54 = x_81; -x_55 = x_82; +x_51 = x_84; +x_52 = x_78; +x_53 = x_80; +x_54 = x_82; +x_55 = x_81; x_56 = x_83; x_57 = x_86; goto block_76; @@ -59318,20 +59972,20 @@ goto block_76; block_99: { lean_object* x_95; lean_object* x_96; -x_95 = l_Lean_replaceRef(x_1, x_93); -lean_dec(x_93); +x_95 = l_Lean_replaceRef(x_1, x_92); +lean_dec(x_92); x_96 = l_Lean_Syntax_getPos_x3f(x_95, x_91); if (lean_obj_tag(x_96) == 0) { lean_object* x_97; x_97 = lean_unsigned_to_nat(0u); x_77 = x_88; -x_78 = x_89; -x_79 = x_90; -x_80 = x_95; +x_78 = x_94; +x_79 = x_95; +x_80 = x_89; x_81 = x_91; -x_82 = x_92; -x_83 = x_94; +x_82 = x_90; +x_83 = x_93; x_84 = x_97; goto block_87; } @@ -59342,12 +59996,12 @@ x_98 = lean_ctor_get(x_96, 0); lean_inc(x_98); lean_dec_ref(x_96); x_77 = x_88; -x_78 = x_89; -x_79 = x_90; -x_80 = x_95; +x_78 = x_94; +x_79 = x_95; +x_80 = x_89; x_81 = x_91; -x_82 = x_92; -x_83 = x_94; +x_82 = x_90; +x_83 = x_93; x_84 = x_98; goto block_87; } @@ -59356,9 +60010,9 @@ goto block_87; { if (x_107 == 0) { -x_88 = x_104; -x_89 = x_101; -x_90 = x_102; +x_88 = x_101; +x_89 = x_102; +x_90 = x_104; x_91 = x_106; x_92 = x_103; x_93 = x_105; @@ -59367,9 +60021,9 @@ goto block_99; } else { -x_88 = x_104; -x_89 = x_101; -x_90 = x_102; +x_88 = x_101; +x_89 = x_102; +x_90 = x_104; x_91 = x_106; x_92 = x_103; x_93 = x_105; @@ -59389,21 +60043,21 @@ x_113 = lean_ctor_get(x_7, 5); x_114 = lean_ctor_get_uint8(x_7, sizeof(void*)*14 + 1); x_115 = lean_box(x_109); x_116 = lean_box(x_114); -x_117 = lean_alloc_closure((void*)(l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___boxed), 3, 2); +x_117 = lean_alloc_closure((void*)(l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___boxed), 3, 2); lean_closure_set(x_117, 0, x_115); lean_closure_set(x_117, 1, x_116); x_118 = 1; x_119 = l_Lean_instBEqMessageSeverity_beq(x_3, x_118); if (x_119 == 0) { -lean_inc(x_113); lean_inc_ref(x_110); +lean_inc(x_113); lean_inc_ref(x_111); -x_101 = x_114; +x_101 = x_117; x_102 = x_111; -x_103 = x_110; -x_104 = x_117; -x_105 = x_113; +x_103 = x_113; +x_104 = x_110; +x_105 = x_114; x_106 = x_109; x_107 = x_119; goto block_108; @@ -59411,16 +60065,16 @@ goto block_108; else { lean_object* x_120; uint8_t x_121; -x_120 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___closed__1; +x_120 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___closed__1; x_121 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13_spec__27(x_112, x_120); -lean_inc(x_113); lean_inc_ref(x_110); +lean_inc(x_113); lean_inc_ref(x_111); -x_101 = x_114; +x_101 = x_117; x_102 = x_111; -x_103 = x_110; -x_104 = x_117; -x_105 = x_113; +x_103 = x_113; +x_104 = x_110; +x_105 = x_114; x_106 = x_109; x_107 = x_121; goto block_108; @@ -67200,7 +67854,7 @@ return x_2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_elabMVCGen___lam__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; uint8_t x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; uint8_t x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; uint8_t x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; uint8_t x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; uint8_t x_353; lean_object* x_354; lean_object* x_355; uint8_t x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_409; lean_object* x_410; uint8_t x_411; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_203; lean_object* x_204; uint8_t x_205; lean_object* x_206; uint8_t x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; uint8_t x_287; lean_object* x_288; uint8_t x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; uint8_t x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; uint8_t x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_409; lean_object* x_410; uint8_t x_411; x_409 = lean_ctor_get(x_9, 2); x_410 = l_Lean_Elab_Tactic_Do_elabMVCGen___lam__5___closed__12; x_411 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13_spec__27(x_409, x_410); @@ -67414,7 +68068,7 @@ lean_inc(x_74); lean_inc_ref(x_73); lean_inc(x_72); lean_inc_ref(x_71); -x_76 = l_Lean_Elab_Term_TermElabM_run___redArg(x_62, x_63, x_65, x_71, x_72, x_73, x_74); +x_76 = l_Lean_Elab_Term_TermElabM_run___redArg(x_62, x_63, x_66, x_71, x_72, x_73, x_74); if (lean_obj_tag(x_76) == 0) { lean_object* x_77; lean_object* x_78; lean_object* x_79; uint8_t x_80; @@ -67436,7 +68090,7 @@ lean_inc(x_81); lean_dec(x_77); x_24 = x_81; x_25 = x_64; -x_26 = x_66; +x_26 = x_65; x_27 = x_67; x_28 = x_68; x_29 = x_69; @@ -67458,8 +68112,8 @@ lean_object* x_83; lean_object* x_84; lean_object* x_85; size_t x_86; size_t x_8 x_83 = lean_ctor_get(x_77, 0); x_84 = lean_ctor_get(x_77, 1); lean_dec(x_84); -lean_inc_ref(x_66); -x_85 = l_Array_append___redArg(x_66, x_83); +lean_inc_ref(x_65); +x_85 = l_Array_append___redArg(x_65, x_83); x_86 = lean_array_size(x_85); x_87 = 0; x_88 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Tactic_Do_elabMVCGen_spec__2___redArg(x_86, x_87, x_85, x_71, x_72, x_73, x_74); @@ -67484,7 +68138,7 @@ if (lean_obj_tag(x_95) == 0) lean_dec_ref(x_95); x_24 = x_83; x_25 = x_64; -x_26 = x_66; +x_26 = x_65; x_27 = x_67; x_28 = x_68; x_29 = x_69; @@ -67507,7 +68161,7 @@ lean_dec(x_70); lean_dec_ref(x_69); lean_dec(x_68); lean_dec_ref(x_67); -lean_dec_ref(x_66); +lean_dec_ref(x_65); lean_dec(x_64); return x_95; } @@ -67525,7 +68179,7 @@ lean_dec(x_70); lean_dec_ref(x_69); lean_dec(x_68); lean_dec_ref(x_67); -lean_dec_ref(x_66); +lean_dec_ref(x_65); lean_dec(x_64); x_96 = !lean_is_exclusive(x_88); if (x_96 == 0) @@ -67550,8 +68204,8 @@ lean_object* x_99; lean_object* x_100; size_t x_101; size_t x_102; lean_object* x_99 = lean_ctor_get(x_77, 0); lean_inc(x_99); lean_dec(x_77); -lean_inc_ref(x_66); -x_100 = l_Array_append___redArg(x_66, x_99); +lean_inc_ref(x_65); +x_100 = l_Array_append___redArg(x_65, x_99); x_101 = lean_array_size(x_100); x_102 = 0; x_103 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Tactic_Do_elabMVCGen_spec__2___redArg(x_101, x_102, x_100, x_71, x_72, x_73, x_74); @@ -67576,7 +68230,7 @@ if (lean_obj_tag(x_111) == 0) lean_dec_ref(x_111); x_24 = x_99; x_25 = x_64; -x_26 = x_66; +x_26 = x_65; x_27 = x_67; x_28 = x_68; x_29 = x_69; @@ -67599,7 +68253,7 @@ lean_dec(x_70); lean_dec_ref(x_69); lean_dec(x_68); lean_dec_ref(x_67); -lean_dec_ref(x_66); +lean_dec_ref(x_65); lean_dec(x_64); return x_111; } @@ -67616,7 +68270,7 @@ lean_dec(x_70); lean_dec_ref(x_69); lean_dec(x_68); lean_dec_ref(x_67); -lean_dec_ref(x_66); +lean_dec_ref(x_65); lean_dec(x_64); x_112 = lean_ctor_get(x_103, 0); lean_inc(x_112); @@ -67649,7 +68303,7 @@ lean_dec(x_70); lean_dec_ref(x_69); lean_dec(x_68); lean_dec_ref(x_67); -lean_dec_ref(x_66); +lean_dec_ref(x_65); lean_dec(x_64); x_115 = !lean_is_exclusive(x_76); if (x_115 == 0) @@ -67671,8 +68325,8 @@ return x_117; block_152: { lean_object* x_134; lean_object* x_135; uint8_t x_136; -lean_inc(x_124); -x_134 = l_Lean_isTracingEnabledFor___at___00Lean_Elab_Tactic_Do_elabMVCGen_spec__1___redArg(x_124, x_120); +lean_inc(x_127); +x_134 = l_Lean_isTracingEnabledFor___at___00Lean_Elab_Tactic_Do_elabMVCGen_spec__1___redArg(x_127, x_123); x_135 = lean_ctor_get(x_134, 0); lean_inc(x_135); lean_dec_ref(x_134); @@ -67680,20 +68334,20 @@ x_136 = lean_unbox(x_135); lean_dec(x_135); if (x_136 == 0) { -lean_dec_ref(x_130); +lean_dec_ref(x_129); x_62 = x_119; x_63 = x_122; -x_64 = x_124; -x_65 = x_128; -x_66 = x_132; -x_67 = x_121; -x_68 = x_126; -x_69 = x_123; -x_70 = x_127; -x_71 = x_131; -x_72 = x_125; -x_73 = x_120; -x_74 = x_129; +x_64 = x_127; +x_65 = x_132; +x_66 = x_128; +x_67 = x_124; +x_68 = x_121; +x_69 = x_126; +x_70 = x_131; +x_71 = x_130; +x_72 = x_120; +x_73 = x_123; +x_74 = x_125; x_75 = lean_box(0); goto block_118; } @@ -67701,11 +68355,11 @@ else { lean_object* x_137; size_t x_138; size_t x_139; lean_object* x_140; lean_inc_ref(x_132); -x_137 = l_Array_append___redArg(x_132, x_130); -lean_dec_ref(x_130); +x_137 = l_Array_append___redArg(x_132, x_129); +lean_dec_ref(x_129); x_138 = lean_array_size(x_137); x_139 = 0; -x_140 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Tactic_Do_elabMVCGen_spec__2___redArg(x_138, x_139, x_137, x_131, x_125, x_120, x_129); +x_140 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Tactic_Do_elabMVCGen_spec__2___redArg(x_138, x_139, x_137, x_130, x_120, x_123, x_125); if (lean_obj_tag(x_140) == 0) { lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; @@ -67720,41 +68374,41 @@ x_146 = l_Lean_MessageData_ofList(x_145); x_147 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_147, 0, x_142); lean_ctor_set(x_147, 1, x_146); -lean_inc(x_124); -x_148 = l_Lean_addTrace___at___00Lean_Elab_Tactic_Do_elabMVCGen_spec__4___redArg(x_124, x_147, x_131, x_125, x_120, x_129); +lean_inc(x_127); +x_148 = l_Lean_addTrace___at___00Lean_Elab_Tactic_Do_elabMVCGen_spec__4___redArg(x_127, x_147, x_130, x_120, x_123, x_125); if (lean_obj_tag(x_148) == 0) { lean_dec_ref(x_148); x_62 = x_119; x_63 = x_122; -x_64 = x_124; -x_65 = x_128; -x_66 = x_132; -x_67 = x_121; -x_68 = x_126; -x_69 = x_123; -x_70 = x_127; -x_71 = x_131; -x_72 = x_125; -x_73 = x_120; -x_74 = x_129; +x_64 = x_127; +x_65 = x_132; +x_66 = x_128; +x_67 = x_124; +x_68 = x_121; +x_69 = x_126; +x_70 = x_131; +x_71 = x_130; +x_72 = x_120; +x_73 = x_123; +x_74 = x_125; x_75 = lean_box(0); goto block_118; } else { lean_dec_ref(x_132); -lean_dec_ref(x_131); -lean_dec(x_129); +lean_dec(x_131); +lean_dec_ref(x_130); lean_dec_ref(x_128); lean_dec(x_127); -lean_dec(x_126); +lean_dec_ref(x_126); lean_dec(x_125); -lean_dec(x_124); +lean_dec_ref(x_124); lean_dec_ref(x_123); lean_dec_ref(x_122); -lean_dec_ref(x_121); -lean_dec_ref(x_120); +lean_dec(x_121); +lean_dec(x_120); lean_dec_ref(x_119); return x_148; } @@ -67763,17 +68417,17 @@ else { uint8_t x_149; lean_dec_ref(x_132); -lean_dec_ref(x_131); -lean_dec(x_129); +lean_dec(x_131); +lean_dec_ref(x_130); lean_dec_ref(x_128); lean_dec(x_127); -lean_dec(x_126); +lean_dec_ref(x_126); lean_dec(x_125); -lean_dec(x_124); +lean_dec_ref(x_124); lean_dec_ref(x_123); lean_dec_ref(x_122); -lean_dec_ref(x_121); -lean_dec_ref(x_120); +lean_dec(x_121); +lean_dec(x_120); lean_dec_ref(x_119); x_149 = !lean_is_exclusive(x_140); if (x_149 == 0) @@ -67807,9 +68461,9 @@ x_121 = x_154; x_122 = x_156; x_123 = x_158; x_124 = x_157; -x_125 = x_159; +x_125 = x_161; x_126 = x_160; -x_127 = x_161; +x_127 = x_159; x_128 = x_162; x_129 = x_163; x_130 = x_164; @@ -67821,18 +68475,18 @@ goto block_152; else { uint8_t x_168; -lean_dec_ref(x_165); +lean_dec(x_165); lean_dec_ref(x_164); -lean_dec(x_163); +lean_dec_ref(x_163); lean_dec_ref(x_162); lean_dec(x_161); -lean_dec(x_160); +lean_dec_ref(x_160); lean_dec(x_159); lean_dec_ref(x_158); -lean_dec(x_157); +lean_dec_ref(x_157); lean_dec_ref(x_156); -lean_dec_ref(x_155); -lean_dec_ref(x_154); +lean_dec(x_155); +lean_dec(x_154); lean_dec_ref(x_153); x_168 = !lean_is_exclusive(x_166); if (x_168 == 0) @@ -67856,9 +68510,9 @@ return x_170; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; x_188 = lean_unsigned_to_nat(3u); x_189 = l_Lean_Syntax_getArg(x_1, x_188); -lean_inc_ref(x_178); +lean_inc_ref(x_177); x_190 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_suggestInvariant___boxed), 11, 1); -lean_closure_set(x_190, 0, x_178); +lean_closure_set(x_190, 0, x_177); lean_inc(x_186); lean_inc_ref(x_185); lean_inc(x_184); @@ -67874,24 +68528,24 @@ if (lean_obj_tag(x_191) == 0) lean_object* x_192; lean_object* x_193; uint8_t x_194; lean_dec_ref(x_191); x_192 = lean_array_get_size(x_172); -x_193 = lean_mk_empty_array_with_capacity(x_176); -x_194 = lean_nat_dec_lt(x_176, x_192); +x_193 = lean_mk_empty_array_with_capacity(x_178); +x_194 = lean_nat_dec_lt(x_178, x_192); if (x_194 == 0) { lean_dec(x_172); x_119 = x_173; -x_120 = x_185; -x_121 = x_179; +x_120 = x_184; +x_121 = x_180; x_122 = x_174; -x_123 = x_181; -x_124 = x_175; -x_125 = x_184; -x_126 = x_180; -x_127 = x_182; -x_128 = x_177; -x_129 = x_186; -x_130 = x_178; -x_131 = x_183; +x_123 = x_185; +x_124 = x_179; +x_125 = x_186; +x_126 = x_181; +x_127 = x_175; +x_128 = x_176; +x_129 = x_177; +x_130 = x_183; +x_131 = x_182; x_132 = x_193; x_133 = lean_box(0); goto block_152; @@ -67906,18 +68560,18 @@ if (x_194 == 0) { lean_dec(x_172); x_119 = x_173; -x_120 = x_185; -x_121 = x_179; +x_120 = x_184; +x_121 = x_180; x_122 = x_174; -x_123 = x_181; -x_124 = x_175; -x_125 = x_184; -x_126 = x_180; -x_127 = x_182; -x_128 = x_177; -x_129 = x_186; -x_130 = x_178; -x_131 = x_183; +x_123 = x_185; +x_124 = x_179; +x_125 = x_186; +x_126 = x_181; +x_127 = x_175; +x_128 = x_176; +x_129 = x_177; +x_130 = x_183; +x_131 = x_182; x_132 = x_193; x_133 = lean_box(0); goto block_152; @@ -67930,18 +68584,18 @@ x_197 = lean_usize_of_nat(x_192); x_198 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__10_spec__15(x_172, x_196, x_197, x_193, x_179, x_180, x_181, x_182, x_183, x_184, x_185, x_186); lean_dec(x_172); x_153 = x_173; -x_154 = x_179; -x_155 = x_185; +x_154 = x_180; +x_155 = x_184; x_156 = x_174; -x_157 = x_175; -x_158 = x_181; -x_159 = x_184; -x_160 = x_180; -x_161 = x_182; -x_162 = x_177; -x_163 = x_186; -x_164 = x_178; -x_165 = x_183; +x_157 = x_179; +x_158 = x_185; +x_159 = x_175; +x_160 = x_181; +x_161 = x_186; +x_162 = x_176; +x_163 = x_177; +x_164 = x_183; +x_165 = x_182; x_166 = x_198; goto block_171; } @@ -67954,18 +68608,18 @@ x_200 = lean_usize_of_nat(x_192); x_201 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Elab_Tactic_Do_elabInvariants_spec__10_spec__15(x_172, x_199, x_200, x_193, x_179, x_180, x_181, x_182, x_183, x_184, x_185, x_186); lean_dec(x_172); x_153 = x_173; -x_154 = x_179; -x_155 = x_185; +x_154 = x_180; +x_155 = x_184; x_156 = x_174; -x_157 = x_175; -x_158 = x_181; -x_159 = x_184; -x_160 = x_180; -x_161 = x_182; -x_162 = x_177; -x_163 = x_186; -x_164 = x_178; -x_165 = x_183; +x_157 = x_179; +x_158 = x_185; +x_159 = x_175; +x_160 = x_181; +x_161 = x_186; +x_162 = x_176; +x_163 = x_177; +x_164 = x_183; +x_165 = x_182; x_166 = x_201; goto block_171; } @@ -67981,8 +68635,8 @@ lean_dec(x_182); lean_dec_ref(x_181); lean_dec(x_180); lean_dec_ref(x_179); -lean_dec_ref(x_178); lean_dec_ref(x_177); +lean_dec_ref(x_176); lean_dec(x_175); lean_dec_ref(x_174); lean_dec_ref(x_173); @@ -67997,14 +68651,14 @@ x_221 = lean_ctor_get_uint8(x_208, sizeof(void*)*1); x_222 = lean_ctor_get_uint8(x_208, sizeof(void*)*1 + 1); lean_dec_ref(x_208); x_223 = lean_box(x_222); -x_224 = lean_box(x_203); +x_224 = lean_box(x_205); lean_inc_ref(x_2); -lean_inc_ref(x_205); +lean_inc_ref(x_203); x_225 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_elabMVCGen___lam__1___boxed), 12, 5); lean_closure_set(x_225, 0, x_223); lean_closure_set(x_225, 1, x_206); lean_closure_set(x_225, 2, x_224); -lean_closure_set(x_225, 3, x_205); +lean_closure_set(x_225, 3, x_203); lean_closure_set(x_225, 4, x_2); x_226 = 1; x_227 = lean_box(0); @@ -68016,7 +68670,7 @@ x_232 = lean_alloc_ctor(0, 8, 10); lean_ctor_set(x_232, 0, x_227); lean_ctor_set(x_232, 1, x_228); lean_ctor_set(x_232, 2, x_227); -lean_ctor_set(x_232, 3, x_210); +lean_ctor_set(x_232, 3, x_209); lean_ctor_set(x_232, 4, x_229); lean_ctor_set(x_232, 5, x_229); lean_ctor_set(x_232, 6, x_227); @@ -68025,11 +68679,11 @@ lean_ctor_set_uint8(x_232, sizeof(void*)*8, x_226); lean_ctor_set_uint8(x_232, sizeof(void*)*8 + 1, x_226); lean_ctor_set_uint8(x_232, sizeof(void*)*8 + 2, x_226); lean_ctor_set_uint8(x_232, sizeof(void*)*8 + 3, x_226); -lean_ctor_set_uint8(x_232, sizeof(void*)*8 + 4, x_209); -lean_ctor_set_uint8(x_232, sizeof(void*)*8 + 5, x_209); -lean_ctor_set_uint8(x_232, sizeof(void*)*8 + 6, x_209); +lean_ctor_set_uint8(x_232, sizeof(void*)*8 + 4, x_207); +lean_ctor_set_uint8(x_232, sizeof(void*)*8 + 5, x_207); +lean_ctor_set_uint8(x_232, sizeof(void*)*8 + 6, x_207); lean_ctor_set_uint8(x_232, sizeof(void*)*8 + 7, x_226); -lean_ctor_set_uint8(x_232, sizeof(void*)*8 + 8, x_209); +lean_ctor_set_uint8(x_232, sizeof(void*)*8 + 8, x_207); lean_ctor_set_uint8(x_232, sizeof(void*)*8 + 9, x_226); x_233 = l_Lean_Elab_Tactic_Do_elabMVCGen___lam__5___closed__7; lean_inc(x_219); @@ -68044,8 +68698,8 @@ lean_object* x_235; lean_object* x_236; lean_object* x_237; uint8_t x_238; x_235 = lean_ctor_get(x_234, 0); lean_inc(x_235); lean_dec_ref(x_234); -lean_inc(x_207); -x_236 = l_Lean_isTracingEnabledFor___at___00Lean_Elab_Tactic_Do_elabMVCGen_spec__1___redArg(x_207, x_218); +lean_inc(x_210); +x_236 = l_Lean_isTracingEnabledFor___at___00Lean_Elab_Tactic_Do_elabMVCGen_spec__1___redArg(x_210, x_218); x_237 = lean_ctor_get(x_236, 0); lean_inc(x_237); lean_dec_ref(x_236); @@ -68057,22 +68711,22 @@ x_239 = lean_ctor_get(x_235, 0); x_240 = lean_ctor_get(x_235, 1); lean_dec(x_240); x_241 = lean_box(x_222); -x_242 = lean_box(x_203); +x_242 = lean_box(x_205); lean_inc_ref(x_2); -lean_inc_ref(x_205); +lean_inc_ref(x_203); x_243 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_elabMVCGen___lam__2___boxed), 12, 4); lean_closure_set(x_243, 0, x_241); lean_closure_set(x_243, 1, x_242); -lean_closure_set(x_243, 2, x_205); +lean_closure_set(x_243, 2, x_203); lean_closure_set(x_243, 3, x_2); x_244 = lean_box(x_221); -x_245 = lean_box(x_203); +x_245 = lean_box(x_205); x_246 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_elabMVCGen___lam__3___boxed), 13, 6); lean_closure_set(x_246, 0, x_244); lean_closure_set(x_246, 1, x_243); lean_closure_set(x_246, 2, x_204); lean_closure_set(x_246, 3, x_245); -lean_closure_set(x_246, 4, x_205); +lean_closure_set(x_246, 4, x_203); lean_closure_set(x_246, 5, x_2); x_247 = lean_unbox(x_237); lean_dec(x_237); @@ -68082,10 +68736,10 @@ lean_free_object(x_235); x_172 = x_239; x_173 = x_246; x_174 = x_232; -x_175 = x_207; -x_176 = x_230; -x_177 = x_233; -x_178 = x_211; +x_175 = x_210; +x_176 = x_233; +x_177 = x_211; +x_178 = x_230; x_179 = x_212; x_180 = x_213; x_181 = x_214; @@ -68118,18 +68772,18 @@ x_256 = l_Lean_MessageData_ofList(x_255); lean_ctor_set_tag(x_235, 7); lean_ctor_set(x_235, 1, x_256); lean_ctor_set(x_235, 0, x_253); -lean_inc(x_207); -x_257 = l_Lean_addTrace___at___00Lean_Elab_Tactic_Do_elabMVCGen_spec__4___redArg(x_207, x_235, x_216, x_217, x_218, x_219); +lean_inc(x_210); +x_257 = l_Lean_addTrace___at___00Lean_Elab_Tactic_Do_elabMVCGen_spec__4___redArg(x_210, x_235, x_216, x_217, x_218, x_219); if (lean_obj_tag(x_257) == 0) { lean_dec_ref(x_257); x_172 = x_239; x_173 = x_246; x_174 = x_232; -x_175 = x_207; -x_176 = x_230; -x_177 = x_233; -x_178 = x_211; +x_175 = x_210; +x_176 = x_233; +x_177 = x_211; +x_178 = x_230; x_179 = x_212; x_180 = x_213; x_181 = x_214; @@ -68155,7 +68809,7 @@ lean_dec_ref(x_214); lean_dec(x_213); lean_dec_ref(x_212); lean_dec_ref(x_211); -lean_dec(x_207); +lean_dec(x_210); return x_257; } } @@ -68175,7 +68829,7 @@ lean_dec_ref(x_214); lean_dec(x_213); lean_dec_ref(x_212); lean_dec_ref(x_211); -lean_dec(x_207); +lean_dec(x_210); x_258 = !lean_is_exclusive(x_251); if (x_258 == 0) { @@ -68201,22 +68855,22 @@ x_261 = lean_ctor_get(x_235, 0); lean_inc(x_261); lean_dec(x_235); x_262 = lean_box(x_222); -x_263 = lean_box(x_203); +x_263 = lean_box(x_205); lean_inc_ref(x_2); -lean_inc_ref(x_205); +lean_inc_ref(x_203); x_264 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_elabMVCGen___lam__2___boxed), 12, 4); lean_closure_set(x_264, 0, x_262); lean_closure_set(x_264, 1, x_263); -lean_closure_set(x_264, 2, x_205); +lean_closure_set(x_264, 2, x_203); lean_closure_set(x_264, 3, x_2); x_265 = lean_box(x_221); -x_266 = lean_box(x_203); +x_266 = lean_box(x_205); x_267 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_elabMVCGen___lam__3___boxed), 13, 6); lean_closure_set(x_267, 0, x_265); lean_closure_set(x_267, 1, x_264); lean_closure_set(x_267, 2, x_204); lean_closure_set(x_267, 3, x_266); -lean_closure_set(x_267, 4, x_205); +lean_closure_set(x_267, 4, x_203); lean_closure_set(x_267, 5, x_2); x_268 = lean_unbox(x_237); lean_dec(x_237); @@ -68225,10 +68879,10 @@ if (x_268 == 0) x_172 = x_261; x_173 = x_267; x_174 = x_232; -x_175 = x_207; -x_176 = x_230; -x_177 = x_233; -x_178 = x_211; +x_175 = x_210; +x_176 = x_233; +x_177 = x_211; +x_178 = x_230; x_179 = x_212; x_180 = x_213; x_181 = x_214; @@ -68261,18 +68915,18 @@ x_277 = l_Lean_MessageData_ofList(x_276); x_278 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_278, 0, x_274); lean_ctor_set(x_278, 1, x_277); -lean_inc(x_207); -x_279 = l_Lean_addTrace___at___00Lean_Elab_Tactic_Do_elabMVCGen_spec__4___redArg(x_207, x_278, x_216, x_217, x_218, x_219); +lean_inc(x_210); +x_279 = l_Lean_addTrace___at___00Lean_Elab_Tactic_Do_elabMVCGen_spec__4___redArg(x_210, x_278, x_216, x_217, x_218, x_219); if (lean_obj_tag(x_279) == 0) { lean_dec_ref(x_279); x_172 = x_261; x_173 = x_267; x_174 = x_232; -x_175 = x_207; -x_176 = x_230; -x_177 = x_233; -x_178 = x_211; +x_175 = x_210; +x_176 = x_233; +x_177 = x_211; +x_178 = x_230; x_179 = x_212; x_180 = x_213; x_181 = x_214; @@ -68298,7 +68952,7 @@ lean_dec_ref(x_214); lean_dec(x_213); lean_dec_ref(x_212); lean_dec_ref(x_211); -lean_dec(x_207); +lean_dec(x_210); return x_279; } } @@ -68317,7 +68971,7 @@ lean_dec_ref(x_214); lean_dec(x_213); lean_dec_ref(x_212); lean_dec_ref(x_211); -lean_dec(x_207); +lean_dec(x_210); x_280 = lean_ctor_get(x_272, 0); lean_inc(x_280); if (lean_is_exclusive(x_272)) { @@ -68351,9 +69005,9 @@ lean_dec_ref(x_214); lean_dec(x_213); lean_dec_ref(x_212); lean_dec_ref(x_211); -lean_dec(x_207); -lean_dec_ref(x_205); +lean_dec(x_210); lean_dec_ref(x_204); +lean_dec_ref(x_203); lean_dec_ref(x_2); x_283 = !lean_is_exclusive(x_234); if (x_283 == 0) @@ -68404,14 +69058,14 @@ if (x_312 == 0) { lean_free_object(x_304); lean_inc_ref(x_307); -x_203 = x_287; +x_203 = x_308; x_204 = x_307; -x_205 = x_308; +x_205 = x_287; x_206 = x_306; -x_207 = x_309; -x_208 = x_289; -x_209 = x_291; -x_210 = x_292; +x_207 = x_289; +x_208 = x_291; +x_209 = x_292; +x_210 = x_309; x_211 = x_307; x_212 = x_294; x_213 = x_295; @@ -68451,14 +69105,14 @@ if (lean_obj_tag(x_323) == 0) { lean_dec_ref(x_323); lean_inc_ref(x_307); -x_203 = x_287; +x_203 = x_308; x_204 = x_307; -x_205 = x_308; +x_205 = x_287; x_206 = x_306; -x_207 = x_309; -x_208 = x_289; -x_209 = x_291; -x_210 = x_292; +x_207 = x_289; +x_208 = x_291; +x_209 = x_292; +x_210 = x_309; x_211 = x_307; x_212 = x_294; x_213 = x_295; @@ -68484,7 +69138,7 @@ lean_dec_ref(x_296); lean_dec(x_295); lean_dec_ref(x_294); lean_dec_ref(x_292); -lean_dec_ref(x_289); +lean_dec_ref(x_291); lean_dec_ref(x_2); return x_323; } @@ -68504,7 +69158,7 @@ lean_dec_ref(x_296); lean_dec(x_295); lean_dec_ref(x_294); lean_dec_ref(x_292); -lean_dec_ref(x_289); +lean_dec_ref(x_291); lean_dec_ref(x_2); x_324 = !lean_is_exclusive(x_316); if (x_324 == 0) @@ -68543,14 +69197,14 @@ lean_dec(x_332); if (x_333 == 0) { lean_inc_ref(x_328); -x_203 = x_287; +x_203 = x_329; x_204 = x_328; -x_205 = x_329; +x_205 = x_287; x_206 = x_327; -x_207 = x_330; -x_208 = x_289; -x_209 = x_291; -x_210 = x_292; +x_207 = x_289; +x_208 = x_291; +x_209 = x_292; +x_210 = x_330; x_211 = x_328; x_212 = x_294; x_213 = x_295; @@ -68590,14 +69244,14 @@ if (lean_obj_tag(x_345) == 0) { lean_dec_ref(x_345); lean_inc_ref(x_328); -x_203 = x_287; +x_203 = x_329; x_204 = x_328; -x_205 = x_329; +x_205 = x_287; x_206 = x_327; -x_207 = x_330; -x_208 = x_289; -x_209 = x_291; -x_210 = x_292; +x_207 = x_289; +x_208 = x_291; +x_209 = x_292; +x_210 = x_330; x_211 = x_328; x_212 = x_294; x_213 = x_295; @@ -68623,7 +69277,7 @@ lean_dec_ref(x_296); lean_dec(x_295); lean_dec_ref(x_294); lean_dec_ref(x_292); -lean_dec_ref(x_289); +lean_dec_ref(x_291); lean_dec_ref(x_2); return x_345; } @@ -68642,7 +69296,7 @@ lean_dec_ref(x_296); lean_dec(x_295); lean_dec_ref(x_294); lean_dec_ref(x_292); -lean_dec_ref(x_289); +lean_dec_ref(x_291); lean_dec_ref(x_2); x_346 = lean_ctor_get(x_337, 0); lean_inc(x_346); @@ -68676,7 +69330,7 @@ lean_dec_ref(x_296); lean_dec(x_295); lean_dec_ref(x_294); lean_dec_ref(x_292); -lean_dec_ref(x_289); +lean_dec_ref(x_291); lean_dec_ref(x_2); x_349 = !lean_is_exclusive(x_303); if (x_349 == 0) @@ -68698,7 +69352,7 @@ return x_351; block_380: { lean_object* x_368; -x_368 = l_Lean_Elab_Tactic_getMainGoal___redArg(x_357, x_360, x_365, x_366, x_355); +x_368 = l_Lean_Elab_Tactic_getMainGoal___redArg(x_362, x_363, x_357, x_358, x_354); if (lean_obj_tag(x_368) == 0) { uint8_t x_369; @@ -68710,20 +69364,20 @@ x_370 = lean_ctor_get(x_368, 0); lean_inc(x_370); lean_dec_ref(x_368); x_287 = x_353; -x_288 = x_354; -x_289 = x_361; +x_288 = x_359; +x_289 = x_360; x_290 = x_367; -x_291 = x_356; -x_292 = x_363; +x_291 = x_361; +x_292 = x_355; x_293 = x_370; -x_294 = x_362; -x_295 = x_357; -x_296 = x_359; -x_297 = x_364; -x_298 = x_360; -x_299 = x_365; -x_300 = x_366; -x_301 = x_355; +x_294 = x_366; +x_295 = x_362; +x_296 = x_365; +x_297 = x_356; +x_298 = x_363; +x_299 = x_357; +x_300 = x_358; +x_301 = x_354; x_302 = lean_box(0); goto block_352; } @@ -68733,11 +69387,11 @@ lean_object* x_371; lean_object* x_372; x_371 = lean_ctor_get(x_368, 0); lean_inc(x_371); lean_dec_ref(x_368); -lean_inc(x_355); -lean_inc_ref(x_366); -lean_inc(x_365); -lean_inc_ref(x_360); -x_372 = l_Lean_Elab_Tactic_Do_elimLets(x_371, x_369, x_360, x_365, x_366, x_355); +lean_inc(x_354); +lean_inc_ref(x_358); +lean_inc(x_357); +lean_inc_ref(x_363); +x_372 = l_Lean_Elab_Tactic_Do_elimLets(x_371, x_369, x_363, x_357, x_358, x_354); if (lean_obj_tag(x_372) == 0) { lean_object* x_373; @@ -68745,20 +69399,20 @@ x_373 = lean_ctor_get(x_372, 0); lean_inc(x_373); lean_dec_ref(x_372); x_287 = x_353; -x_288 = x_354; -x_289 = x_361; +x_288 = x_359; +x_289 = x_360; x_290 = x_367; -x_291 = x_356; -x_292 = x_363; +x_291 = x_361; +x_292 = x_355; x_293 = x_373; -x_294 = x_362; -x_295 = x_357; -x_296 = x_359; -x_297 = x_364; -x_298 = x_360; -x_299 = x_365; -x_300 = x_366; -x_301 = x_355; +x_294 = x_366; +x_295 = x_362; +x_296 = x_365; +x_297 = x_356; +x_298 = x_363; +x_299 = x_357; +x_300 = x_358; +x_301 = x_354; x_302 = lean_box(0); goto block_352; } @@ -68767,16 +69421,16 @@ else uint8_t x_374; lean_dec(x_367); lean_dec_ref(x_366); -lean_dec(x_365); -lean_dec(x_364); +lean_dec_ref(x_365); lean_dec_ref(x_363); -lean_dec_ref(x_362); +lean_dec(x_362); lean_dec_ref(x_361); -lean_dec_ref(x_360); lean_dec_ref(x_359); +lean_dec_ref(x_358); lean_dec(x_357); -lean_dec(x_355); -lean_dec_ref(x_354); +lean_dec(x_356); +lean_dec_ref(x_355); +lean_dec(x_354); lean_dec_ref(x_2); x_374 = !lean_is_exclusive(x_372); if (x_374 == 0) @@ -68801,16 +69455,16 @@ else uint8_t x_377; lean_dec(x_367); lean_dec_ref(x_366); -lean_dec(x_365); -lean_dec(x_364); +lean_dec_ref(x_365); lean_dec_ref(x_363); -lean_dec_ref(x_362); +lean_dec(x_362); lean_dec_ref(x_361); -lean_dec_ref(x_360); lean_dec_ref(x_359); +lean_dec_ref(x_358); lean_dec(x_357); -lean_dec(x_355); -lean_dec_ref(x_354); +lean_dec(x_356); +lean_dec_ref(x_355); +lean_dec(x_354); lean_dec_ref(x_2); x_377 = !lean_is_exclusive(x_368); if (x_377 == 0) @@ -68865,19 +69519,19 @@ if (lean_obj_tag(x_398) == 0) lean_object* x_401; x_401 = lean_box(1); x_353 = x_394; -x_354 = x_396; -x_355 = x_388; -x_356 = x_394; -x_357 = x_382; -x_358 = lean_box(0); -x_359 = x_383; -x_360 = x_385; +x_354 = x_388; +x_355 = x_400; +x_356 = x_384; +x_357 = x_386; +x_358 = x_387; +x_359 = x_396; +x_360 = x_394; x_361 = x_397; -x_362 = x_381; -x_363 = x_400; -x_364 = x_384; -x_365 = x_386; -x_366 = x_387; +x_362 = x_382; +x_363 = x_385; +x_364 = lean_box(0); +x_365 = x_383; +x_366 = x_381; x_367 = x_401; goto block_380; } @@ -68889,19 +69543,19 @@ if (x_402 == 0) { lean_ctor_set_tag(x_398, 0); x_353 = x_394; -x_354 = x_396; -x_355 = x_388; -x_356 = x_394; -x_357 = x_382; -x_358 = lean_box(0); -x_359 = x_383; -x_360 = x_385; +x_354 = x_388; +x_355 = x_400; +x_356 = x_384; +x_357 = x_386; +x_358 = x_387; +x_359 = x_396; +x_360 = x_394; x_361 = x_397; -x_362 = x_381; -x_363 = x_400; -x_364 = x_384; -x_365 = x_386; -x_366 = x_387; +x_362 = x_382; +x_363 = x_385; +x_364 = lean_box(0); +x_365 = x_383; +x_366 = x_381; x_367 = x_398; goto block_380; } @@ -68914,19 +69568,19 @@ lean_dec(x_398); x_404 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_404, 0, x_403); x_353 = x_394; -x_354 = x_396; -x_355 = x_388; -x_356 = x_394; -x_357 = x_382; -x_358 = lean_box(0); -x_359 = x_383; -x_360 = x_385; +x_354 = x_388; +x_355 = x_400; +x_356 = x_384; +x_357 = x_386; +x_358 = x_387; +x_359 = x_396; +x_360 = x_394; x_361 = x_397; -x_362 = x_381; -x_363 = x_400; -x_364 = x_384; -x_365 = x_386; -x_366 = x_387; +x_362 = x_382; +x_363 = x_385; +x_364 = lean_box(0); +x_365 = x_383; +x_366 = x_381; x_367 = x_404; goto block_380; } @@ -69769,8 +70423,8 @@ l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_T lean_mark_persistent(l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13___redArg___closed__0); l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13___redArg___closed__1 = _init_l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13___redArg___closed__1(); lean_mark_persistent(l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13___redArg___closed__1); -l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___closed__0 = _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___closed__0(); -lean_mark_persistent(l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___closed__0); +l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___closed__0 = _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___closed__0(); +lean_mark_persistent(l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___closed__0); l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13___redArg___closed__2 = _init_l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13___redArg___closed__2(); l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13___redArg___closed__3 = _init_l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13___redArg___closed__3(); lean_mark_persistent(l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13___redArg___closed__3); @@ -69787,10 +70441,10 @@ l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_T lean_mark_persistent(l_Lean_withTraceNode___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__13___redArg___closed__6); l_Lean_MVarId_isAssigned___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__6___closed__0 = _init_l_Lean_MVarId_isAssigned___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__6___closed__0(); lean_mark_persistent(l_Lean_MVarId_isAssigned___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__6___closed__0); -l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__0___closed__1 = _init_l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__0___closed__1(); -lean_mark_persistent(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__0___closed__1); -l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__0___closed__0 = _init_l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__0___closed__0(); -lean_mark_persistent(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__40___lam__0___closed__0); +l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__0___closed__1 = _init_l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__0___closed__1(); +lean_mark_persistent(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__0___closed__1); +l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__0___closed__0 = _init_l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__0___closed__0(); +lean_mark_persistent(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__41___lam__0___closed__0); l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__11___closed__0 = _init_l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__11___closed__0(); lean_mark_persistent(l_Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10___lam__11___closed__0); l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__3___closed__0 = _init_l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__3___closed__0(); @@ -69811,22 +70465,22 @@ l_List_mapTR_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tact lean_mark_persistent(l_List_mapTR_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__8___closed__2); l_List_mapTR_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__8___closed__3 = _init_l_List_mapTR_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__8___closed__3(); lean_mark_persistent(l_List_mapTR_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__8___closed__3); -l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__0 = _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__0(); -lean_mark_persistent(l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__0); -l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__1 = _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__1(); -lean_mark_persistent(l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__1); -l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__2 = _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__2(); -lean_mark_persistent(l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__2); -l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__3 = _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__3(); -lean_mark_persistent(l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__3); -l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__4 = _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__4(); -lean_mark_persistent(l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__4); -l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__5 = _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__5(); -lean_mark_persistent(l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__5); -l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__6 = _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__6(); -lean_mark_persistent(l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__6); -l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__7 = _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__7(); -lean_mark_persistent(l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__44___redArg___closed__7); +l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__0 = _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__0(); +lean_mark_persistent(l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__0); +l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__1 = _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__1(); +lean_mark_persistent(l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__1); +l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__2 = _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__2(); +lean_mark_persistent(l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__2); +l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__3 = _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__3(); +lean_mark_persistent(l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__3); +l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__4 = _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__4(); +lean_mark_persistent(l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__4); +l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__5 = _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__5(); +lean_mark_persistent(l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__5); +l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__6 = _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__6(); +lean_mark_persistent(l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__6); +l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__7 = _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__7(); +lean_mark_persistent(l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_assignMVars_spec__22_spec__45___redArg___closed__7); l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__16___closed__2 = _init_l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__16___closed__2(); lean_mark_persistent(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__16___closed__2); l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__9___closed__0 = _init_l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__9___closed__0(); @@ -69903,8 +70557,8 @@ l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails lean_mark_persistent(l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___closed__9); l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___closed__10 = _init_l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___closed__10(); lean_mark_persistent(l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20___closed__10); -l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__8 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__8(); -lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__8); +l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__8 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__8(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__8); l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___closed__5 = _init_l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___closed__5(); lean_mark_persistent(l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___closed__5); l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___closed__6 = _init_l___private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargeFailEntails___at___00__private_Lean_Elab_Tactic_Do_Spec_0__Lean_Elab_Tactic_Do_dischargePostEntails___at___00Lean_Elab_Tactic_Do_mSpec___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp_spec__10_spec__20_spec__26___closed__6(); @@ -69990,112 +70644,132 @@ l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPAp lean_mark_persistent(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__16___closed__15); l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__16___closed__16 = _init_l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__16___closed__16(); lean_mark_persistent(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onWPApp___lam__16___closed__16); -l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59___lam__1___closed__0 = _init_l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59___lam__1___closed__0(); -lean_mark_persistent(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59___lam__1___closed__0); -l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59___lam__1___closed__1 = _init_l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59___lam__1___closed__1(); -lean_mark_persistent(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__59___lam__1___closed__1); -l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62___lam__1___closed__0 = _init_l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62___lam__1___closed__0(); -lean_mark_persistent(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62___lam__1___closed__0); -l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62___lam__1___closed__1 = _init_l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62___lam__1___closed__1(); -lean_mark_persistent(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__62___lam__1___closed__1); -l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__0 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__0(); -lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__0); -l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__1 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__1); -l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__2 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__2(); -lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__2); -l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__3 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__3(); -lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__3); -l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__0 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__0(); -lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__0); -l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__1 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__1(); -lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__1); -l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__0 = _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__0(); -lean_mark_persistent(l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__0); -l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__1 = _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__1(); -lean_mark_persistent(l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__1); -l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__2 = _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__2(); -lean_mark_persistent(l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__2); -l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__3 = _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__3(); -lean_mark_persistent(l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__3); -l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__4 = _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__4(); -lean_mark_persistent(l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__4); -l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__5 = _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__5(); -lean_mark_persistent(l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___lam__2___closed__5); -l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___closed__1 = _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___closed__1(); -lean_mark_persistent(l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__52_spec__59_spec__70___closed__1); -l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__0 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__0(); -lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__0); -l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__1 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__1(); -lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__1); -l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__0 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__0(); -lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__0); -l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__1 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__1(); -lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___closed__1); -l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___closed__0 = _init_l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___closed__0(); -lean_mark_persistent(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___closed__0); -l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___closed__1 = _init_l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___closed__1(); -lean_mark_persistent(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__53___closed__1); -l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__2 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__2(); -lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__2); -l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__3 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__3(); -lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__3); -l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__3___closed__0 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__3___closed__0(); -lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__3___closed__0); -l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__3___closed__1 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__3___closed__1(); -lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__3___closed__1); -l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__5 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__5(); -lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__5); -l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__4 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__4(); -lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__4); -l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__6 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__6(); -lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__6); -l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__7 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__7(); -lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__7); -l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__9 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__9(); -lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__9); -l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__10 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__10(); -lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__56___redArg___lam__7___closed__10); -l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__2 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__2(); -lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__2); -l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__3 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__3(); -lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__3); -l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__4 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__4(); -lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__4); -l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__5 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__5(); -lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__5); -l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__6 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__6(); -lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__6); -l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__7 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__7(); -lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__7); -l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__8 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__8(); -lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__8); -l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___boxed__const__1 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___boxed__const__1(); -lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___boxed__const__1); -l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___closed__0 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___closed__0(); -lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___closed__0); -l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___closed__0 = _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___closed__0(); -lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37_spec__47___closed__0); -l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___closed__1 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___closed__1(); -lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___closed__1); -l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___closed__2 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___closed__2(); -lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___closed__2); -l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___closed__3 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___closed__3(); -lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___closed__3); -l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___closed__4 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___closed__4(); -lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___lam__13___closed__4); -l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__9 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__9(); -lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__9); -l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__10 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__10(); -lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__10); -l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__11 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__11(); -lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__11); -l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__12 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__12(); -lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__37___closed__12); +l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__1___closed__0 = _init_l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__1___closed__0(); +lean_mark_persistent(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__1___closed__0); +l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__1___closed__1 = _init_l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__1___closed__1(); +lean_mark_persistent(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__60___lam__1___closed__1); +l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__63___lam__1___closed__0 = _init_l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__63___lam__1___closed__0(); +lean_mark_persistent(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__63___lam__1___closed__0); +l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__63___lam__1___closed__1 = _init_l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__63___lam__1___closed__1(); +lean_mark_persistent(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00Lean_Elab_Tactic_Do_ProofMode_mFrameCore___at___00Lean_Elab_Tactic_Do_ProofMode_mTryFrame___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJumpSite_spec__16_spec__28_spec__33_spec__38___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__63___lam__1___closed__1); +l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__0 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__0(); +lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__0); +l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__1 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__1); +l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__2 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__2); +l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__3 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18___closed__3); +l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__0 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__0(); +lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__0); +l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__1 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__1(); +lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__1); +l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__0 = _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__0(); +lean_mark_persistent(l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__0); +l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__1 = _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__1(); +lean_mark_persistent(l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__1); +l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__2 = _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__2(); +lean_mark_persistent(l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__2); +l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__3 = _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__3(); +lean_mark_persistent(l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__3); +l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__4 = _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__4(); +lean_mark_persistent(l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__4); +l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__5 = _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__5(); +lean_mark_persistent(l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___lam__2___closed__5); +l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___closed__1 = _init_l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___closed__1(); +lean_mark_persistent(l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__53_spec__60_spec__71___closed__1); +l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__0 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__0(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__0); +l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__1 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__1(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__1); +l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__0 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__0(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__0); +l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__1 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__1(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___closed__1); +l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___closed__0 = _init_l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___closed__0(); +lean_mark_persistent(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___closed__0); +l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___closed__1 = _init_l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___closed__1(); +lean_mark_persistent(l_panic___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__54___closed__1); +l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__2 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__2(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__2); +l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__3 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__3(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__3); +l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__3___closed__0 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__3___closed__0(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__3___closed__0); +l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__3___closed__1 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__3___closed__1(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__3___closed__1); +l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__5 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__5(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__5); +l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__4 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__4(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__4); +l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__6 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__6(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__6); +l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__7 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__7(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__7); +l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__9 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__9(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__9); +l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__10 = _init_l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__10(); +lean_mark_persistent(l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__57___redArg___lam__7___closed__10); +l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__2 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__2(); +lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__2); +l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__3 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__3(); +lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__3); +l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__4 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__4(); +lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__4); +l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__5 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__5(); +lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__5); +l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__6 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__6(); +lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__6); +l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__7 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__7(); +lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__7); +l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__8 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__8(); +lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__8); +l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___boxed__const__1 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___boxed__const__1(); +lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___boxed__const__1); +l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___closed__0 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___closed__0(); +lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___closed__0); +l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___closed__0 = _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___closed__0(); +lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38_spec__48___closed__0); +l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___closed__1 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___closed__1(); +lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___closed__1); +l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___closed__2 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___closed__2(); +lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___closed__2); +l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___closed__3 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___closed__3(); +lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___closed__3); +l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___closed__4 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___closed__4(); +lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___lam__13___closed__4); +l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__9 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__9(); +lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__9); +l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__10 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__10(); +lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__10); +l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__11 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__11(); +lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__11); +l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__12 = _init_l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__12(); +lean_mark_persistent(l_Lean_Meta_MatcherApp_transform___at___00Lean_Elab_Tactic_Do_SplitInfo_splitWith___at___00__private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint_spec__18_spec__38___closed__12); l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___closed__0 = _init_l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___closed__0(); lean_mark_persistent(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___closed__0); l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___closed__1 = _init_l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onJoinPoint___closed__1); +l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__0 = _init_l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__0(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__0); +l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__1 = _init_l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__1(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__1); +l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__2 = _init_l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__2(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__2); +l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__3 = _init_l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__3(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__3); +l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__4 = _init_l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__4(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__4); +l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__5 = _init_l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__5(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__5); +l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__6 = _init_l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__6(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__6); +l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__7 = _init_l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__7(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__7); +l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__8 = _init_l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__8(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__8); +l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__9 = _init_l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__9(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___lam__7___closed__9); l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___closed__0 = _init_l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___closed__0(); lean_mark_persistent(l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___closed__0); l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___closed__1 = _init_l___private_Lean_Elab_Tactic_Do_VCGen_0__Lean_Elab_Tactic_Do_VCGen_genVCs_onSplit___closed__1(); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Do/VCGen/Basic.c b/stage0/stdlib/Lean/Elab/Tactic/Do/VCGen/Basic.c index 654d81018e7f..c45b182097ba 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Do/VCGen/Basic.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Do/VCGen/Basic.c @@ -44,6 +44,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_freshenUserNam static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Tactic_Do_mkSpecContext_spec__4___redArg___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_Fuel_ctorIdx___boxed(lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__7_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_540456248____hygCtx___hyg_2_; +static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__1_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_; lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLetDeclImp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_withJP___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_Raw_u2080_Const_get_x3f___at___00Lean_recordExtraModUseFromDecl___at___00Lean_Elab_Tactic_Do_elabConfig_spec__0_spec__2___redArg(lean_object*, lean_object*); @@ -117,6 +118,7 @@ static lean_object* l_Lean_Elab_Tactic_Do_initFn___closed__0_00___x40_Lean_Elab_ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_burnOne___redArg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_recordExtraModUseFromDecl___at___00Lean_Elab_Tactic_Do_elabConfig_spec__0_spec__2_spec__6___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_recordExtraModUseFromDecl___at___00Lean_Elab_Tactic_Do_elabConfig_spec__0___closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_(); lean_object* lean_array_fget_borrowed(lean_object*, lean_object*); static lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_Elab_Tactic_Do_mkSpecContext_spec__3_spec__3___redArg___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_freshenLCtxUserNamesSinceIdx___at___00Lean_Elab_Tactic_Do_addSubGoalAsVC_spec__0_spec__1_spec__3_spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -124,6 +126,7 @@ uint8_t l_Lean_Expr_isAppOf(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at___00Lean_MVarId_freshenLCtxUserNamesSinceIdx___at___00Lean_Elab_Tactic_Do_addSubGoalAsVC_spec__0_spec__1___redArg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_instMonadLiftSimpMVCGenM; static lean_object* l_Lean_recordExtraModUseFromDecl___at___00Lean_Elab_Tactic_Do_elabConfig_spec__0___closed__3; +lean_object* l_Lean_Name_mkStr5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_withLetDeclShared(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_instDecidableEqFuel___boxed(lean_object*, lean_object*); static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_Elab_Tactic_Do_mkSpecContext_spec__3_spec__3_spec__4_spec__7_spec__9___redArg___closed__5; @@ -159,6 +162,7 @@ lean_object* l_Array_mkArray0(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_initFn_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_1593450264____hygCtx___hyg_4_(); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_540456248____hygCtx___hyg_2_(); static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Tactic_Do_mkSpecContext_spec__4___redArg___closed__2; +static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__0_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_; uint8_t l_Lean_Expr_hasSyntheticSorry(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_MVarId_freshenLCtxUserNamesSinceIdx___at___00Lean_Elab_Tactic_Do_addSubGoalAsVC_spec__0_spec__2___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___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_recordExtraModUseFromDecl___at___00Lean_Elab_Tactic_Do_elabConfig_spec__0_spec__2_spec__6___redArg(lean_object*, lean_object*); @@ -188,6 +192,7 @@ LEAN_EXPORT lean_object* l_Lean_throwUnknownIdentifierAt___at___00Lean_throwUnkn static lean_object* l_Lean_Elab_Tactic_Do_mkSpecContext___closed__7; static size_t l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00__private_Lean_ExtraModUses_0__Lean_recordExtraModUseCore___at___00Lean_recordExtraModUseFromDecl___at___00Lean_Elab_Tactic_Do_elabConfig_spec__0_spec__0_spec__1_spec__4___redArg___closed__0; static lean_object* l___private_Lean_ExtraModUses_0__Lean_recordExtraModUseCore___at___00Lean_recordExtraModUseFromDecl___at___00Lean_Elab_Tactic_Do_elabConfig_spec__0_spec__0___closed__0; +static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__2_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_; static lean_object* l_Lean_MVarId_freshenLCtxUserNamesSinceIdx___at___00Lean_Elab_Tactic_Do_addSubGoalAsVC_spec__0___closed__0; static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Tactic_Do_mkSpecContext_spec__4___redArg___closed__4; static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_Elab_Tactic_Do_mkSpecContext_spec__3_spec__3_spec__4_spec__7_spec__9___redArg___closed__1; @@ -263,6 +268,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_instMonadLiftSimpMVCGenM___lam__0 lean_object* lean_st_mk_ref(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at___00Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00__private_Lean_ExtraModUses_0__Lean_recordExtraModUseCore___at___00Lean_recordExtraModUseFromDecl___at___00Lean_Elab_Tactic_Do_elabConfig_spec__0_spec__0_spec__1_spec__4_spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__4_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_540456248____hygCtx___hyg_2_; +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2____boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_addMacroStack___at___00Lean_throwError___at___00Lean_Elab_Tactic_Do_elabConfig_spec__1_spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_Elab_Tactic_Do_mkSpecContext_spec__3_spec__3_spec__4_spec__7_spec__9___redArg___closed__10; @@ -420,6 +426,7 @@ static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknow lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_unfoldDefinition_x3f(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__34_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_540456248____hygCtx___hyg_2_; +static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__5_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_; static lean_object* l_Lean_Elab_Tactic_Do_mkSpecContext___closed__14; static lean_object* l_Lean_Elab_Tactic_Do_mkSpecContext___closed__21; LEAN_EXPORT lean_object* l_Lean_Elab_addMacroStack___at___00Lean_throwError___at___00Lean_Elab_Tactic_Do_elabConfig_spec__1_spec__5___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -453,6 +460,7 @@ lean_object* l_Lean_Meta_Simp_mkDefaultMethodsCore(lean_object*); uint8_t l_Lean_isPrivateName(lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_addSubGoalAsVC___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_withJP(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__3_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_; lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); lean_object* l_Lean_PersistentEnvExtension_addEntry___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Elab_Tactic_Do_SpecAttr_SpecTheorems_isErased(lean_object*, lean_object*); @@ -483,6 +491,7 @@ LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkU lean_object* l_Lean_PersistentHashMap_empty(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___00Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_freshenLCtxUserNamesSinceIdx___at___00Lean_Elab_Tactic_Do_addSubGoalAsVC_spec__0_spec__1_spec__3_spec__4_spec__6___redArg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Option_get___at___00Lean_Elab_addMacroStack___at___00Lean_throwError___at___00Lean_Elab_Tactic_Do_elabConfig_spec__1_spec__5_spec__10___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__4_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_; static lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__2_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_540456248____hygCtx___hyg_2_; lean_object* l_Lean_Expr_betaRev(lean_object*, lean_object*, uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___00Lean_Elab_Tactic_Do_withLetDeclShared_spec__0___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -934,6 +943,86 @@ x_2 = l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn_ return x_2; } } +static lean_object* _init_l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__0_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("split", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__1_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__0_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_; +x_2 = l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__3_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_540456248____hygCtx___hyg_2_; +x_3 = l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__2_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_540456248____hygCtx___hyg_2_; +x_4 = l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__1_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_540456248____hygCtx___hyg_2_; +x_5 = l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__0_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_540456248____hygCtx___hyg_2_; +x_6 = l_Lean_Name_mkStr5(x_5, x_4, x_3, x_2, x_1); +return x_6; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__2_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(4153898153u); +x_2 = l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__30_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_540456248____hygCtx___hyg_2_; +x_3 = l_Lean_Name_num___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__3_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__32_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_540456248____hygCtx___hyg_2_; +x_2 = l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__2_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_; +x_3 = l_Lean_Name_str___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__4_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__34_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_540456248____hygCtx___hyg_2_; +x_2 = l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__3_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_; +x_3 = l_Lean_Name_str___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__5_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(2u); +x_2 = l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__4_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_; +x_3 = l_Lean_Name_num___override(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_() { +_start: +{ +lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; +x_2 = l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__1_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_; +x_3 = 0; +x_4 = l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__5_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_; +x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2____boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_(); +return x_2; +} +} LEAN_EXPORT lean_object* l_Lean_Option_register___at___00Lean_Elab_Tactic_Do_initFn_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_1593450264____hygCtx___hyg_4__spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { @@ -3298,9 +3387,9 @@ return x_128; if (x_20 == 0) { lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -lean_dec_ref(x_16); +lean_dec_ref(x_17); x_21 = l_Lean_Elab_Tactic_Do_elabConfig___redArg___closed__1; -x_22 = l_Lean_MessageData_ofExpr(x_19); +x_22 = l_Lean_MessageData_ofExpr(x_18); x_23 = l_Lean_indentD(x_22); x_24 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_24, 0, x_21); @@ -3309,29 +3398,29 @@ x_25 = l_Lean_Elab_Tactic_Do_elabConfig___redArg___closed__3; x_26 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_26, 0, x_24); lean_ctor_set(x_26, 1, x_25); -x_27 = l_Lean_Exception_toMessageData(x_12); +x_27 = l_Lean_Exception_toMessageData(x_10); x_28 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_28, 0, x_26); lean_ctor_set(x_28, 1, x_27); -x_29 = l_Lean_throwError___at___00Lean_Elab_Tactic_Do_elabConfig_spec__1___redArg(x_28, x_15, x_14, x_11, x_13, x_18, x_17); -lean_dec(x_17); -lean_dec_ref(x_18); +x_29 = l_Lean_throwError___at___00Lean_Elab_Tactic_Do_elabConfig_spec__1___redArg(x_28, x_19, x_16, x_14, x_12, x_15, x_13); lean_dec(x_13); -lean_dec_ref(x_11); -lean_dec(x_14); +lean_dec_ref(x_15); +lean_dec(x_12); +lean_dec_ref(x_14); +lean_dec(x_16); return x_29; } else { lean_dec_ref(x_19); lean_dec_ref(x_18); -lean_dec(x_17); +lean_dec(x_16); lean_dec_ref(x_15); -lean_dec(x_14); +lean_dec_ref(x_14); lean_dec(x_13); -lean_dec_ref(x_12); -lean_dec_ref(x_11); -return x_16; +lean_dec(x_12); +lean_dec_ref(x_10); +return x_17; } } block_43: @@ -3365,31 +3454,31 @@ if (x_41 == 0) uint8_t x_42; lean_inc(x_40); x_42 = l_Lean_Exception_isRuntime(x_40); -x_10 = lean_box(0); -x_11 = x_34; -x_12 = x_40; -x_13 = x_35; -x_14 = x_33; -x_15 = x_32; -x_16 = x_39; -x_17 = x_37; -x_18 = x_36; -x_19 = x_31; +x_10 = x_40; +x_11 = lean_box(0); +x_12 = x_35; +x_13 = x_37; +x_14 = x_34; +x_15 = x_36; +x_16 = x_33; +x_17 = x_39; +x_18 = x_31; +x_19 = x_32; x_20 = x_42; goto block_30; } else { -x_10 = lean_box(0); -x_11 = x_34; -x_12 = x_40; -x_13 = x_35; -x_14 = x_33; -x_15 = x_32; -x_16 = x_39; -x_17 = x_37; -x_18 = x_36; -x_19 = x_31; +x_10 = x_40; +x_11 = lean_box(0); +x_12 = x_35; +x_13 = x_37; +x_14 = x_34; +x_15 = x_36; +x_16 = x_33; +x_17 = x_39; +x_18 = x_31; +x_19 = x_32; x_20 = x_41; goto block_30; } @@ -8798,17 +8887,17 @@ if (x_159 == 0) uint8_t x_160; lean_inc(x_158); x_160 = l_Lean_Exception_isRuntime(x_158); -x_52 = lean_box(0); +x_52 = x_158; x_53 = x_153; -x_54 = x_158; +x_54 = lean_box(0); x_55 = x_160; goto block_62; } else { -x_52 = lean_box(0); +x_52 = x_158; x_53 = x_153; -x_54 = x_158; +x_54 = lean_box(0); x_55 = x_159; goto block_62; } @@ -8928,16 +9017,16 @@ if (x_176 == 0) uint8_t x_177; lean_inc(x_175); x_177 = l_Lean_Exception_isRuntime(x_175); -x_63 = x_170; -x_64 = lean_box(0); +x_63 = lean_box(0); +x_64 = x_170; x_65 = x_175; x_66 = x_177; goto block_73; } else { -x_63 = x_170; -x_64 = lean_box(0); +x_63 = lean_box(0); +x_64 = x_170; x_65 = x_175; x_66 = x_176; goto block_73; @@ -9617,7 +9706,7 @@ goto block_20; if (x_55 == 0) { lean_object* x_56; -lean_dec_ref(x_54); +lean_dec_ref(x_52); x_56 = l_Lean_Elab_Tactic_SavedState_restore___redArg(x_53, x_55, x_7, x_8, x_9, x_10, x_11, x_12, x_13); if (lean_obj_tag(x_56) == 0) { @@ -9674,7 +9763,7 @@ lean_dec_ref(x_10); lean_dec(x_9); lean_dec_ref(x_8); x_61 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_61, 0, x_54); +lean_ctor_set(x_61, 0, x_52); return x_61; } } @@ -9684,7 +9773,7 @@ if (x_66 == 0) { lean_object* x_67; lean_dec_ref(x_65); -x_67 = l_Lean_Elab_Tactic_SavedState_restore___redArg(x_63, x_66, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_67 = l_Lean_Elab_Tactic_SavedState_restore___redArg(x_64, x_66, x_7, x_8, x_9, x_10, x_11, x_12, x_13); if (lean_obj_tag(x_67) == 0) { lean_object* x_68; @@ -9730,7 +9819,7 @@ return x_71; else { lean_object* x_72; -lean_dec_ref(x_63); +lean_dec_ref(x_64); lean_dec(x_47); lean_dec(x_28); lean_dec(x_27); @@ -10272,17 +10361,17 @@ if (x_159 == 0) uint8_t x_160; lean_inc(x_158); x_160 = l_Lean_Exception_isRuntime(x_158); -x_52 = lean_box(0); -x_53 = x_158; -x_54 = x_153; +x_52 = x_153; +x_53 = lean_box(0); +x_54 = x_158; x_55 = x_160; goto block_62; } else { -x_52 = lean_box(0); -x_53 = x_158; -x_54 = x_153; +x_52 = x_153; +x_53 = lean_box(0); +x_54 = x_158; x_55 = x_159; goto block_62; } @@ -10402,16 +10491,16 @@ if (x_176 == 0) uint8_t x_177; lean_inc(x_175); x_177 = l_Lean_Exception_isRuntime(x_175); -x_63 = x_170; -x_64 = x_175; +x_63 = x_175; +x_64 = x_170; x_65 = lean_box(0); x_66 = x_177; goto block_73; } else { -x_63 = x_170; -x_64 = x_175; +x_63 = x_175; +x_64 = x_170; x_65 = lean_box(0); x_66 = x_176; goto block_73; @@ -10868,7 +10957,7 @@ goto block_220; if (x_208 == 0) { lean_object* x_209; -lean_dec_ref(x_206); +lean_dec_ref(x_207); lean_dec(x_205); x_209 = l_Lean_Elab_Tactic_SavedState_restore___redArg(x_204, x_208, x_7, x_8, x_9, x_10, x_11, x_12, x_13); if (lean_obj_tag(x_209) == 0) @@ -10931,7 +11020,7 @@ if (lean_is_scalar(x_205)) { x_214 = x_205; lean_ctor_set_tag(x_214, 1); } -lean_ctor_set(x_214, 0, x_206); +lean_ctor_set(x_214, 0, x_207); return x_214; } } @@ -10944,15 +11033,15 @@ if (x_218 == 0) uint8_t x_219; lean_inc_ref(x_216); x_219 = l_Lean_Exception_isRuntime(x_216); -x_206 = x_216; -x_207 = lean_box(0); +x_206 = lean_box(0); +x_207 = x_216; x_208 = x_219; goto block_215; } else { -x_206 = x_216; -x_207 = lean_box(0); +x_206 = lean_box(0); +x_207 = x_216; x_208 = x_218; goto block_215; } @@ -11091,8 +11180,8 @@ goto block_20; if (x_55 == 0) { lean_object* x_56; -lean_dec_ref(x_53); -x_56 = l_Lean_Elab_Tactic_SavedState_restore___redArg(x_54, x_55, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec_ref(x_54); +x_56 = l_Lean_Elab_Tactic_SavedState_restore___redArg(x_52, x_55, x_7, x_8, x_9, x_10, x_11, x_12, x_13); if (lean_obj_tag(x_56) == 0) { lean_object* x_57; @@ -11136,7 +11225,7 @@ return x_60; else { lean_object* x_61; -lean_dec_ref(x_54); +lean_dec_ref(x_52); lean_dec(x_47); lean_dec(x_27); lean_dec(x_26); @@ -11148,7 +11237,7 @@ lean_dec_ref(x_10); lean_dec(x_9); lean_dec_ref(x_8); x_61 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_61, 0, x_53); +lean_ctor_set(x_61, 0, x_54); return x_61; } } @@ -11157,8 +11246,8 @@ return x_61; if (x_66 == 0) { lean_object* x_67; -lean_dec_ref(x_64); -x_67 = l_Lean_Elab_Tactic_SavedState_restore___redArg(x_63, x_66, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec_ref(x_63); +x_67 = l_Lean_Elab_Tactic_SavedState_restore___redArg(x_64, x_66, x_7, x_8, x_9, x_10, x_11, x_12, x_13); if (lean_obj_tag(x_67) == 0) { lean_object* x_68; @@ -11204,7 +11293,7 @@ return x_71; else { lean_object* x_72; -lean_dec_ref(x_63); +lean_dec_ref(x_64); lean_dec(x_47); lean_dec(x_28); lean_dec(x_27); @@ -11218,7 +11307,7 @@ lean_dec_ref(x_10); lean_dec(x_9); lean_dec_ref(x_8); x_72 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_72, 0, x_64); +lean_ctor_set(x_72, 0, x_63); return x_72; } } @@ -13186,6 +13275,21 @@ lean_mark_persistent(l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Ta if (builtin) {res = l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_540456248____hygCtx___hyg_2_(); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +}l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__0_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_ = _init_l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__0_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__0_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_); +l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__1_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_ = _init_l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__1_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__1_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_); +l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__2_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_ = _init_l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__2_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__2_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_); +l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__3_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_ = _init_l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__3_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__3_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_); +l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__4_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_ = _init_l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__4_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__4_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_); +l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__5_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_ = _init_l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__5_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn___closed__5_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_); +if (builtin) {res = l___private_Lean_Elab_Tactic_Do_VCGen_Basic_0__Lean_Elab_Tactic_Do_initFn_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_4153898153____hygCtx___hyg_2_(); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); }l_Lean_Elab_Tactic_Do_initFn___closed__0_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_1593450264____hygCtx___hyg_4_ = _init_l_Lean_Elab_Tactic_Do_initFn___closed__0_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_1593450264____hygCtx___hyg_4_(); lean_mark_persistent(l_Lean_Elab_Tactic_Do_initFn___closed__0_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_1593450264____hygCtx___hyg_4_); l_Lean_Elab_Tactic_Do_initFn___closed__1_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_1593450264____hygCtx___hyg_4_ = _init_l_Lean_Elab_Tactic_Do_initFn___closed__1_00___x40_Lean_Elab_Tactic_Do_VCGen_Basic_1593450264____hygCtx___hyg_4_(); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Do/VCGen/Split.c b/stage0/stdlib/Lean/Elab/Tactic/Do/VCGen/Split.c index 4bd570385a3e..d7d037acf051 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Do/VCGen/Split.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Do/VCGen/Split.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Elab.Tactic.Do.VCGen.Split -// Imports: public import Lean.Meta.Tactic.FunInd public import Lean.Meta.Match.MatcherApp.Transform import Lean.Meta.Tactic.Simp.Rewrite import Lean.Meta.Tactic.Assumption +// Imports: public import Lean.Meta.Tactic.Simp.Types public import Lean.Meta.Match.MatcherApp.Transform public import Lean.Data.Array import Lean.Meta.Match.Rewrite import Lean.Meta.Tactic.Simp.Rewrite import Lean.Meta.Tactic.Assumption #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -13,159 +13,173 @@ #ifdef __cplusplus extern "C" { #endif +static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__3___closed__2; static lean_object* l_Lean_Elab_Tactic_Do_instInhabitedSplitInfo_default___closed__3; lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addMessageContextFull___at___00Lean_throwError___at___00Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__8_spec__10_spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_altInfos___closed__0; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_looseBVarRange(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); lean_object* l_List_lengthTR___redArg(lean_object*); lean_object* l_Lean_Core_instMonadCoreM___lam__0___boxed(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_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7_spec__8___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_matcher_elim(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_getSplitInfo_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__3___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__17(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at___00Lean_Elab_Tactic_Do_SplitInfo_altInfos_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_instInhabitedSplitInfo; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__9(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7___closed__0; -static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__11___closed__0; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Environment_header(lean_object*); static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7_spec__8___redArg___closed__15; +static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__9; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__11; static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7_spec__8___redArg___closed__0; static lean_object* l_Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0___closed__2; -lean_object* l_Lean_Tactic_FunInd_rwMatcher(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Array_mask___redArg(lean_object*, lean_object*); lean_object* l_Lean_Meta_MatcherApp_transform___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_ite_elim(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7_spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isApp(lean_object*); -static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__4___closed__0; lean_object* l_Lean_Expr_sort___override(lean_object*); lean_object* l_Lean_Meta_mkLambdaFVars___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__3___closed__0; +static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__3___closed__0; static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7_spec__8___redArg___closed__8; extern lean_object* l_Lean_unknownIdentifierMessageTag; static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7_spec__8___redArg___closed__10; uint8_t l_Lean_Name_isAnonymous(lean_object*); lean_object* l_ReaderT_instFunctorOfMonad___redArg___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__4___closed__0; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_ctorElim___redArg(lean_object*, lean_object*); static lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4___redArg___closed__3; lean_object* l_Lean_replaceRef(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__5; LEAN_EXPORT lean_object* l_Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Match_MatcherInfo_arity(lean_object*); static lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4___redArg___closed__2; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__8___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*, uint8_t); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__0; +static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__4___closed__1; +static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__1___closed__0; lean_object* lean_array_fget_borrowed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_rwIfOrMatcher(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isAppOf(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instApplicativeOfMonad___redArg___lam__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_altInfos(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__2___redArg(lean_object*, lean_object*); lean_object* l_Lean_MessageData_note(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__17___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_panic___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__1___closed__3; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_getLevel___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Id_instMonad___lam__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_ctorElim___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__9(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instMonad___redArg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__12___closed__1; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7_spec__8___redArg___closed__12; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__17(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__0___closed__0; LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__8_spec__10___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at___00Lean_Elab_Tactic_Do_SplitInfo_altInfos_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7_spec__8___redArg___closed__6; static lean_object* l_panic___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__1___closed__2; static lean_object* l_panic___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__1___closed__0; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_ctorElim(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_findLocalDeclWithType_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7_spec__8___redArg___closed__13; static lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4___redArg___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_dite_elim(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getRevArg_x21(lean_object*, lean_object*); lean_object* l_Lean_Meta_Match_Extension_getMatcherInfo_x3f(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__1___closed__1; +lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__3___closed__3; static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7_spec__8___redArg___closed__14; +static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__1; lean_object* l_Lean_Meta_Match_MatcherInfo_getMotivePos(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instMonadEST(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__8_spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_resTy(lean_object*); lean_object* l_Lean_Name_getPrefix(lean_object*); -static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__3___closed__2; +static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__6; static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7_spec__8___redArg___closed__7; lean_object* l_Lean_Meta_Match_MatcherInfo_numAlts(lean_object*); +lean_object* l_Id_instMonad___lam__2___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__1___closed__1; static lean_object* l_Lean_Elab_Tactic_Do_instInhabitedSplitInfo_default___closed__2; static lean_object* l_Lean_Elab_Tactic_Do_instInhabitedSplitInfo_default___closed__0; -static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__3___closed__1; lean_object* l_ReaderT_instApplicativeOfMonad___redArg___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__8_spec__10___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instInhabitedOfMonad___redArg(lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__5(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive(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); -static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__0___closed__1; -static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___closed__0; +static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__7; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__3(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__1___closed__0; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_rwIfOrMatcher___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkConst(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__14(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_withLocalDecl___redArg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t); extern lean_object* l_Lean_instInhabitedExpr; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); static lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4___redArg___closed__0; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7_spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addMessageContextFull___at___00Lean_throwError___at___00Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__8_spec__10_spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Id_instMonad___lam__3(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__0___closed__1; lean_object* l_Lean_Core_instMonadCoreM___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__10; LEAN_EXPORT lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0___closed__1; lean_object* l_Lean_mkNot(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_simpDiscrs_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Basic_0__Nat_repeatTR_loop___at___00Lean_Elab_Tactic_Do_SplitInfo_resTy_spec__0(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_extract___redArg(lean_object*, lean_object*, lean_object*); lean_object* l_Subarray_toArray___redArg(lean_object*); +lean_object* l_Lean_Expr_abstractM___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__8___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__8_spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0___closed__0; +static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__3___closed__1; lean_object* lean_array_get_borrowed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Core_mkFreshUserName(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__0___closed__0; lean_object* l_Lean_MessageData_ofConstName(lean_object*, uint8_t); uint8_t l_Lean_Environment_contains(lean_object*, lean_object*, uint8_t); lean_object* l_Lean_MessageData_ofExpr(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__2___redArg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__1___closed__4; +lean_object* l_Lean_Meta_rwMatcher(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppNumArgs(lean_object*); lean_object* lean_name_append_index_after(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__12___closed__1; static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7_spec__8___redArg___closed__19; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0___closed__3; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__11(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Id_instMonad___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7_spec__8___redArg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_instMonadMetaM___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_isFVar___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_ite_elim___redArg(lean_object*, lean_object*); lean_object* l_Lean_Meta_instMonadMetaM___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -175,6 +189,8 @@ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7_spec__8___redArg___closed__1; LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_rwIfOrMatcher___closed__0; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__5(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_EnvironmentHeader_moduleNames(lean_object*); static lean_object* l_Lean_Elab_Tactic_Do_instInhabitedSplitInfo_default___closed__1; @@ -182,12 +198,13 @@ static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknow LEAN_EXPORT lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instFunctorOfMonad___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7_spec__8___redArg___closed__18; -lean_object* l_Lean_Tactic_FunInd_rwIfWith(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_MatcherApp_toExpr(lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_toSubarray___redArg(lean_object*, lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); extern lean_object* l_Lean_Meta_Match_instInhabitedAltParamInfo_default; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); lean_object* l_ReaderT_instApplicativeOfMonad___redArg___lam__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -195,19 +212,21 @@ LEAN_EXPORT lean_object* l_panic___at___00Lean_Meta_matchMatcherApp_x3f___at___0 uint8_t l_Lean_isPrivateName(lean_object*); static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7_spec__8___redArg___closed__17; LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__10(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Environment_getModuleIdxFor_x3f(lean_object*, lean_object*); +lean_object* l_Id_instMonad___lam__6(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__4; lean_object* l_Subarray_size___redArg(lean_object*); lean_object* l_Lean_Meta_Simp_simpMatchDiscrs_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_expr_instantiate_rev(lean_object*, lean_object*); lean_object* l_Lean_Meta_MatcherApp_altNumParams(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_mk(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); lean_object* l_mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_ctorIdx___boxed(lean_object*); static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7_spec__8___redArg___closed__9; LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -217,46 +236,52 @@ size_t lean_array_size(lean_object*); static lean_object* l_panic___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkFVar(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr1(lean_object*); static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7_spec__8___redArg___closed__11; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); +static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__3; static lean_object* l_Lean_Elab_Tactic_Do_rwIfOrMatcher___closed__1; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7_spec__8___redArg___closed__4; lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7_spec__8___redArg___closed__3; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__1___closed__5; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); +lean_object* l_Id_instMonad___lam__0(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__12___closed__0; +lean_object* l_Lean_Meta_rwIfWith(lean_object*, lean_object*, lean_object*, lean_object*, 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*); uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_simpDiscrs_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__10(lean_object*, lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkApp5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7_spec__8___redArg___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_dite_elim___redArg(lean_object*, lean_object*); lean_object* l_Lean_Environment_setExporting(lean_object*, uint8_t); static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7_spec__8___redArg___closed__20; -static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__4___closed__1; lean_object* l_Lean_InductiveVal_numCtors(lean_object*); LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at___00Lean_Elab_Tactic_Do_SplitInfo_altInfos_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__11___closed__0; static lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__3___closed__2; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__12___closed__0; lean_object* l_Lean_MessageData_ofName(lean_object*); uint8_t l_Lean_isCasesOnRecursor(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_ctorIdx(lean_object*); +lean_object* l_Id_instMonad___lam__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7_spec__8___redArg___closed__16; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_matcher_elim___redArg(lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at___00Lean_Elab_Tactic_Do_SplitInfo_altInfos_spec__0___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_getSplitInfo_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__2; +static lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__8; static lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__3___closed__0; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_ctorIdx(lean_object* x_1) { _start: @@ -801,7 +826,7 @@ lean_dec_ref(x_1); return x_7; } } -static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__3___closed__0() { +static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__3___closed__0() { _start: { lean_object* x_1; @@ -809,16 +834,16 @@ x_1 = lean_mk_string_unchecked("isTrue", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__3___closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__3___closed__1() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__3___closed__0; +x_1 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__3___closed__0; x_2 = l_Lean_Name_mkStr1(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__3___closed__2() { +static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; @@ -827,7 +852,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__0___closed__0() { +static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__0___closed__0() { _start: { lean_object* x_1; @@ -835,20 +860,20 @@ x_1 = lean_mk_string_unchecked("ite", 3, 3); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__0___closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__0___closed__1() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__0___closed__0; +x_1 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__0___closed__0; x_2 = l_Lean_Name_mkStr1(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_8 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__0___closed__1; +x_8 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__0___closed__1; x_9 = lean_box(0); x_10 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_10, 0, x_1); @@ -859,7 +884,7 @@ x_13 = lean_apply_2(x_6, lean_box(0), x_12); return x_13; } } -static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__1___closed__0() { +static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__1___closed__0() { _start: { lean_object* x_1; @@ -867,33 +892,34 @@ x_1 = lean_mk_string_unchecked("isFalse", 7, 7); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__1___closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__1___closed__1() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__1___closed__0; +x_1 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__1___closed__0; x_2 = l_Lean_Name_mkStr1(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_11 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__0), 7, 6); +lean_inc_ref(x_2); +x_11 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__0), 7, 6); lean_closure_set(x_11, 0, x_1); lean_closure_set(x_11, 1, x_2); lean_closure_set(x_11, 2, x_3); lean_closure_set(x_11, 3, x_4); lean_closure_set(x_11, 4, x_10); lean_closure_set(x_11, 5, x_5); -x_12 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__1___closed__1; -x_13 = lean_apply_3(x_6, x_12, x_7, x_8); +x_12 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__1___closed__1; +x_13 = lean_apply_4(x_6, x_12, x_2, x_7, x_8); x_14 = lean_apply_4(x_9, lean_box(0), lean_box(0), x_13, x_11); return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__2(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__2(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5) { _start: { 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; lean_object* x_16; @@ -918,71 +944,71 @@ x_16 = lean_apply_2(x_4, lean_box(0), x_15); return x_16; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { uint8_t x_6; lean_object* x_7; x_6 = lean_unbox(x_3); -x_7 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__2(x_1, x_2, x_6, x_4, x_5); +x_7 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__2(x_1, x_2, x_6, x_4, x_5); lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__3(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__3(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_7 = lean_box(x_2); -x_8 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__2___boxed), 5, 4); -lean_closure_set(x_8, 0, x_1); -lean_closure_set(x_8, 1, x_6); -lean_closure_set(x_8, 2, x_7); -lean_closure_set(x_8, 3, x_3); -x_9 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__3___closed__1; -x_10 = lean_unsigned_to_nat(0u); -x_11 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__3___closed__2; -x_12 = lean_apply_3(x_4, x_9, x_10, x_11); -x_13 = lean_apply_4(x_5, lean_box(0), lean_box(0), x_12, x_8); -return x_13; +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_8 = lean_box(x_2); +x_9 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__2___boxed), 5, 4); +lean_closure_set(x_9, 0, x_1); +lean_closure_set(x_9, 1, x_7); +lean_closure_set(x_9, 2, x_8); +lean_closure_set(x_9, 3, x_3); +x_10 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__3___closed__1; +x_11 = lean_unsigned_to_nat(0u); +x_12 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__3___closed__2; +x_13 = lean_apply_4(x_4, x_10, x_5, x_11, x_12); +x_14 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_13, x_9); +return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -uint8_t x_7; lean_object* x_8; -x_7 = lean_unbox(x_2); -x_8 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__3(x_1, x_7, x_3, x_4, x_5, x_6); -return x_8; +uint8_t x_8; lean_object* x_9; +x_8 = lean_unbox(x_2); +x_9 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__3(x_1, x_8, x_3, x_4, x_5, x_6, x_7); +return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__5(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__5(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_7 = lean_box(x_2); +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_8 = lean_box(x_2); lean_inc(x_1); -x_8 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__2___boxed), 5, 4); -lean_closure_set(x_8, 0, x_1); -lean_closure_set(x_8, 1, x_6); -lean_closure_set(x_8, 2, x_7); -lean_closure_set(x_8, 3, x_3); -x_9 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__1___closed__1; -x_10 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__3___closed__2; -x_11 = lean_apply_3(x_4, x_9, x_1, x_10); -x_12 = lean_apply_4(x_5, lean_box(0), lean_box(0), x_11, x_8); -return x_12; +x_9 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__2___boxed), 5, 4); +lean_closure_set(x_9, 0, x_1); +lean_closure_set(x_9, 1, x_7); +lean_closure_set(x_9, 2, x_8); +lean_closure_set(x_9, 3, x_3); +x_10 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__1___closed__1; +x_11 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__3___closed__2; +x_12 = lean_apply_4(x_4, x_10, x_5, x_1, x_11); +x_13 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_12, x_9); +return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -uint8_t x_7; lean_object* x_8; -x_7 = lean_unbox(x_2); -x_8 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__5(x_1, x_7, x_3, x_4, x_5, x_6); -return x_8; +uint8_t x_8; lean_object* x_9; +x_8 = lean_unbox(x_2); +x_9 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__5(x_1, x_8, x_3, x_4, x_5, x_6, x_7); +return x_9; } } -static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__4___closed__0() { +static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__4___closed__0() { _start: { lean_object* x_1; @@ -990,20 +1016,20 @@ x_1 = lean_mk_string_unchecked("dite", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__4___closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__4___closed__1() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__4___closed__0; +x_1 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__4___closed__0; x_2 = l_Lean_Name_mkStr1(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_8 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__4___closed__1; +x_8 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__4___closed__1; x_9 = lean_box(0); x_10 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_10, 0, x_1); @@ -1014,12 +1040,12 @@ x_13 = lean_apply_2(x_6, lean_box(0), x_12); return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, uint8_t x_9, lean_object* x_10, uint8_t x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, uint8_t x_9, lean_object* x_10, uint8_t x_11, lean_object* x_12, lean_object* x_13) { _start: { lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_inc_ref(x_3); -x_14 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__4), 7, 6); +x_14 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__4), 7, 6); lean_closure_set(x_14, 0, x_1); lean_closure_set(x_14, 1, x_2); lean_closure_set(x_14, 2, x_3); @@ -1032,17 +1058,17 @@ x_17 = lean_apply_4(x_12, lean_box(0), lean_box(0), x_16, x_14); return x_17; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { uint8_t x_14; uint8_t x_15; lean_object* x_16; x_14 = lean_unbox(x_9); x_15 = lean_unbox(x_11); -x_16 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_14, x_10, x_15, x_12, x_13); +x_16 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_14, x_10, x_15, x_12, x_13); return x_16; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { uint8_t x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; @@ -1055,7 +1081,7 @@ lean_inc(x_11); lean_inc_ref(x_7); lean_inc_ref(x_6); lean_inc_ref(x_3); -x_16 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__6___boxed), 13, 12); +x_16 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__6___boxed), 13, 12); lean_closure_set(x_16, 0, x_1); lean_closure_set(x_16, 1, x_2); lean_closure_set(x_16, 2, x_3); @@ -1073,7 +1099,7 @@ x_18 = lean_apply_4(x_9, lean_box(0), lean_box(0), x_17, x_16); return x_18; } } -static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__12___closed__0() { +static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__12___closed__0() { _start: { lean_object* x_1; @@ -1081,16 +1107,16 @@ x_1 = lean_mk_string_unchecked("h", 1, 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__12___closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__12___closed__1() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__12___closed__0; +x_1 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__12___closed__0; x_2 = l_Lean_Name_mkStr1(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_7; @@ -1098,17 +1124,17 @@ x_7 = l_Lean_Core_mkFreshUserName(x_1, x_4, x_5); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; -x_7 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__8(x_1, x_2, x_3, x_4, x_5); +x_7 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__8(x_1, x_2, x_3, x_4, x_5); lean_dec(x_3); lean_dec_ref(x_2); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__9(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__9(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; @@ -1130,12 +1156,13 @@ lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean lean_dec_ref(x_9); lean_dec_ref(x_8); lean_dec(x_7); -x_20 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__3___closed__1; +x_20 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__3___closed__1; x_21 = lean_unsigned_to_nat(0u); -x_22 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__3___closed__2; +x_22 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__3___closed__2; lean_inc(x_6); lean_inc(x_5); -x_23 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__1), 10, 9); +lean_inc_ref(x_3); +x_23 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__1), 10, 9); lean_closure_set(x_23, 0, x_10); lean_closure_set(x_23, 1, x_3); lean_closure_set(x_23, 2, x_15); @@ -1145,7 +1172,7 @@ lean_closure_set(x_23, 5, x_5); lean_closure_set(x_23, 6, x_11); lean_closure_set(x_23, 7, x_22); lean_closure_set(x_23, 8, x_6); -x_24 = lean_apply_3(x_5, x_20, x_21, x_22); +x_24 = lean_apply_4(x_5, x_20, x_3, x_21, x_22); x_25 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_24, x_23); return x_25; } @@ -1154,25 +1181,29 @@ else lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; x_26 = lean_box(x_2); lean_inc(x_6); +lean_inc_ref(x_3); lean_inc(x_5); lean_inc(x_7); -x_27 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__3___boxed), 6, 5); +x_27 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__3___boxed), 7, 6); lean_closure_set(x_27, 0, x_11); lean_closure_set(x_27, 1, x_26); lean_closure_set(x_27, 2, x_7); lean_closure_set(x_27, 3, x_5); -lean_closure_set(x_27, 4, x_6); +lean_closure_set(x_27, 4, x_3); +lean_closure_set(x_27, 5, x_6); x_28 = lean_box(x_2); lean_inc(x_6); +lean_inc_ref(x_3); lean_inc(x_7); -x_29 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__5___boxed), 6, 5); +x_29 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__5___boxed), 7, 6); lean_closure_set(x_29, 0, x_11); lean_closure_set(x_29, 1, x_28); lean_closure_set(x_29, 2, x_7); lean_closure_set(x_29, 3, x_5); -lean_closure_set(x_29, 4, x_6); +lean_closure_set(x_29, 4, x_3); +lean_closure_set(x_29, 5, x_6); lean_inc(x_6); -x_30 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__7), 11, 10); +x_30 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__7), 11, 10); lean_closure_set(x_30, 0, x_10); lean_closure_set(x_30, 1, x_3); lean_closure_set(x_30, 2, x_15); @@ -1183,8 +1214,8 @@ lean_closure_set(x_30, 6, x_9); lean_closure_set(x_30, 7, x_29); lean_closure_set(x_30, 8, x_6); lean_closure_set(x_30, 9, x_27); -x_31 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__12___closed__1; -x_32 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__8___boxed), 6, 1); +x_31 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__12___closed__1; +x_32 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__8___boxed), 6, 1); lean_closure_set(x_32, 0, x_31); x_33 = lean_apply_2(x_7, lean_box(0), x_32); x_34 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_33, x_30); @@ -1192,17 +1223,17 @@ return x_34; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { uint8_t x_11; lean_object* x_12; x_11 = lean_unbox(x_2); -x_12 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__9(x_1, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_12 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__9(x_1, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec_ref(x_1); return x_12; } } -static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__11___closed__0() { +static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__11___closed__0() { _start: { lean_object* x_1; lean_object* x_2; @@ -1211,7 +1242,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__10(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__10(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { uint8_t x_4; uint8_t x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; @@ -1235,140 +1266,279 @@ x_13 = lean_apply_2(x_2, lean_box(0), x_12); return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_5 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__3___closed__1; -x_6 = lean_unsigned_to_nat(0u); -x_7 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__11___closed__0; -x_8 = lean_array_push(x_7, x_4); -lean_inc_ref(x_8); -x_9 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__10), 3, 2); -lean_closure_set(x_9, 0, x_8); -lean_closure_set(x_9, 1, x_1); -x_10 = lean_apply_3(x_2, x_5, x_6, x_8); -x_11 = lean_apply_4(x_3, lean_box(0), lean_box(0), x_10, x_9); -return x_11; +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_6 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__3___closed__1; +x_7 = lean_unsigned_to_nat(0u); +x_8 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__11___closed__0; +x_9 = lean_array_push(x_8, x_5); +lean_inc_ref(x_9); +x_10 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__10), 3, 2); +lean_closure_set(x_10, 0, x_9); +lean_closure_set(x_10, 1, x_1); +x_11 = lean_apply_4(x_2, x_6, x_3, x_7, x_9); +x_12 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_11, x_10); +return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__13(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__13(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_6 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__1___closed__1; -x_7 = lean_mk_empty_array_with_capacity(x_1); -x_8 = lean_array_push(x_7, x_5); -lean_inc_ref(x_8); -x_9 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__10), 3, 2); -lean_closure_set(x_9, 0, x_8); -lean_closure_set(x_9, 1, x_2); -x_10 = lean_apply_3(x_3, x_6, x_1, x_8); -x_11 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_10, x_9); -return x_11; +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_7 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__1___closed__1; +x_8 = lean_mk_empty_array_with_capacity(x_1); +x_9 = lean_array_push(x_8, x_6); +lean_inc_ref(x_9); +x_10 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__10), 3, 2); +lean_closure_set(x_10, 0, x_9); +lean_closure_set(x_10, 1, x_2); +x_11 = lean_apply_4(x_3, x_7, x_4, x_1, x_9); +x_12 = lean_apply_4(x_5, lean_box(0), lean_box(0), x_11, x_10); +return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__17(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__17(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; x_11 = lean_unsigned_to_nat(1u); -lean_inc(x_3); +lean_inc(x_4); +lean_inc_ref(x_3); lean_inc(x_1); -x_12 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__13), 5, 4); +x_12 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__13), 6, 5); lean_closure_set(x_12, 0, x_11); lean_closure_set(x_12, 1, x_1); lean_closure_set(x_12, 2, x_2); lean_closure_set(x_12, 3, x_3); -x_13 = l_Lean_Expr_getAppNumArgs(x_4); +lean_closure_set(x_12, 4, x_4); +x_13 = l_Lean_Expr_getAppNumArgs(x_5); x_14 = lean_nat_sub(x_13, x_11); x_15 = lean_nat_sub(x_14, x_11); lean_dec(x_14); -x_16 = l_Lean_Expr_getRevArg_x21(x_4, x_15); +x_16 = l_Lean_Expr_getRevArg_x21(x_5, x_15); x_17 = lean_unsigned_to_nat(2u); x_18 = lean_nat_sub(x_13, x_17); lean_dec(x_13); x_19 = lean_nat_sub(x_18, x_11); lean_dec(x_18); -x_20 = l_Lean_Expr_getRevArg_x21(x_4, x_19); -lean_inc(x_3); -x_21 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__7), 11, 10); +x_20 = l_Lean_Expr_getRevArg_x21(x_5, x_19); +lean_inc(x_4); +x_21 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__7), 11, 10); lean_closure_set(x_21, 0, x_10); -lean_closure_set(x_21, 1, x_5); +lean_closure_set(x_21, 1, x_3); lean_closure_set(x_21, 2, x_16); lean_closure_set(x_21, 3, x_20); lean_closure_set(x_21, 4, x_6); lean_closure_set(x_21, 5, x_7); lean_closure_set(x_21, 6, x_8); lean_closure_set(x_21, 7, x_12); -lean_closure_set(x_21, 8, x_3); +lean_closure_set(x_21, 8, x_4); lean_closure_set(x_21, 9, x_9); -x_22 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__12___closed__1; -x_23 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__8___boxed), 6, 1); +x_22 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__12___closed__1; +x_23 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__8___boxed), 6, 1); lean_closure_set(x_23, 0, x_22); x_24 = lean_apply_2(x_1, lean_box(0), x_23); -x_25 = lean_apply_4(x_3, lean_box(0), lean_box(0), x_24, x_21); +x_25 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_24, x_21); return x_25; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__17___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__17___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; -x_11 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__17(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec_ref(x_4); +x_11 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__17(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec_ref(x_5); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_toExpr), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_6 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__12___closed__1; +x_6 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__12___closed__1; x_7 = lean_unsigned_to_nat(1u); x_8 = lean_nat_add(x_2, x_7); x_9 = lean_name_append_index_after(x_6, x_8); -x_10 = lean_apply_3(x_1, x_9, x_2, x_4); +x_10 = lean_apply_4(x_1, x_9, x_3, x_2, x_4); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__12___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_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; -x_6 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__12(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__12(x_1, x_2, x_3, x_4, x_5); lean_dec_ref(x_5); -lean_dec_ref(x_3); return x_6; } } -static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___closed__0() { +static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_toExpr), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Expr_isFVar___boxed), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__14(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__8() { _start: { -lean_object* x_5; -x_5 = lean_apply_2(x_1, lean_box(0), x_2); -return x_5; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Id_instMonad___lam__6), 4, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__7() { _start: { -lean_object* x_5; -x_5 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__14(x_1, x_2, x_3, x_4); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -return x_5; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Id_instMonad___lam__5___boxed), 4, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Id_instMonad___lam__4___boxed), 4, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Id_instMonad___lam__3), 4, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Id_instMonad___lam__2___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Id_instMonad___lam__1___boxed), 4, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, uint8_t x_12) { +static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Id_instMonad___lam__0), 4, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__3; +x_2 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__2; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___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_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__7; +x_2 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__6; +x_3 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__5; +x_4 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__4; +x_5 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__9; +x_6 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_4); +lean_ctor_set(x_6, 2, x_3); +lean_ctor_set(x_6, 3, x_2); +lean_ctor_set(x_6, 4, x_1); +return x_6; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__8; +x_2 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__10; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__14(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_6 = l_Array_mask___redArg(x_1, x_4); +x_7 = lean_expr_instantiate_rev(x_2, x_6); +lean_dec_ref(x_6); +x_8 = lean_apply_2(x_3, lean_box(0), x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__14(x_1, x_2, x_3, x_4, x_5); +lean_dec_ref(x_5); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__15(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, uint8_t x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_17 = lean_ctor_get(x_1, 0); +lean_inc(x_17); +lean_dec_ref(x_1); +lean_inc(x_3); +x_18 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__14___boxed), 5, 3); +lean_closure_set(x_18, 0, x_2); +lean_closure_set(x_18, 1, x_16); +lean_closure_set(x_18, 2, x_3); +x_19 = lean_apply_1(x_3, lean_box(0)); +lean_inc(x_19); +x_20 = l_Lean_Meta_MatcherApp_transform___redArg(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_13, x_19, x_18, x_14, x_19); +x_21 = lean_apply_4(x_17, lean_box(0), lean_box(0), x_15, x_20); +return x_21; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__15___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +uint8_t x_17; lean_object* x_18; +x_17 = lean_unbox(x_13); +x_18 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__15(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_17, x_14, x_15, x_16); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, uint8_t x_12) { _start: { switch (lean_obj_tag(x_9)) { @@ -1392,7 +1562,7 @@ x_17 = lean_box(x_12); lean_inc(x_1); lean_inc(x_14); lean_inc_ref(x_10); -x_18 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__9___boxed), 10, 9); +x_18 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__9___boxed), 10, 9); lean_closure_set(x_18, 0, x_16); lean_closure_set(x_18, 1, x_17); lean_closure_set(x_18, 2, x_10); @@ -1425,21 +1595,23 @@ x_25 = lean_ctor_get(x_9, 0); lean_inc_ref(x_25); lean_dec_ref(x_9); lean_inc(x_23); +lean_inc_ref(x_10); lean_inc(x_11); lean_inc(x_1); -x_26 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__11), 4, 3); +x_26 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__11), 5, 4); lean_closure_set(x_26, 0, x_1); lean_closure_set(x_26, 1, x_11); -lean_closure_set(x_26, 2, x_23); -lean_inc_ref(x_10); +lean_closure_set(x_26, 2, x_10); +lean_closure_set(x_26, 3, x_23); lean_inc(x_23); +lean_inc_ref(x_10); lean_inc(x_1); -x_27 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__17___boxed), 10, 9); +x_27 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__17___boxed), 10, 9); lean_closure_set(x_27, 0, x_1); lean_closure_set(x_27, 1, x_11); -lean_closure_set(x_27, 2, x_23); -lean_closure_set(x_27, 3, x_25); -lean_closure_set(x_27, 4, x_10); +lean_closure_set(x_27, 2, x_10); +lean_closure_set(x_27, 3, x_23); +lean_closure_set(x_27, 4, x_25); lean_closure_set(x_27, 5, x_24); lean_closure_set(x_27, 6, x_2); lean_closure_set(x_27, 7, x_3); @@ -1452,55 +1624,82 @@ return x_30; } default: { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; size_t x_41; size_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; x_31 = lean_ctor_get(x_3, 0); -x_32 = lean_ctor_get(x_31, 0); -x_33 = lean_ctor_get(x_31, 1); -x_34 = lean_ctor_get(x_9, 0); -lean_inc_ref(x_34); +x_32 = lean_ctor_get(x_9, 0); +lean_inc_ref(x_32); lean_dec_ref(x_9); -x_35 = lean_ctor_get(x_32, 0); -lean_inc(x_35); -x_36 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__12___boxed), 5, 1); -lean_closure_set(x_36, 0, x_11); -x_37 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___closed__0; -lean_inc(x_33); -x_38 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__14___boxed), 4, 2); -lean_closure_set(x_38, 0, x_33); -lean_closure_set(x_38, 1, x_10); +x_33 = lean_ctor_get(x_3, 1); lean_inc(x_33); -x_39 = lean_apply_1(x_33, lean_box(0)); -lean_inc(x_39); -x_40 = l_Lean_Meta_MatcherApp_transform___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_34, x_12, x_12, x_39, x_38, x_36, x_39); -x_41 = lean_apply_4(x_35, lean_box(0), lean_box(0), x_37, x_40); -return x_41; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, uint8_t x_13) { +x_34 = lean_ctor_get(x_31, 0); +lean_inc_ref(x_34); +x_35 = lean_ctor_get(x_31, 1); +lean_inc(x_35); +x_36 = lean_ctor_get(x_32, 5); +lean_inc_ref(x_36); +x_37 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__0; +x_38 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__12___boxed), 5, 1); +lean_closure_set(x_38, 0, x_11); +x_39 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__1; +x_40 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__11; +x_41 = lean_array_size(x_36); +x_42 = 0; +lean_inc_ref(x_36); +x_43 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map(lean_box(0), lean_box(0), lean_box(0), x_40, x_39, x_41, x_42, x_36); +x_44 = lean_box(x_12); +lean_inc(x_1); +lean_inc(x_43); +x_45 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__15___boxed), 16, 15); +lean_closure_set(x_45, 0, x_34); +lean_closure_set(x_45, 1, x_43); +lean_closure_set(x_45, 2, x_35); +lean_closure_set(x_45, 3, x_1); +lean_closure_set(x_45, 4, x_2); +lean_closure_set(x_45, 5, x_3); +lean_closure_set(x_45, 6, x_4); +lean_closure_set(x_45, 7, x_5); +lean_closure_set(x_45, 8, x_6); +lean_closure_set(x_45, 9, x_7); +lean_closure_set(x_45, 10, x_8); +lean_closure_set(x_45, 11, x_32); +lean_closure_set(x_45, 12, x_44); +lean_closure_set(x_45, 13, x_38); +lean_closure_set(x_45, 14, x_37); +x_46 = l_Array_mask___redArg(x_43, x_36); +lean_dec(x_43); +x_47 = lean_alloc_closure((void*)(l_Lean_Expr_abstractM___boxed), 7, 2); +lean_closure_set(x_47, 0, x_10); +lean_closure_set(x_47, 1, x_46); +x_48 = lean_apply_2(x_1, lean_box(0), x_47); +x_49 = lean_apply_4(x_33, lean_box(0), lean_box(0), x_48, x_45); +return x_49; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, uint8_t x_13) { _start: { lean_object* x_14; -x_14 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_14 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { uint8_t x_14; lean_object* x_15; x_14 = lean_unbox(x_13); -x_15 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_14); +x_15 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_14); return x_15; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { uint8_t x_13; lean_object* x_14; x_13 = lean_unbox(x_12); -x_14 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_13); +x_14 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_13); return x_14; } } @@ -4252,12 +4451,12 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_getSplitInfo_x3f(lean_object* x_1 _start: { lean_object* x_7; uint8_t x_8; -x_7 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__0___closed__1; +x_7 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__0___closed__1; x_8 = l_Lean_Expr_isAppOf(x_1, x_7); if (x_8 == 0) { lean_object* x_9; uint8_t x_10; -x_9 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__4___closed__1; +x_9 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__4___closed__1; x_10 = l_Lean_Expr_isAppOf(x_1, x_9); if (x_10 == 0) { @@ -4655,12 +4854,12 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Do_rwIfOrMatcher(lean_object* x_1, l _start: { lean_object* x_8; uint8_t x_22; lean_object* x_33; uint8_t x_34; -x_33 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__0___closed__1; +x_33 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__0___closed__1; x_34 = l_Lean_Expr_isAppOf(x_2, x_33); if (x_34 == 0) { lean_object* x_35; uint8_t x_36; -x_35 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__4___closed__1; +x_35 = l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__4___closed__1; x_36 = l_Lean_Expr_isAppOf(x_2, x_35); x_22 = x_36; goto block_32; @@ -4693,7 +4892,7 @@ x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); lean_dec_ref(x_10); x_12 = l_Lean_mkFVar(x_11); -x_13 = l_Lean_Tactic_FunInd_rwIfWith(x_12, x_2, x_3, x_4, x_5, x_6); +x_13 = l_Lean_Meta_rwIfWith(x_12, x_2, x_3, x_4, x_5, x_6); return x_13; } else @@ -4745,7 +4944,7 @@ return x_20; if (x_22 == 0) { lean_object* x_23; -x_23 = l_Lean_Tactic_FunInd_rwMatcher(x_1, x_2, x_3, x_4, x_5, x_6); +x_23 = l_Lean_Meta_rwMatcher(x_1, x_2, x_3, x_4, x_5, x_6); return x_23; } else @@ -4785,8 +4984,10 @@ x_8 = l_Lean_Elab_Tactic_Do_rwIfOrMatcher(x_1, x_2, x_3, x_4, x_5, x_6); return x_8; } } -lean_object* initialize_Lean_Meta_Tactic_FunInd(uint8_t builtin); +lean_object* initialize_Lean_Meta_Tactic_Simp_Types(uint8_t builtin); lean_object* initialize_Lean_Meta_Match_MatcherApp_Transform(uint8_t builtin); +lean_object* initialize_Lean_Data_Array(uint8_t builtin); +lean_object* initialize_Lean_Meta_Match_Rewrite(uint8_t builtin); lean_object* initialize_Lean_Meta_Tactic_Simp_Rewrite(uint8_t builtin); lean_object* initialize_Lean_Meta_Tactic_Assumption(uint8_t builtin); static bool _G_initialized = false; @@ -4794,12 +4995,18 @@ LEAN_EXPORT lean_object* initialize_Lean_Elab_Tactic_Do_VCGen_Split(uint8_t buil lean_object * res; if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); _G_initialized = true; -res = initialize_Lean_Meta_Tactic_FunInd(builtin); +res = initialize_Lean_Meta_Tactic_Simp_Types(builtin); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); res = initialize_Lean_Meta_Match_MatcherApp_Transform(builtin); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Data_Array(builtin); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Match_Rewrite(builtin); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Simp_Rewrite(builtin); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); @@ -4820,32 +5027,54 @@ l_Lean_Elab_Tactic_Do_instInhabitedSplitInfo = _init_l_Lean_Elab_Tactic_Do_instI lean_mark_persistent(l_Lean_Elab_Tactic_Do_instInhabitedSplitInfo); l_Lean_Elab_Tactic_Do_SplitInfo_altInfos___closed__0 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_altInfos___closed__0(); lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_altInfos___closed__0); -l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__3___closed__0 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__3___closed__0(); -lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__3___closed__0); -l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__3___closed__1 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__3___closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__3___closed__1); -l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__3___closed__2 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__3___closed__2(); -lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__3___closed__2); -l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__0___closed__0 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__0___closed__0(); -lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__0___closed__0); -l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__0___closed__1 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__0___closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__0___closed__1); -l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__1___closed__0 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__1___closed__0(); -lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__1___closed__0); -l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__1___closed__1 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__1___closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__1___closed__1); -l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__4___closed__0 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__4___closed__0(); -lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__4___closed__0); -l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__4___closed__1 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__4___closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__4___closed__1); -l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__12___closed__0 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__12___closed__0(); -lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__12___closed__0); -l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__12___closed__1 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__12___closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__12___closed__1); -l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__11___closed__0 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__11___closed__0(); -lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___lam__11___closed__0); -l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___closed__0 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___closed__0(); -lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWithConstantMotive___redArg___closed__0); +l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__3___closed__0 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__3___closed__0(); +lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__3___closed__0); +l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__3___closed__1 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__3___closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__3___closed__1); +l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__3___closed__2 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__3___closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__3___closed__2); +l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__0___closed__0 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__0___closed__0(); +lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__0___closed__0); +l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__0___closed__1 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__0___closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__0___closed__1); +l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__1___closed__0 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__1___closed__0(); +lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__1___closed__0); +l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__1___closed__1 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__1___closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__1___closed__1); +l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__4___closed__0 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__4___closed__0(); +lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__4___closed__0); +l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__4___closed__1 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__4___closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__4___closed__1); +l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__12___closed__0 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__12___closed__0(); +lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__12___closed__0); +l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__12___closed__1 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__12___closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__12___closed__1); +l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__11___closed__0 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__11___closed__0(); +lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___lam__11___closed__0); +l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__0 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__0(); +lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__0); +l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__1 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__1); +l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__8 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__8(); +lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__8); +l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__7 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__7(); +lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__7); +l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__6 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__6(); +lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__6); +l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__5 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__5(); +lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__5); +l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__4 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__4(); +lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__4); +l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__3 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__3); +l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__2 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__2); +l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__9 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__9(); +lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__9); +l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__10 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__10(); +lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__10); +l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__11 = _init_l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__11(); +lean_mark_persistent(l_Lean_Elab_Tactic_Do_SplitInfo_splitWith___redArg___closed__11); l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7_spec__8___redArg___closed__0 = _init_l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7_spec__8___redArg___closed__0(); lean_mark_persistent(l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7_spec__8___redArg___closed__0); l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7_spec__8___redArg___closed__1 = _init_l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00Lean_Elab_Tactic_Do_getSplitInfo_x3f_spec__0_spec__0_spec__1_spec__4_spec__6_spec__7_spec__8___redArg___closed__1(); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Grind/Param.c b/stage0/stdlib/Lean/Elab/Tactic/Grind/Param.c index 11d8017ffb4b..3256bf536923 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Grind/Param.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Grind/Param.c @@ -9780,7 +9780,7 @@ return x_9; LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logWarning___at___00Lean_checkPrivateInPublic___at___00Lean_resolveGlobalName___at___00__private_Lean_ResolveName_0__Lean_resolveLocalName_loop___at___00Lean_resolveLocalName___at___00__private_Lean_Elab_Tactic_Grind_Param_0__Lean_Elab_Tactic_processParam_spec__5_spec__8_spec__13_spec__16_spec__18_spec__20_spec__21___redArg(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; uint8_t x_55; uint8_t x_56; lean_object* x_57; lean_object* x_77; lean_object* x_78; lean_object* x_79; uint8_t x_80; lean_object* x_81; uint8_t x_82; uint8_t x_83; lean_object* x_84; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; uint8_t x_92; uint8_t x_93; uint8_t x_94; uint8_t x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; uint8_t x_104; lean_object* x_105; uint8_t x_106; uint8_t x_107; uint8_t x_109; uint8_t x_125; +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; uint8_t x_54; lean_object* x_55; uint8_t x_56; lean_object* x_57; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; uint8_t x_81; uint8_t x_82; uint8_t x_83; lean_object* x_84; lean_object* x_88; lean_object* x_89; lean_object* x_90; uint8_t x_91; uint8_t x_92; lean_object* x_93; uint8_t x_94; uint8_t x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; uint8_t x_104; lean_object* x_105; uint8_t x_106; uint8_t x_107; uint8_t x_109; uint8_t x_125; x_100 = 2; x_125 = l_Lean_instBEqMessageSeverity_beq(x_3, x_100); if (x_125 == 0) @@ -9815,15 +9815,15 @@ lean_ctor_set(x_25, 0, x_21); lean_ctor_set(x_25, 1, x_22); x_26 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_10); +lean_ctor_set(x_26, 1, x_11); x_27 = lean_alloc_ctor(0, 5, 3); -lean_ctor_set(x_27, 0, x_11); +lean_ctor_set(x_27, 0, x_12); lean_ctor_set(x_27, 1, x_16); -lean_ctor_set(x_27, 2, x_12); +lean_ctor_set(x_27, 2, x_10); lean_ctor_set(x_27, 3, x_14); lean_ctor_set(x_27, 4, x_26); -lean_ctor_set_uint8(x_27, sizeof(void*)*5, x_15); -lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 1, x_13); +lean_ctor_set_uint8(x_27, sizeof(void*)*5, x_13); +lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 1, x_15); lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 2, x_4); x_28 = l_Lean_MessageLog_add(x_27, x_24); lean_ctor_set(x_20, 6, x_28); @@ -9860,15 +9860,15 @@ lean_ctor_set(x_41, 0, x_21); lean_ctor_set(x_41, 1, x_22); x_42 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_10); +lean_ctor_set(x_42, 1, x_11); x_43 = lean_alloc_ctor(0, 5, 3); -lean_ctor_set(x_43, 0, x_11); +lean_ctor_set(x_43, 0, x_12); lean_ctor_set(x_43, 1, x_16); -lean_ctor_set(x_43, 2, x_12); +lean_ctor_set(x_43, 2, x_10); lean_ctor_set(x_43, 3, x_14); lean_ctor_set(x_43, 4, x_42); -lean_ctor_set_uint8(x_43, sizeof(void*)*5, x_15); -lean_ctor_set_uint8(x_43, sizeof(void*)*5 + 1, x_13); +lean_ctor_set_uint8(x_43, sizeof(void*)*5, x_13); +lean_ctor_set_uint8(x_43, sizeof(void*)*5 + 1, x_15); lean_ctor_set_uint8(x_43, sizeof(void*)*5 + 2, x_4); x_44 = l_Lean_MessageLog_add(x_43, x_38); x_45 = lean_alloc_ctor(0, 9, 0); @@ -9898,22 +9898,22 @@ if (x_60 == 0) { lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; x_61 = lean_ctor_get(x_59, 0); -lean_inc_ref(x_53); -x_62 = l_Lean_FileMap_toPosition(x_53, x_51); -lean_dec(x_51); -x_63 = l_Lean_FileMap_toPosition(x_53, x_57); +lean_inc_ref(x_52); +x_62 = l_Lean_FileMap_toPosition(x_52, x_55); +lean_dec(x_55); +x_63 = l_Lean_FileMap_toPosition(x_52, x_57); lean_dec(x_57); x_64 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_64, 0, x_63); x_65 = l_Lean_logAt___at___00Lean_log___at___00Lean_logWarning___at___00__private_Lean_Elab_Tactic_Grind_Param_0__Lean_Elab_Tactic_warnRedundantEMatchArg_spec__0_spec__0_spec__1___closed__0; -if (x_54 == 0) +if (x_53 == 0) { lean_free_object(x_59); lean_dec_ref(x_50); -x_10 = x_61; -x_11 = x_52; -x_12 = x_64; -x_13 = x_55; +x_10 = x_64; +x_11 = x_61; +x_12 = x_51; +x_13 = x_54; x_14 = x_65; x_15 = x_56; x_16 = x_62; @@ -9933,7 +9933,7 @@ lean_object* x_67; lean_dec_ref(x_64); lean_dec_ref(x_62); lean_dec(x_61); -lean_dec_ref(x_52); +lean_dec_ref(x_51); lean_dec_ref(x_7); x_67 = lean_box(0); lean_ctor_set(x_59, 0, x_67); @@ -9942,10 +9942,10 @@ return x_59; else { lean_free_object(x_59); -x_10 = x_61; -x_11 = x_52; -x_12 = x_64; -x_13 = x_55; +x_10 = x_64; +x_11 = x_61; +x_12 = x_51; +x_13 = x_54; x_14 = x_65; x_15 = x_56; x_16 = x_62; @@ -9962,21 +9962,21 @@ lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean x_68 = lean_ctor_get(x_59, 0); lean_inc(x_68); lean_dec(x_59); -lean_inc_ref(x_53); -x_69 = l_Lean_FileMap_toPosition(x_53, x_51); -lean_dec(x_51); -x_70 = l_Lean_FileMap_toPosition(x_53, x_57); +lean_inc_ref(x_52); +x_69 = l_Lean_FileMap_toPosition(x_52, x_55); +lean_dec(x_55); +x_70 = l_Lean_FileMap_toPosition(x_52, x_57); lean_dec(x_57); x_71 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_71, 0, x_70); x_72 = l_Lean_logAt___at___00Lean_log___at___00Lean_logWarning___at___00__private_Lean_Elab_Tactic_Grind_Param_0__Lean_Elab_Tactic_warnRedundantEMatchArg_spec__0_spec__0_spec__1___closed__0; -if (x_54 == 0) +if (x_53 == 0) { lean_dec_ref(x_50); -x_10 = x_68; -x_11 = x_52; -x_12 = x_71; -x_13 = x_55; +x_10 = x_71; +x_11 = x_68; +x_12 = x_51; +x_13 = x_54; x_14 = x_72; x_15 = x_56; x_16 = x_69; @@ -9996,7 +9996,7 @@ lean_object* x_74; lean_object* x_75; lean_dec_ref(x_71); lean_dec_ref(x_69); lean_dec(x_68); -lean_dec_ref(x_52); +lean_dec_ref(x_51); lean_dec_ref(x_7); x_74 = lean_box(0); x_75 = lean_alloc_ctor(0, 1, 0); @@ -10005,10 +10005,10 @@ return x_75; } else { -x_10 = x_68; -x_11 = x_52; -x_12 = x_71; -x_13 = x_55; +x_10 = x_71; +x_11 = x_68; +x_12 = x_51; +x_13 = x_54; x_14 = x_72; x_15 = x_56; x_16 = x_69; @@ -10023,17 +10023,17 @@ goto block_49; block_87: { lean_object* x_85; -x_85 = l_Lean_Syntax_getTailPos_x3f(x_81, x_83); -lean_dec(x_81); +x_85 = l_Lean_Syntax_getTailPos_x3f(x_79, x_82); +lean_dec(x_79); if (lean_obj_tag(x_85) == 0) { lean_inc(x_84); x_50 = x_77; -x_51 = x_84; -x_52 = x_79; -x_53 = x_78; -x_54 = x_80; -x_55 = x_82; +x_51 = x_78; +x_52 = x_80; +x_53 = x_81; +x_54 = x_82; +x_55 = x_84; x_56 = x_83; x_57 = x_84; goto block_76; @@ -10045,11 +10045,11 @@ x_86 = lean_ctor_get(x_85, 0); lean_inc(x_86); lean_dec_ref(x_85); x_50 = x_77; -x_51 = x_84; -x_52 = x_79; -x_53 = x_78; -x_54 = x_80; -x_55 = x_82; +x_51 = x_78; +x_52 = x_80; +x_53 = x_81; +x_54 = x_82; +x_55 = x_84; x_56 = x_83; x_57 = x_86; goto block_76; @@ -10058,20 +10058,20 @@ goto block_76; block_99: { lean_object* x_95; lean_object* x_96; -x_95 = l_Lean_replaceRef(x_1, x_89); -lean_dec(x_89); -x_96 = l_Lean_Syntax_getPos_x3f(x_95, x_93); +x_95 = l_Lean_replaceRef(x_1, x_93); +lean_dec(x_93); +x_96 = l_Lean_Syntax_getPos_x3f(x_95, x_92); if (lean_obj_tag(x_96) == 0) { lean_object* x_97; x_97 = lean_unsigned_to_nat(0u); x_77 = x_88; -x_78 = x_91; -x_79 = x_90; -x_80 = x_92; -x_81 = x_95; -x_82 = x_94; -x_83 = x_93; +x_78 = x_89; +x_79 = x_95; +x_80 = x_90; +x_81 = x_91; +x_82 = x_92; +x_83 = x_94; x_84 = x_97; goto block_87; } @@ -10082,12 +10082,12 @@ x_98 = lean_ctor_get(x_96, 0); lean_inc(x_98); lean_dec_ref(x_96); x_77 = x_88; -x_78 = x_91; -x_79 = x_90; -x_80 = x_92; -x_81 = x_95; -x_82 = x_94; -x_83 = x_93; +x_78 = x_89; +x_79 = x_95; +x_80 = x_90; +x_81 = x_91; +x_82 = x_92; +x_83 = x_94; x_84 = x_98; goto block_87; } @@ -10096,23 +10096,23 @@ goto block_87; { if (x_107 == 0) { -x_88 = x_105; -x_89 = x_101; +x_88 = x_101; +x_89 = x_102; x_90 = x_103; -x_91 = x_102; -x_92 = x_104; -x_93 = x_106; +x_91 = x_104; +x_92 = x_106; +x_93 = x_105; x_94 = x_3; goto block_99; } else { -x_88 = x_105; -x_89 = x_101; +x_88 = x_101; +x_89 = x_102; x_90 = x_103; -x_91 = x_102; -x_92 = x_104; -x_93 = x_106; +x_91 = x_104; +x_92 = x_106; +x_93 = x_105; x_94 = x_100; goto block_99; } @@ -10136,14 +10136,14 @@ x_118 = 1; x_119 = l_Lean_instBEqMessageSeverity_beq(x_3, x_118); if (x_119 == 0) { -lean_inc_ref(x_110); -lean_inc_ref(x_111); lean_inc(x_113); -x_101 = x_113; -x_102 = x_111; -x_103 = x_110; +lean_inc_ref(x_111); +lean_inc_ref(x_110); +x_101 = x_117; +x_102 = x_110; +x_103 = x_111; x_104 = x_114; -x_105 = x_117; +x_105 = x_113; x_106 = x_109; x_107 = x_119; goto block_108; @@ -10153,14 +10153,14 @@ else lean_object* x_120; uint8_t x_121; x_120 = l_Lean_logAt___at___00Lean_log___at___00Lean_logWarning___at___00__private_Lean_Elab_Tactic_Grind_Param_0__Lean_Elab_Tactic_warnRedundantEMatchArg_spec__0_spec__0_spec__1___closed__1; x_121 = l_Lean_Option_get___at___00Lean_logAt___at___00Lean_log___at___00Lean_logWarning___at___00__private_Lean_Elab_Tactic_Grind_Param_0__Lean_Elab_Tactic_warnRedundantEMatchArg_spec__0_spec__0_spec__1_spec__5(x_112, x_120); -lean_inc_ref(x_110); -lean_inc_ref(x_111); lean_inc(x_113); -x_101 = x_113; -x_102 = x_111; -x_103 = x_110; +lean_inc_ref(x_111); +lean_inc_ref(x_110); +x_101 = x_117; +x_102 = x_110; +x_103 = x_111; x_104 = x_114; -x_105 = x_117; +x_105 = x_113; x_106 = x_109; x_107 = x_121; goto block_108; @@ -17179,12 +17179,12 @@ lean_dec(x_9); lean_dec_ref(x_8); lean_dec_ref(x_7); x_24 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 0, x_21); return x_24; } else { -lean_dec_ref(x_22); +lean_dec_ref(x_21); x_15 = x_7; x_16 = lean_box(0); goto block_20; @@ -17201,7 +17201,7 @@ lean_dec(x_9); lean_dec_ref(x_8); lean_dec_ref(x_7); x_25 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_25, 0, x_22); +lean_ctor_set(x_25, 0, x_21); return x_25; } } @@ -17214,15 +17214,15 @@ if (x_29 == 0) uint8_t x_30; lean_inc_ref(x_27); x_30 = l_Lean_Exception_isRuntime(x_27); -x_21 = lean_box(0); -x_22 = x_27; +x_21 = x_27; +x_22 = lean_box(0); x_23 = x_30; goto block_26; } else { -x_21 = lean_box(0); -x_22 = x_27; +x_21 = x_27; +x_22 = lean_box(0); x_23 = x_29; goto block_26; } diff --git a/stage0/stdlib/Lean/Elab/Tactic/Simp.c b/stage0/stdlib/Lean/Elab/Tactic/Simp.c index 83f0febbe464..9ea69c1ded2e 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Simp.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Simp.c @@ -22251,7 +22251,7 @@ return x_2; LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logWarning___at___00Lean_checkPrivateInPublic___at___00Lean_resolveGlobalName___at___00Lean_unresolveNameGlobal___at___00Lean_unresolveNameGlobalAvoidingLocals___at___00Lean_Elab_Tactic_mkSimpOnly_spec__0_spec__1_spec__10_spec__20_spec__25_spec__26_spec__27(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_50; uint8_t x_51; lean_object* x_52; uint8_t x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; lean_object* x_57; lean_object* x_77; uint8_t x_78; lean_object* x_79; uint8_t x_80; lean_object* x_81; lean_object* x_82; uint8_t x_83; lean_object* x_84; lean_object* x_88; uint8_t x_89; lean_object* x_90; uint8_t x_91; lean_object* x_92; lean_object* x_93; uint8_t x_94; uint8_t x_100; lean_object* x_101; uint8_t x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; uint8_t x_106; uint8_t x_107; uint8_t x_109; uint8_t x_125; +lean_object* x_10; uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_50; lean_object* x_51; uint8_t x_52; uint8_t x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; lean_object* x_57; lean_object* x_77; lean_object* x_78; uint8_t x_79; uint8_t x_80; lean_object* x_81; lean_object* x_82; uint8_t x_83; lean_object* x_84; lean_object* x_88; lean_object* x_89; uint8_t x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; uint8_t x_94; uint8_t x_100; lean_object* x_101; uint8_t x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; uint8_t x_106; uint8_t x_107; uint8_t x_109; uint8_t x_125; x_100 = 2; x_125 = l_Lean_instBEqMessageSeverity_beq(x_3, x_100); if (x_125 == 0) @@ -22286,15 +22286,15 @@ lean_ctor_set(x_25, 0, x_21); lean_ctor_set(x_25, 1, x_22); x_26 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_10); +lean_ctor_set(x_26, 1, x_16); x_27 = lean_alloc_ctor(0, 5, 3); -lean_ctor_set(x_27, 0, x_14); -lean_ctor_set(x_27, 1, x_15); -lean_ctor_set(x_27, 2, x_13); -lean_ctor_set(x_27, 3, x_11); +lean_ctor_set(x_27, 0, x_10); +lean_ctor_set(x_27, 1, x_13); +lean_ctor_set(x_27, 2, x_12); +lean_ctor_set(x_27, 3, x_14); lean_ctor_set(x_27, 4, x_26); -lean_ctor_set_uint8(x_27, sizeof(void*)*5, x_12); -lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 1, x_16); +lean_ctor_set_uint8(x_27, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 1, x_11); lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 2, x_4); x_28 = l_Lean_MessageLog_add(x_27, x_24); lean_ctor_set(x_20, 6, x_28); @@ -22331,15 +22331,15 @@ lean_ctor_set(x_41, 0, x_21); lean_ctor_set(x_41, 1, x_22); x_42 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_10); +lean_ctor_set(x_42, 1, x_16); x_43 = lean_alloc_ctor(0, 5, 3); -lean_ctor_set(x_43, 0, x_14); -lean_ctor_set(x_43, 1, x_15); -lean_ctor_set(x_43, 2, x_13); -lean_ctor_set(x_43, 3, x_11); +lean_ctor_set(x_43, 0, x_10); +lean_ctor_set(x_43, 1, x_13); +lean_ctor_set(x_43, 2, x_12); +lean_ctor_set(x_43, 3, x_14); lean_ctor_set(x_43, 4, x_42); -lean_ctor_set_uint8(x_43, sizeof(void*)*5, x_12); -lean_ctor_set_uint8(x_43, sizeof(void*)*5 + 1, x_16); +lean_ctor_set_uint8(x_43, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_43, sizeof(void*)*5 + 1, x_11); lean_ctor_set_uint8(x_43, sizeof(void*)*5 + 2, x_4); x_44 = l_Lean_MessageLog_add(x_43, x_38); x_45 = lean_alloc_ctor(0, 9, 0); @@ -22369,25 +22369,25 @@ if (x_60 == 0) { lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; x_61 = lean_ctor_get(x_59, 0); -lean_inc_ref(x_55); -x_62 = l_Lean_FileMap_toPosition(x_55, x_52); -lean_dec(x_52); -x_63 = l_Lean_FileMap_toPosition(x_55, x_57); +lean_inc_ref(x_54); +x_62 = l_Lean_FileMap_toPosition(x_54, x_55); +lean_dec(x_55); +x_63 = l_Lean_FileMap_toPosition(x_54, x_57); lean_dec(x_57); x_64 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_64, 0, x_63); x_65 = l_Lean_addTrace___at___00__private_Lean_ExtraModUses_0__Lean_recordExtraModUseCore___at___00Lean_recordExtraModUseFromDecl___at___00Lean_Elab_Tactic_elabSimpConfigCore_spec__0_spec__0_spec__3___redArg___closed__1; -if (x_51 == 0) +if (x_53 == 0) { lean_free_object(x_59); lean_dec_ref(x_50); -x_10 = x_61; -x_11 = x_65; -x_12 = x_53; -x_13 = x_64; -x_14 = x_54; -x_15 = x_62; -x_16 = x_56; +x_10 = x_51; +x_11 = x_52; +x_12 = x_64; +x_13 = x_62; +x_14 = x_65; +x_15 = x_56; +x_16 = x_61; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -22404,7 +22404,7 @@ lean_object* x_67; lean_dec_ref(x_64); lean_dec_ref(x_62); lean_dec(x_61); -lean_dec_ref(x_54); +lean_dec_ref(x_51); lean_dec_ref(x_7); x_67 = lean_box(0); lean_ctor_set(x_59, 0, x_67); @@ -22413,13 +22413,13 @@ return x_59; else { lean_free_object(x_59); -x_10 = x_61; -x_11 = x_65; -x_12 = x_53; -x_13 = x_64; -x_14 = x_54; -x_15 = x_62; -x_16 = x_56; +x_10 = x_51; +x_11 = x_52; +x_12 = x_64; +x_13 = x_62; +x_14 = x_65; +x_15 = x_56; +x_16 = x_61; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -22433,24 +22433,24 @@ lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean x_68 = lean_ctor_get(x_59, 0); lean_inc(x_68); lean_dec(x_59); -lean_inc_ref(x_55); -x_69 = l_Lean_FileMap_toPosition(x_55, x_52); -lean_dec(x_52); -x_70 = l_Lean_FileMap_toPosition(x_55, x_57); +lean_inc_ref(x_54); +x_69 = l_Lean_FileMap_toPosition(x_54, x_55); +lean_dec(x_55); +x_70 = l_Lean_FileMap_toPosition(x_54, x_57); lean_dec(x_57); x_71 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_71, 0, x_70); x_72 = l_Lean_addTrace___at___00__private_Lean_ExtraModUses_0__Lean_recordExtraModUseCore___at___00Lean_recordExtraModUseFromDecl___at___00Lean_Elab_Tactic_elabSimpConfigCore_spec__0_spec__0_spec__3___redArg___closed__1; -if (x_51 == 0) +if (x_53 == 0) { lean_dec_ref(x_50); -x_10 = x_68; -x_11 = x_72; -x_12 = x_53; -x_13 = x_71; -x_14 = x_54; -x_15 = x_69; -x_16 = x_56; +x_10 = x_51; +x_11 = x_52; +x_12 = x_71; +x_13 = x_69; +x_14 = x_72; +x_15 = x_56; +x_16 = x_68; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -22467,7 +22467,7 @@ lean_object* x_74; lean_object* x_75; lean_dec_ref(x_71); lean_dec_ref(x_69); lean_dec(x_68); -lean_dec_ref(x_54); +lean_dec_ref(x_51); lean_dec_ref(x_7); x_74 = lean_box(0); x_75 = lean_alloc_ctor(0, 1, 0); @@ -22476,13 +22476,13 @@ return x_75; } else { -x_10 = x_68; -x_11 = x_72; -x_12 = x_53; -x_13 = x_71; -x_14 = x_54; -x_15 = x_69; -x_16 = x_56; +x_10 = x_51; +x_11 = x_52; +x_12 = x_71; +x_13 = x_69; +x_14 = x_72; +x_15 = x_56; +x_16 = x_68; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -22494,17 +22494,17 @@ goto block_49; block_87: { lean_object* x_85; -x_85 = l_Lean_Syntax_getTailPos_x3f(x_79, x_80); -lean_dec(x_79); +x_85 = l_Lean_Syntax_getTailPos_x3f(x_81, x_83); +lean_dec(x_81); if (lean_obj_tag(x_85) == 0) { lean_inc(x_84); x_50 = x_77; x_51 = x_78; -x_52 = x_84; +x_52 = x_79; x_53 = x_80; x_54 = x_82; -x_55 = x_81; +x_55 = x_84; x_56 = x_83; x_57 = x_84; goto block_76; @@ -22517,10 +22517,10 @@ lean_inc(x_86); lean_dec_ref(x_85); x_50 = x_77; x_51 = x_78; -x_52 = x_84; +x_52 = x_79; x_53 = x_80; x_54 = x_82; -x_55 = x_81; +x_55 = x_84; x_56 = x_83; x_57 = x_86; goto block_76; @@ -22529,20 +22529,20 @@ goto block_76; block_99: { lean_object* x_95; lean_object* x_96; -x_95 = l_Lean_replaceRef(x_1, x_90); -lean_dec(x_90); -x_96 = l_Lean_Syntax_getPos_x3f(x_95, x_91); +x_95 = l_Lean_replaceRef(x_1, x_92); +lean_dec(x_92); +x_96 = l_Lean_Syntax_getPos_x3f(x_95, x_93); if (lean_obj_tag(x_96) == 0) { lean_object* x_97; x_97 = lean_unsigned_to_nat(0u); x_77 = x_88; x_78 = x_89; -x_79 = x_95; -x_80 = x_91; -x_81 = x_93; -x_82 = x_92; -x_83 = x_94; +x_79 = x_94; +x_80 = x_90; +x_81 = x_95; +x_82 = x_91; +x_83 = x_93; x_84 = x_97; goto block_87; } @@ -22554,11 +22554,11 @@ lean_inc(x_98); lean_dec_ref(x_96); x_77 = x_88; x_78 = x_89; -x_79 = x_95; -x_80 = x_91; -x_81 = x_93; -x_82 = x_92; -x_83 = x_94; +x_79 = x_94; +x_80 = x_90; +x_81 = x_95; +x_82 = x_91; +x_83 = x_93; x_84 = x_98; goto block_87; } @@ -22567,23 +22567,23 @@ goto block_87; { if (x_107 == 0) { -x_88 = x_101; -x_89 = x_102; -x_90 = x_103; -x_91 = x_106; -x_92 = x_105; -x_93 = x_104; +x_88 = x_105; +x_89 = x_101; +x_90 = x_102; +x_91 = x_103; +x_92 = x_104; +x_93 = x_106; x_94 = x_3; goto block_99; } else { -x_88 = x_101; -x_89 = x_102; -x_90 = x_103; -x_91 = x_106; -x_92 = x_105; -x_93 = x_104; +x_88 = x_105; +x_89 = x_101; +x_90 = x_102; +x_91 = x_103; +x_92 = x_104; +x_93 = x_106; x_94 = x_100; goto block_99; } @@ -22607,14 +22607,14 @@ x_118 = 1; x_119 = l_Lean_instBEqMessageSeverity_beq(x_3, x_118); if (x_119 == 0) { -lean_inc_ref(x_110); -lean_inc_ref(x_111); lean_inc(x_113); -x_101 = x_117; +lean_inc_ref(x_111); +lean_inc_ref(x_110); +x_101 = x_110; x_102 = x_114; -x_103 = x_113; -x_104 = x_111; -x_105 = x_110; +x_103 = x_111; +x_104 = x_113; +x_105 = x_117; x_106 = x_109; x_107 = x_119; goto block_108; @@ -22624,14 +22624,14 @@ else lean_object* x_120; uint8_t x_121; x_120 = l_Lean_logAt___at___00Lean_logErrorAt___at___00Lean_Elab_logException___at___00__private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArg_spec__1_spec__2_spec__7___redArg___closed__0; x_121 = l_Lean_Option_get___at___00Lean_Elab_addMacroStack___at___00Lean_throwError___at___00Lean_Elab_Tactic_elabSimpConfigCore_spec__1_spec__5_spec__10(x_112, x_120); -lean_inc_ref(x_110); -lean_inc_ref(x_111); lean_inc(x_113); -x_101 = x_117; +lean_inc_ref(x_111); +lean_inc_ref(x_110); +x_101 = x_110; x_102 = x_114; -x_103 = x_113; -x_104 = x_111; -x_105 = x_110; +x_103 = x_111; +x_104 = x_113; +x_105 = x_117; x_106 = x_109; x_107 = x_121; goto block_108; @@ -25848,7 +25848,7 @@ lean_inc_ref(x_5); x_163 = lean_local_ctx_find(x_5, x_162); if (lean_obj_tag(x_163) == 1) { -lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_175; lean_object* x_176; uint8_t x_182; lean_object* x_183; lean_object* x_184; uint8_t x_187; +lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_175; lean_object* x_176; lean_object* x_182; uint8_t x_183; lean_object* x_184; uint8_t x_187; x_164 = lean_ctor_get(x_163, 0); lean_inc(x_164); if (lean_is_exclusive(x_163)) { @@ -25949,23 +25949,23 @@ uint8_t x_185; x_185 = l_Lean_Name_hasMacroScopes(x_184); if (x_185 == 0) { -x_175 = x_183; +x_175 = x_182; x_176 = x_184; goto block_181; } else { -if (x_182 == 0) +if (x_183 == 0) { lean_dec(x_184); -lean_dec_ref(x_183); +lean_dec_ref(x_182); lean_dec(x_165); lean_dec(x_164); goto block_43; } else { -x_175 = x_183; +x_175 = x_182; x_176 = x_184; goto block_181; } @@ -25991,8 +25991,8 @@ lean_inc(x_189); x_190 = lean_is_inaccessible_user_name(x_189); if (x_190 == 0) { -x_182 = x_187; -x_183 = x_188; +x_182 = x_188; +x_183 = x_187; x_184 = x_189; goto block_186; } @@ -26008,8 +26008,8 @@ goto block_43; } else { -x_182 = x_187; -x_183 = x_188; +x_182 = x_188; +x_183 = x_187; x_184 = x_189; goto block_186; } @@ -26718,7 +26718,7 @@ lean_inc_ref(x_6); x_164 = lean_local_ctx_find(x_6, x_163); if (lean_obj_tag(x_164) == 1) { -lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_176; lean_object* x_177; lean_object* x_183; uint8_t x_184; lean_object* x_185; uint8_t x_188; +lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_176; lean_object* x_177; uint8_t x_183; lean_object* x_184; lean_object* x_185; uint8_t x_188; x_165 = lean_ctor_get(x_164, 0); lean_inc(x_165); if (lean_is_exclusive(x_164)) { @@ -26768,15 +26768,15 @@ lean_dec(x_171); lean_dec(x_170); if (x_172 == 0) { -lean_dec_ref(x_168); -lean_dec(x_167); +lean_dec(x_168); +lean_dec_ref(x_167); lean_dec(x_166); goto block_44; } else { lean_object* x_173; lean_object* x_174; -x_173 = lean_array_push(x_168, x_167); +x_173 = lean_array_push(x_167, x_168); if (lean_is_scalar(x_166)) { x_174 = lean_alloc_ctor(1, 1, 0); } else { @@ -26790,7 +26790,7 @@ goto block_42; block_182: { lean_object* x_178; -x_178 = l_Lean_LocalContext_findFromUserName_x3f(x_6, x_176); +x_178 = l_Lean_LocalContext_findFromUserName_x3f(x_6, x_177); if (lean_obj_tag(x_178) == 0) { lean_object* x_179; lean_object* x_180; @@ -26816,26 +26816,26 @@ goto block_175; block_187: { uint8_t x_186; -x_186 = l_Lean_Name_hasMacroScopes(x_183); +x_186 = l_Lean_Name_hasMacroScopes(x_185); if (x_186 == 0) { -x_176 = x_183; +x_176 = x_184; x_177 = x_185; goto block_182; } else { -if (x_184 == 0) +if (x_183 == 0) { -lean_dec_ref(x_185); -lean_dec(x_183); +lean_dec(x_185); +lean_dec_ref(x_184); lean_dec(x_166); lean_dec(x_165); goto block_44; } else { -x_176 = x_183; +x_176 = x_184; x_177 = x_185; goto block_182; } @@ -26861,9 +26861,9 @@ lean_inc(x_190); x_191 = lean_is_inaccessible_user_name(x_190); if (x_191 == 0) { -x_183 = x_190; -x_184 = x_188; -x_185 = x_189; +x_183 = x_188; +x_184 = x_189; +x_185 = x_190; goto block_187; } else @@ -26878,9 +26878,9 @@ goto block_44; } else { -x_183 = x_190; -x_184 = x_188; -x_185 = x_189; +x_183 = x_188; +x_184 = x_189; +x_185 = x_190; goto block_187; } } @@ -31330,8 +31330,8 @@ return x_175; if (x_24 == 0) { lean_object* x_25; -lean_dec(x_23); -lean_dec_ref(x_20); +lean_dec_ref(x_21); +lean_dec(x_20); lean_dec(x_19); lean_dec_ref(x_18); lean_dec(x_16); @@ -31341,15 +31341,15 @@ if (lean_is_scalar(x_17)) { x_25 = x_17; lean_ctor_set_tag(x_25, 1); } -lean_ctor_set(x_25, 0, x_21); +lean_ctor_set(x_25, 0, x_23); return x_25; } else { lean_object* x_26; lean_dec(x_17); -x_26 = l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_withLoopChecking_go___at___00Lean_Elab_Tactic_evalSimp_spec__0___redArg(x_16, x_24, x_20, x_23, x_18, x_19); -lean_dec_ref(x_18); +x_26 = l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_withLoopChecking_go___at___00Lean_Elab_Tactic_evalSimp_spec__0___redArg(x_16, x_24, x_18, x_20, x_21, x_19); +lean_dec_ref(x_21); if (lean_obj_tag(x_26) == 0) { uint8_t x_27; @@ -31360,7 +31360,7 @@ lean_object* x_28; x_28 = lean_ctor_get(x_26, 0); lean_dec(x_28); lean_ctor_set_tag(x_26, 1); -lean_ctor_set(x_26, 0, x_21); +lean_ctor_set(x_26, 0, x_23); return x_26; } else @@ -31368,14 +31368,14 @@ else lean_object* x_29; lean_dec(x_26); x_29 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_29, 0, x_21); +lean_ctor_set(x_29, 0, x_23); return x_29; } } else { uint8_t x_30; -lean_dec_ref(x_21); +lean_dec_ref(x_23); x_30 = !lean_is_exclusive(x_26); if (x_30 == 0) { @@ -31410,11 +31410,11 @@ lean_inc_ref(x_38); x_42 = l_Lean_Exception_toMessageData(x_38); x_43 = l_Lean_MessageData_hasTag(x_5, x_42); x_18 = x_34; -x_19 = x_35; -x_20 = x_36; -x_21 = x_38; +x_19 = x_36; +x_20 = x_35; +x_21 = x_37; x_22 = lean_box(0); -x_23 = x_37; +x_23 = x_38; x_24 = x_43; goto block_33; } @@ -31422,11 +31422,11 @@ else { lean_dec_ref(x_5); x_18 = x_34; -x_19 = x_35; -x_20 = x_36; -x_21 = x_38; +x_19 = x_36; +x_20 = x_35; +x_21 = x_37; x_22 = lean_box(0); -x_23 = x_37; +x_23 = x_38; x_24 = x_41; goto block_33; } @@ -31434,8 +31434,8 @@ goto block_33; else { lean_object* x_44; -lean_dec(x_37); -lean_dec_ref(x_36); +lean_dec_ref(x_37); +lean_dec(x_36); lean_dec(x_35); lean_dec_ref(x_34); lean_dec(x_17); @@ -32006,10 +32006,10 @@ lean_dec(x_1); x_166 = lean_ctor_get(x_135, 0); lean_inc(x_166); lean_dec_ref(x_135); -x_34 = x_129; -x_35 = x_130; -x_36 = x_127; -x_37 = x_128; +x_34 = x_127; +x_35 = x_128; +x_36 = x_130; +x_37 = x_129; x_38 = x_166; x_39 = lean_box(0); goto block_45; @@ -32023,10 +32023,10 @@ lean_dec(x_1); x_167 = lean_ctor_get(x_132, 0); lean_inc(x_167); lean_dec_ref(x_132); -x_34 = x_129; -x_35 = x_130; -x_36 = x_127; -x_37 = x_128; +x_34 = x_127; +x_35 = x_128; +x_36 = x_130; +x_37 = x_129; x_38 = x_167; x_39 = lean_box(0); goto block_45; diff --git a/stage0/stdlib/Lean/Meta/IndPredBelow.c b/stage0/stdlib/Lean/Meta/IndPredBelow.c index 410c6ce28759..5191560ff954 100644 --- a/stage0/stdlib/Lean/Meta/IndPredBelow.c +++ b/stage0/stdlib/Lean/Meta/IndPredBelow.c @@ -11245,7 +11245,7 @@ return x_34; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00__private_Lean_Meta_IndPredBelow_0__Lean_Meta_IndPredBelow_mkBelowMatcher_convertToBelow_spec__8___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; double x_30; uint8_t x_31; lean_object* x_32; double x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; double x_41; uint8_t x_42; lean_object* x_43; lean_object* x_44; double x_45; lean_object* x_46; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_67; double x_68; lean_object* x_69; double x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; uint8_t x_78; lean_object* x_79; double x_80; lean_object* x_81; double x_82; lean_object* x_83; lean_object* x_88; lean_object* x_89; double x_90; uint8_t x_91; lean_object* x_92; lean_object* x_93; double x_94; lean_object* x_95; uint8_t x_96; double x_130; uint8_t x_131; lean_object* x_132; lean_object* x_133; double x_134; lean_object* x_135; double x_136; lean_object* x_140; lean_object* x_141; uint8_t x_142; lean_object* x_143; lean_object* x_144; uint8_t x_162; double x_163; lean_object* x_164; lean_object* x_165; double x_166; lean_object* x_167; uint8_t x_168; uint8_t x_202; lean_object* x_203; double x_204; lean_object* x_205; double x_206; lean_object* x_207; double x_208; lean_object* x_212; lean_object* x_213; uint8_t x_214; lean_object* x_215; lean_object* x_216; uint8_t x_252; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_30; lean_object* x_31; uint8_t x_32; double x_33; double x_34; lean_object* x_35; lean_object* x_36; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; double x_45; double x_46; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_67; double x_68; double x_69; lean_object* x_70; uint8_t x_71; lean_object* x_72; lean_object* x_73; lean_object* x_78; double x_79; lean_object* x_80; lean_object* x_81; double x_82; uint8_t x_83; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; double x_94; double x_95; uint8_t x_96; lean_object* x_130; lean_object* x_131; lean_object* x_132; uint8_t x_133; double x_134; double x_135; double x_136; lean_object* x_140; uint8_t x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_162; double x_163; lean_object* x_164; double x_165; lean_object* x_166; uint8_t x_167; uint8_t x_168; lean_object* x_202; double x_203; lean_object* x_204; lean_object* x_205; double x_206; uint8_t x_207; double x_208; lean_object* x_212; lean_object* x_213; uint8_t x_214; lean_object* x_215; lean_object* x_216; uint8_t x_252; x_12 = lean_ctor_get(x_9, 2); x_13 = lean_ctor_get(x_9, 5); lean_inc(x_1); @@ -11284,7 +11284,7 @@ goto block_251; { lean_object* x_24; lean_dec(x_18); -x_24 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00__private_Lean_Meta_IndPredBelow_0__Lean_Meta_IndPredBelow_mkBelowMatcher_convertToBelow_spec__8_spec__11___redArg(x_14, x_17, x_13, x_15, x_19, x_20, x_21, x_22); +x_24 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00__private_Lean_Meta_IndPredBelow_0__Lean_Meta_IndPredBelow_mkBelowMatcher_convertToBelow_spec__8_spec__11___redArg(x_14, x_17, x_13, x_16, x_19, x_20, x_21, x_22); lean_dec(x_22); lean_dec(x_20); lean_dec_ref(x_19); @@ -11292,13 +11292,13 @@ if (lean_obj_tag(x_24) == 0) { lean_object* x_25; lean_dec_ref(x_24); -x_25 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00__private_Lean_Meta_IndPredBelow_0__Lean_Meta_IndPredBelow_mkBelowMatcher_convertToBelow_spec__8_spec__12___redArg(x_16); +x_25 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00__private_Lean_Meta_IndPredBelow_0__Lean_Meta_IndPredBelow_mkBelowMatcher_convertToBelow_spec__8_spec__12___redArg(x_15); return x_25; } else { uint8_t x_26; -lean_dec_ref(x_16); +lean_dec_ref(x_15); x_26 = !lean_is_exclusive(x_24); if (x_26 == 0) { @@ -11318,7 +11318,7 @@ return x_28; } block_40: { -if (x_31 == 0) +if (x_32 == 0) { double x_37; lean_object* x_38; x_37 = l_Lean_withTraceNode___at___00Lean_Meta_IndPredBelow_mkBelow_spec__8___redArg___closed__0; @@ -11328,9 +11328,9 @@ lean_ctor_set(x_38, 1, x_5); lean_ctor_set_float(x_38, sizeof(void*)*2, x_37); lean_ctor_set_float(x_38, sizeof(void*)*2 + 8, x_37); lean_ctor_set_uint8(x_38, sizeof(void*)*2 + 16, x_4); -x_14 = x_32; -x_15 = x_35; -x_16 = x_34; +x_14 = x_30; +x_15 = x_31; +x_16 = x_35; x_17 = x_38; x_18 = x_6; x_19 = x_7; @@ -11346,12 +11346,12 @@ lean_object* x_39; x_39 = lean_alloc_ctor(0, 2, 17); lean_ctor_set(x_39, 0, x_1); lean_ctor_set(x_39, 1, x_5); -lean_ctor_set_float(x_39, sizeof(void*)*2, x_30); -lean_ctor_set_float(x_39, sizeof(void*)*2 + 8, x_33); +lean_ctor_set_float(x_39, sizeof(void*)*2, x_33); +lean_ctor_set_float(x_39, sizeof(void*)*2 + 8, x_34); lean_ctor_set_uint8(x_39, sizeof(void*)*2 + 16, x_4); -x_14 = x_32; -x_15 = x_35; -x_16 = x_34; +x_14 = x_30; +x_15 = x_31; +x_16 = x_35; x_17 = x_39; x_18 = x_6; x_19 = x_7; @@ -11370,17 +11370,17 @@ lean_inc_ref(x_9); lean_inc(x_8); lean_inc_ref(x_7); lean_inc(x_6); -lean_inc_ref(x_46); -x_47 = lean_apply_7(x_2, x_46, x_6, x_7, x_8, x_9, x_10, lean_box(0)); +lean_inc_ref(x_43); +x_47 = lean_apply_7(x_2, x_43, x_6, x_7, x_8, x_9, x_10, lean_box(0)); if (lean_obj_tag(x_47) == 0) { lean_object* x_48; x_48 = lean_ctor_get(x_47, 0); lean_inc(x_48); lean_dec_ref(x_47); -x_30 = x_41; -x_31 = x_42; -x_32 = x_43; +x_30 = x_42; +x_31 = x_43; +x_32 = x_44; x_33 = x_45; x_34 = x_46; x_35 = x_48; @@ -11392,9 +11392,9 @@ else lean_object* x_49; lean_dec_ref(x_47); x_49 = l_Lean_withTraceNode___at___00Lean_Meta_IndPredBelow_mkBelow_spec__8___redArg___closed__2; -x_30 = x_41; -x_31 = x_42; -x_32 = x_43; +x_30 = x_42; +x_31 = x_43; +x_32 = x_44; x_33 = x_45; x_34 = x_46; x_35 = x_49; @@ -11406,7 +11406,7 @@ goto block_40; { lean_object* x_61; lean_dec(x_55); -x_61 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00__private_Lean_Meta_IndPredBelow_0__Lean_Meta_IndPredBelow_mkBelowMatcher_convertToBelow_spec__8_spec__11___redArg(x_51, x_54, x_13, x_52, x_56, x_57, x_58, x_59); +x_61 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00__private_Lean_Meta_IndPredBelow_0__Lean_Meta_IndPredBelow_mkBelowMatcher_convertToBelow_spec__8_spec__11___redArg(x_52, x_54, x_13, x_51, x_56, x_57, x_58, x_59); lean_dec(x_59); lean_dec(x_57); lean_dec_ref(x_56); @@ -11440,7 +11440,7 @@ return x_65; } block_77: { -if (x_67 == 0) +if (x_71 == 0) { double x_74; lean_object* x_75; x_74 = l_Lean_withTraceNode___at___00Lean_Meta_IndPredBelow_mkBelow_spec__8___redArg___closed__0; @@ -11450,9 +11450,9 @@ lean_ctor_set(x_75, 1, x_5); lean_ctor_set_float(x_75, sizeof(void*)*2, x_74); lean_ctor_set_float(x_75, sizeof(void*)*2 + 8, x_74); lean_ctor_set_uint8(x_75, sizeof(void*)*2 + 16, x_4); -x_51 = x_69; -x_52 = x_72; -x_53 = x_71; +x_51 = x_72; +x_52 = x_67; +x_53 = x_70; x_54 = x_75; x_55 = x_6; x_56 = x_7; @@ -11468,12 +11468,12 @@ lean_object* x_76; x_76 = lean_alloc_ctor(0, 2, 17); lean_ctor_set(x_76, 0, x_1); lean_ctor_set(x_76, 1, x_5); -lean_ctor_set_float(x_76, sizeof(void*)*2, x_68); -lean_ctor_set_float(x_76, sizeof(void*)*2 + 8, x_70); +lean_ctor_set_float(x_76, sizeof(void*)*2, x_69); +lean_ctor_set_float(x_76, sizeof(void*)*2 + 8, x_68); lean_ctor_set_uint8(x_76, sizeof(void*)*2 + 16, x_4); -x_51 = x_69; -x_52 = x_72; -x_53 = x_71; +x_51 = x_72; +x_52 = x_67; +x_53 = x_70; x_54 = x_76; x_55 = x_6; x_56 = x_7; @@ -11492,8 +11492,8 @@ lean_inc_ref(x_9); lean_inc(x_8); lean_inc_ref(x_7); lean_inc(x_6); -lean_inc_ref(x_83); -x_84 = lean_apply_7(x_2, x_83, x_6, x_7, x_8, x_9, x_10, lean_box(0)); +lean_inc_ref(x_81); +x_84 = lean_apply_7(x_2, x_81, x_6, x_7, x_8, x_9, x_10, lean_box(0)); if (lean_obj_tag(x_84) == 0) { lean_object* x_85; @@ -11501,9 +11501,9 @@ x_85 = lean_ctor_get(x_84, 0); lean_inc(x_85); lean_dec_ref(x_84); x_67 = x_78; -x_68 = x_80; -x_69 = x_79; -x_70 = x_82; +x_68 = x_79; +x_69 = x_82; +x_70 = x_81; x_71 = x_83; x_72 = x_85; x_73 = lean_box(0); @@ -11515,9 +11515,9 @@ lean_object* x_86; lean_dec_ref(x_84); x_86 = l_Lean_withTraceNode___at___00Lean_Meta_IndPredBelow_mkBelow_spec__8___redArg___closed__2; x_67 = x_78; -x_68 = x_80; -x_69 = x_79; -x_70 = x_82; +x_68 = x_79; +x_69 = x_82; +x_70 = x_81; x_71 = x_83; x_72 = x_86; x_73 = lean_box(0); @@ -11553,12 +11553,12 @@ if (x_101 == 0) { lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; x_102 = lean_ctor_get(x_100, 0); -x_103 = l_Lean_PersistentArray_append___redArg(x_92, x_102); +x_103 = l_Lean_PersistentArray_append___redArg(x_90, x_102); lean_dec_ref(x_102); lean_ctor_set(x_100, 0, x_103); x_104 = lean_st_ref_set(x_10, x_98); lean_dec(x_10); -x_105 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00__private_Lean_Meta_IndPredBelow_0__Lean_Meta_IndPredBelow_mkBelowMatcher_convertToBelow_spec__8_spec__12___redArg(x_95); +x_105 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00__private_Lean_Meta_IndPredBelow_0__Lean_Meta_IndPredBelow_mkBelowMatcher_convertToBelow_spec__8_spec__12___redArg(x_92); return x_105; } else @@ -11568,7 +11568,7 @@ x_106 = lean_ctor_get_uint64(x_100, sizeof(void*)*1); x_107 = lean_ctor_get(x_100, 0); lean_inc(x_107); lean_dec(x_100); -x_108 = l_Lean_PersistentArray_append___redArg(x_92, x_107); +x_108 = l_Lean_PersistentArray_append___redArg(x_90, x_107); lean_dec_ref(x_107); x_109 = lean_alloc_ctor(0, 1, 8); lean_ctor_set(x_109, 0, x_108); @@ -11576,7 +11576,7 @@ lean_ctor_set_uint64(x_109, sizeof(void*)*1, x_106); lean_ctor_set(x_98, 4, x_109); x_110 = lean_st_ref_set(x_10, x_98); lean_dec(x_10); -x_111 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00__private_Lean_Meta_IndPredBelow_0__Lean_Meta_IndPredBelow_mkBelowMatcher_convertToBelow_spec__8_spec__12___redArg(x_95); +x_111 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00__private_Lean_Meta_IndPredBelow_0__Lean_Meta_IndPredBelow_mkBelowMatcher_convertToBelow_spec__8_spec__12___redArg(x_92); return x_111; } } @@ -11612,7 +11612,7 @@ if (lean_is_exclusive(x_112)) { lean_dec_ref(x_112); x_123 = lean_box(0); } -x_124 = l_Lean_PersistentArray_append___redArg(x_92, x_122); +x_124 = l_Lean_PersistentArray_append___redArg(x_90, x_122); lean_dec_ref(x_122); if (lean_is_scalar(x_123)) { x_125 = lean_alloc_ctor(0, 1, 8); @@ -11633,16 +11633,16 @@ lean_ctor_set(x_126, 7, x_119); lean_ctor_set(x_126, 8, x_120); x_127 = lean_st_ref_set(x_10, x_126); lean_dec(x_10); -x_128 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00__private_Lean_Meta_IndPredBelow_0__Lean_Meta_IndPredBelow_mkBelowMatcher_convertToBelow_spec__8_spec__12___redArg(x_95); +x_128 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00__private_Lean_Meta_IndPredBelow_0__Lean_Meta_IndPredBelow_mkBelowMatcher_convertToBelow_spec__8_spec__12___redArg(x_92); return x_128; } } else { -x_41 = x_90; -x_42 = x_91; +x_41 = lean_box(0); +x_42 = x_90; x_43 = x_92; -x_44 = lean_box(0); +x_44 = x_93; x_45 = x_94; x_46 = x_95; goto block_50; @@ -11650,10 +11650,10 @@ goto block_50; } else { -x_41 = x_90; -x_42 = x_91; +x_41 = lean_box(0); +x_42 = x_90; x_43 = x_92; -x_44 = lean_box(0); +x_44 = x_93; x_45 = x_94; x_46 = x_95; goto block_50; @@ -11662,12 +11662,12 @@ goto block_50; block_139: { double x_137; uint8_t x_138; -x_137 = lean_float_sub(x_134, x_130); +x_137 = lean_float_sub(x_135, x_134); x_138 = lean_float_decLt(x_136, x_137); -x_90 = x_130; -x_91 = x_131; +x_90 = x_131; +x_91 = lean_box(0); x_92 = x_132; -x_93 = lean_box(0); +x_93 = x_133; x_94 = x_134; x_95 = x_135; x_96 = x_138; @@ -11677,7 +11677,7 @@ goto block_129; { lean_object* x_145; double x_146; double x_147; double x_148; double x_149; double x_150; lean_object* x_151; uint8_t x_152; x_145 = lean_io_mono_nanos_now(); -x_146 = lean_float_of_nat(x_140); +x_146 = lean_float_of_nat(x_142); x_147 = l_Lean_withTraceNode___at___00Lean_Meta_IndPredBelow_mkBelow_spec__8___redArg___closed__3; x_148 = lean_float_div(x_146, x_147); x_149 = lean_float_of_nat(x_145); @@ -11686,18 +11686,18 @@ x_151 = l_Lean_withTraceNode___at___00Lean_Meta_IndPredBelow_mkBelow_spec__8___r x_152 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00Lean_Meta_IndPredBelow_mkBelow_spec__8_spec__12(x_12, x_151); if (x_152 == 0) { -x_90 = x_148; -x_91 = x_152; -x_92 = x_141; -x_93 = lean_box(0); -x_94 = x_150; -x_95 = x_143; +x_90 = x_140; +x_91 = lean_box(0); +x_92 = x_143; +x_93 = x_152; +x_94 = x_148; +x_95 = x_150; x_96 = x_152; goto block_129; } else { -if (x_142 == 0) +if (x_141 == 0) { lean_object* x_153; lean_object* x_154; double x_155; double x_156; double x_157; x_153 = l_Lean_withTraceNode___at___00Lean_Meta_IndPredBelow_mkBelow_spec__8___redArg___closed__5; @@ -11705,12 +11705,12 @@ x_154 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00Lean_Meta_IndPred x_155 = lean_float_of_nat(x_154); x_156 = l_Lean_withTraceNode___at___00Lean_Meta_IndPredBelow_mkBelow_spec__8___redArg___closed__6; x_157 = lean_float_div(x_155, x_156); -x_130 = x_148; -x_131 = x_152; -x_132 = x_141; -x_133 = lean_box(0); -x_134 = x_150; -x_135 = x_143; +x_130 = lean_box(0); +x_131 = x_140; +x_132 = x_143; +x_133 = x_152; +x_134 = x_148; +x_135 = x_150; x_136 = x_157; goto block_139; } @@ -11720,12 +11720,12 @@ lean_object* x_158; lean_object* x_159; double x_160; x_158 = l_Lean_withTraceNode___at___00Lean_Meta_IndPredBelow_mkBelow_spec__8___redArg___closed__5; x_159 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00Lean_Meta_IndPredBelow_mkBelow_spec__8_spec__15(x_12, x_158); x_160 = lean_float_of_nat(x_159); -x_130 = x_148; -x_131 = x_152; -x_132 = x_141; -x_133 = lean_box(0); -x_134 = x_150; -x_135 = x_143; +x_130 = lean_box(0); +x_131 = x_140; +x_132 = x_143; +x_133 = x_152; +x_134 = x_148; +x_135 = x_150; x_136 = x_160; goto block_139; } @@ -11760,12 +11760,12 @@ if (x_173 == 0) { lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; x_174 = lean_ctor_get(x_172, 0); -x_175 = l_Lean_PersistentArray_append___redArg(x_164, x_174); +x_175 = l_Lean_PersistentArray_append___redArg(x_162, x_174); lean_dec_ref(x_174); lean_ctor_set(x_172, 0, x_175); x_176 = lean_st_ref_set(x_10, x_170); lean_dec(x_10); -x_177 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00__private_Lean_Meta_IndPredBelow_0__Lean_Meta_IndPredBelow_mkBelowMatcher_convertToBelow_spec__8_spec__12___redArg(x_167); +x_177 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00__private_Lean_Meta_IndPredBelow_0__Lean_Meta_IndPredBelow_mkBelowMatcher_convertToBelow_spec__8_spec__12___redArg(x_166); return x_177; } else @@ -11775,7 +11775,7 @@ x_178 = lean_ctor_get_uint64(x_172, sizeof(void*)*1); x_179 = lean_ctor_get(x_172, 0); lean_inc(x_179); lean_dec(x_172); -x_180 = l_Lean_PersistentArray_append___redArg(x_164, x_179); +x_180 = l_Lean_PersistentArray_append___redArg(x_162, x_179); lean_dec_ref(x_179); x_181 = lean_alloc_ctor(0, 1, 8); lean_ctor_set(x_181, 0, x_180); @@ -11783,7 +11783,7 @@ lean_ctor_set_uint64(x_181, sizeof(void*)*1, x_178); lean_ctor_set(x_170, 4, x_181); x_182 = lean_st_ref_set(x_10, x_170); lean_dec(x_10); -x_183 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00__private_Lean_Meta_IndPredBelow_0__Lean_Meta_IndPredBelow_mkBelowMatcher_convertToBelow_spec__8_spec__12___redArg(x_167); +x_183 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00__private_Lean_Meta_IndPredBelow_0__Lean_Meta_IndPredBelow_mkBelowMatcher_convertToBelow_spec__8_spec__12___redArg(x_166); return x_183; } } @@ -11819,7 +11819,7 @@ if (lean_is_exclusive(x_184)) { lean_dec_ref(x_184); x_195 = lean_box(0); } -x_196 = l_Lean_PersistentArray_append___redArg(x_164, x_194); +x_196 = l_Lean_PersistentArray_append___redArg(x_162, x_194); lean_dec_ref(x_194); if (lean_is_scalar(x_195)) { x_197 = lean_alloc_ctor(0, 1, 8); @@ -11840,17 +11840,17 @@ lean_ctor_set(x_198, 7, x_191); lean_ctor_set(x_198, 8, x_192); x_199 = lean_st_ref_set(x_10, x_198); lean_dec(x_10); -x_200 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00__private_Lean_Meta_IndPredBelow_0__Lean_Meta_IndPredBelow_mkBelowMatcher_convertToBelow_spec__8_spec__12___redArg(x_167); +x_200 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00__private_Lean_Meta_IndPredBelow_0__Lean_Meta_IndPredBelow_mkBelowMatcher_convertToBelow_spec__8_spec__12___redArg(x_166); return x_200; } } else { x_78 = x_162; -x_79 = x_164; -x_80 = x_163; -x_81 = lean_box(0); -x_82 = x_166; +x_79 = x_163; +x_80 = lean_box(0); +x_81 = x_166; +x_82 = x_165; x_83 = x_167; goto block_87; } @@ -11858,10 +11858,10 @@ goto block_87; else { x_78 = x_162; -x_79 = x_164; -x_80 = x_163; -x_81 = lean_box(0); -x_82 = x_166; +x_79 = x_163; +x_80 = lean_box(0); +x_81 = x_166; +x_82 = x_165; x_83 = x_167; goto block_87; } @@ -11869,13 +11869,13 @@ goto block_87; block_211: { double x_209; uint8_t x_210; -x_209 = lean_float_sub(x_206, x_204); +x_209 = lean_float_sub(x_203, x_206); x_210 = lean_float_decLt(x_208, x_209); x_162 = x_202; -x_163 = x_204; -x_164 = x_203; -x_165 = lean_box(0); -x_166 = x_206; +x_163 = x_203; +x_164 = lean_box(0); +x_165 = x_206; +x_166 = x_205; x_167 = x_207; x_168 = x_210; goto block_201; @@ -11884,18 +11884,18 @@ goto block_201; { lean_object* x_217; double x_218; double x_219; lean_object* x_220; uint8_t x_221; x_217 = lean_io_get_num_heartbeats(); -x_218 = lean_float_of_nat(x_212); +x_218 = lean_float_of_nat(x_213); x_219 = lean_float_of_nat(x_217); x_220 = l_Lean_withTraceNode___at___00Lean_Meta_IndPredBelow_mkBelow_spec__8___redArg___closed__4; x_221 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00Lean_Meta_IndPredBelow_mkBelow_spec__8_spec__12(x_12, x_220); if (x_221 == 0) { -x_162 = x_221; -x_163 = x_218; -x_164 = x_213; -x_165 = lean_box(0); -x_166 = x_219; -x_167 = x_215; +x_162 = x_212; +x_163 = x_219; +x_164 = lean_box(0); +x_165 = x_218; +x_166 = x_215; +x_167 = x_221; x_168 = x_221; goto block_201; } @@ -11909,12 +11909,12 @@ x_223 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00Lean_Meta_IndPred x_224 = lean_float_of_nat(x_223); x_225 = l_Lean_withTraceNode___at___00Lean_Meta_IndPredBelow_mkBelow_spec__8___redArg___closed__6; x_226 = lean_float_div(x_224, x_225); -x_202 = x_221; -x_203 = x_213; -x_204 = x_218; -x_205 = lean_box(0); -x_206 = x_219; -x_207 = x_215; +x_202 = x_212; +x_203 = x_219; +x_204 = lean_box(0); +x_205 = x_215; +x_206 = x_218; +x_207 = x_221; x_208 = x_226; goto block_211; } @@ -11924,12 +11924,12 @@ lean_object* x_227; lean_object* x_228; double x_229; x_227 = l_Lean_withTraceNode___at___00Lean_Meta_IndPredBelow_mkBelow_spec__8___redArg___closed__5; x_228 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00Lean_Meta_IndPredBelow_mkBelow_spec__8_spec__15(x_12, x_227); x_229 = lean_float_of_nat(x_228); -x_202 = x_221; -x_203 = x_213; -x_204 = x_218; -x_205 = lean_box(0); -x_206 = x_219; -x_207 = x_215; +x_202 = x_212; +x_203 = x_219; +x_204 = lean_box(0); +x_205 = x_215; +x_206 = x_218; +x_207 = x_221; x_208 = x_229; goto block_211; } @@ -11961,9 +11961,9 @@ x_237 = !lean_is_exclusive(x_236); if (x_237 == 0) { lean_ctor_set_tag(x_236, 1); -x_140 = x_235; -x_141 = x_232; -x_142 = x_234; +x_140 = x_232; +x_141 = x_234; +x_142 = x_235; x_143 = x_236; x_144 = lean_box(0); goto block_161; @@ -11976,9 +11976,9 @@ lean_inc(x_238); lean_dec(x_236); x_239 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_239, 0, x_238); -x_140 = x_235; -x_141 = x_232; -x_142 = x_234; +x_140 = x_232; +x_141 = x_234; +x_142 = x_235; x_143 = x_239; x_144 = lean_box(0); goto block_161; @@ -11991,9 +11991,9 @@ x_240 = !lean_is_exclusive(x_236); if (x_240 == 0) { lean_ctor_set_tag(x_236, 0); -x_140 = x_235; -x_141 = x_232; -x_142 = x_234; +x_140 = x_232; +x_141 = x_234; +x_142 = x_235; x_143 = x_236; x_144 = lean_box(0); goto block_161; @@ -12006,9 +12006,9 @@ lean_inc(x_241); lean_dec(x_236); x_242 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_242, 0, x_241); -x_140 = x_235; -x_141 = x_232; -x_142 = x_234; +x_140 = x_232; +x_141 = x_234; +x_142 = x_235; x_143 = x_242; x_144 = lean_box(0); goto block_161; @@ -12032,8 +12032,8 @@ x_245 = !lean_is_exclusive(x_244); if (x_245 == 0) { lean_ctor_set_tag(x_244, 1); -x_212 = x_243; -x_213 = x_232; +x_212 = x_232; +x_213 = x_243; x_214 = x_234; x_215 = x_244; x_216 = lean_box(0); @@ -12047,8 +12047,8 @@ lean_inc(x_246); lean_dec(x_244); x_247 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_247, 0, x_246); -x_212 = x_243; -x_213 = x_232; +x_212 = x_232; +x_213 = x_243; x_214 = x_234; x_215 = x_247; x_216 = lean_box(0); @@ -12062,8 +12062,8 @@ x_248 = !lean_is_exclusive(x_244); if (x_248 == 0) { lean_ctor_set_tag(x_244, 0); -x_212 = x_243; -x_213 = x_232; +x_212 = x_232; +x_213 = x_243; x_214 = x_234; x_215 = x_244; x_216 = lean_box(0); @@ -12077,8 +12077,8 @@ lean_inc(x_249); lean_dec(x_244); x_250 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_250, 0, x_249); -x_212 = x_243; -x_213 = x_232; +x_212 = x_232; +x_213 = x_243; x_214 = x_234; x_215 = x_250; x_216 = lean_box(0); @@ -17018,7 +17018,7 @@ goto _start; LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at___00Lean_Meta_IndPredBelow_mkBelowMatcher_spec__12___redArg___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, uint8_t x_11, uint8_t x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { -lean_object* x_18; lean_object* x_19; size_t x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; size_t x_114; lean_object* x_115; lean_object* x_116; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_172; lean_object* x_173; uint8_t x_174; +lean_object* x_18; size_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; size_t x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_172; lean_object* x_173; uint8_t x_174; lean_inc(x_1); x_172 = l_Lean_isTracingEnabledFor___at___00Lean_withTraceNode___at___00Lean_Meta_IndPredBelow_mkBelow_spec__8_spec__10___redArg(x_1, x_15); x_173 = lean_ctor_get(x_172, 0); @@ -17096,7 +17096,7 @@ return x_184; { size_t x_26; lean_object* x_27; x_26 = lean_array_size(x_2); -x_27 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_IndPredBelow_mkBelowMatcher_spec__9(x_2, x_26, x_20, x_3, x_21, x_22, x_23, x_24); +x_27 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_IndPredBelow_mkBelowMatcher_spec__9(x_2, x_26, x_19, x_3, x_21, x_22, x_23, x_24); lean_dec_ref(x_2); if (lean_obj_tag(x_27) == 0) { @@ -17117,8 +17117,8 @@ lean_inc(x_30); lean_dec_ref(x_29); x_31 = 1; x_32 = 1; -x_33 = l_Lean_Meta_mkLambdaFVars(x_19, x_30, x_5, x_31, x_5, x_31, x_32, x_21, x_22, x_23, x_24); -lean_dec_ref(x_19); +x_33 = l_Lean_Meta_mkLambdaFVars(x_20, x_30, x_5, x_31, x_5, x_31, x_32, x_21, x_22, x_23, x_24); +lean_dec_ref(x_20); if (lean_obj_tag(x_33) == 0) { uint8_t x_34; @@ -17455,7 +17455,7 @@ lean_dec(x_24); lean_dec_ref(x_23); lean_dec(x_22); lean_dec_ref(x_21); -lean_dec_ref(x_19); +lean_dec_ref(x_20); lean_dec_ref(x_7); lean_dec(x_6); lean_dec(x_1); @@ -17469,7 +17469,7 @@ lean_dec(x_24); lean_dec_ref(x_23); lean_dec(x_22); lean_dec_ref(x_21); -lean_dec_ref(x_19); +lean_dec_ref(x_20); lean_dec_ref(x_18); lean_dec_ref(x_7); lean_dec(x_6); @@ -17496,7 +17496,7 @@ return x_107; { lean_object* x_117; uint8_t x_118; lean_inc(x_1); -x_117 = l_Lean_isTracingEnabledFor___at___00Lean_withTraceNode___at___00Lean_Meta_IndPredBelow_mkBelow_spec__8_spec__10___redArg(x_1, x_113); +x_117 = l_Lean_isTracingEnabledFor___at___00Lean_withTraceNode___at___00Lean_Meta_IndPredBelow_mkBelow_spec__8_spec__10___redArg(x_1, x_112); x_118 = !lean_is_exclusive(x_117); if (x_118 == 0) { @@ -17509,12 +17509,12 @@ if (x_120 == 0) lean_free_object(x_117); lean_dec(x_8); x_18 = x_116; -x_19 = x_112; -x_20 = x_114; +x_19 = x_113; +x_20 = x_115; x_21 = x_110; -x_22 = x_111; -x_23 = x_113; -x_24 = x_109; +x_22 = x_109; +x_23 = x_112; +x_24 = x_114; x_25 = lean_box(0); goto block_108; } @@ -17533,8 +17533,8 @@ x_125 = l_Array_mapFinIdxM_map___at___00Lean_Meta_IndPredBelow_mkBelowMatcher_sp x_126 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_126, 0, x_124); lean_ctor_set(x_126, 1, x_125); -lean_inc_ref(x_112); -x_127 = lean_array_to_list(x_112); +lean_inc_ref(x_115); +x_127 = lean_array_to_list(x_115); x_128 = lean_box(0); x_129 = l_List_mapTR_loop___at___00__private_Lean_Meta_IndPredBelow_0__Lean_Meta_IndPredBelow_mkBelowMatcher_convertToBelow_spec__7(x_127, x_128); x_130 = l_Lean_MessageData_ofList(x_129); @@ -17542,17 +17542,17 @@ x_131 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_131, 0, x_126); lean_ctor_set(x_131, 1, x_130); lean_inc(x_1); -x_132 = l_Lean_addTrace___at___00Lean_Meta_IndPredBelow_mkBelowMatcher_spec__10(x_1, x_131, x_110, x_111, x_113, x_109); +x_132 = l_Lean_addTrace___at___00Lean_Meta_IndPredBelow_mkBelowMatcher_spec__10(x_1, x_131, x_110, x_109, x_112, x_114); if (lean_obj_tag(x_132) == 0) { lean_dec_ref(x_132); x_18 = x_116; -x_19 = x_112; -x_20 = x_114; +x_19 = x_113; +x_20 = x_115; x_21 = x_110; -x_22 = x_111; -x_23 = x_113; -x_24 = x_109; +x_22 = x_109; +x_23 = x_112; +x_24 = x_114; x_25 = lean_box(0); goto block_108; } @@ -17560,9 +17560,9 @@ else { uint8_t x_133; lean_dec_ref(x_116); -lean_dec_ref(x_113); +lean_dec_ref(x_115); +lean_dec(x_114); lean_dec_ref(x_112); -lean_dec(x_111); lean_dec_ref(x_110); lean_dec(x_109); lean_dec_ref(x_7); @@ -17601,12 +17601,12 @@ if (x_137 == 0) { lean_dec(x_8); x_18 = x_116; -x_19 = x_112; -x_20 = x_114; +x_19 = x_113; +x_20 = x_115; x_21 = x_110; -x_22 = x_111; -x_23 = x_113; -x_24 = x_109; +x_22 = x_109; +x_23 = x_112; +x_24 = x_114; x_25 = lean_box(0); goto block_108; } @@ -17625,8 +17625,8 @@ x_143 = l_Array_mapFinIdxM_map___at___00Lean_Meta_IndPredBelow_mkBelowMatcher_sp x_144 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_144, 0, x_142); lean_ctor_set(x_144, 1, x_143); -lean_inc_ref(x_112); -x_145 = lean_array_to_list(x_112); +lean_inc_ref(x_115); +x_145 = lean_array_to_list(x_115); x_146 = lean_box(0); x_147 = l_List_mapTR_loop___at___00__private_Lean_Meta_IndPredBelow_0__Lean_Meta_IndPredBelow_mkBelowMatcher_convertToBelow_spec__7(x_145, x_146); x_148 = l_Lean_MessageData_ofList(x_147); @@ -17634,17 +17634,17 @@ x_149 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_149, 0, x_144); lean_ctor_set(x_149, 1, x_148); lean_inc(x_1); -x_150 = l_Lean_addTrace___at___00Lean_Meta_IndPredBelow_mkBelowMatcher_spec__10(x_1, x_149, x_110, x_111, x_113, x_109); +x_150 = l_Lean_addTrace___at___00Lean_Meta_IndPredBelow_mkBelowMatcher_spec__10(x_1, x_149, x_110, x_109, x_112, x_114); if (lean_obj_tag(x_150) == 0) { lean_dec_ref(x_150); x_18 = x_116; -x_19 = x_112; -x_20 = x_114; +x_19 = x_113; +x_20 = x_115; x_21 = x_110; -x_22 = x_111; -x_23 = x_113; -x_24 = x_109; +x_22 = x_109; +x_23 = x_112; +x_24 = x_114; x_25 = lean_box(0); goto block_108; } @@ -17652,9 +17652,9 @@ else { lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_dec_ref(x_116); -lean_dec_ref(x_113); +lean_dec_ref(x_115); +lean_dec(x_114); lean_dec_ref(x_112); -lean_dec(x_111); lean_dec_ref(x_110); lean_dec(x_109); lean_dec_ref(x_7); @@ -17696,13 +17696,13 @@ lean_inc_ref(x_7); x_165 = l_Lean_Expr_beta(x_7, x_164); if (x_11 == 0) { -x_109 = x_158; +x_109 = x_156; x_110 = x_155; -x_111 = x_156; -x_112 = x_163; -x_113 = x_157; -x_114 = x_162; -x_115 = lean_box(0); +x_111 = lean_box(0); +x_112 = x_157; +x_113 = x_162; +x_114 = x_158; +x_115 = x_163; x_116 = x_165; goto block_154; } @@ -17719,38 +17719,38 @@ if (x_168 == 0) lean_object* x_169; lean_object* x_170; x_169 = l_Array_mapFinIdxM_map___at___00Lean_Meta_IndPredBelow_mkBelowMatcher_spec__12___redArg___lam__0___closed__15; x_170 = l_Lean_Expr_betaRev(x_165, x_169, x_5, x_5); -x_109 = x_158; +x_109 = x_156; x_110 = x_155; -x_111 = x_156; -x_112 = x_163; -x_113 = x_157; -x_114 = x_162; -x_115 = lean_box(0); +x_111 = lean_box(0); +x_112 = x_157; +x_113 = x_162; +x_114 = x_158; +x_115 = x_163; x_116 = x_170; goto block_154; } else { -x_109 = x_158; +x_109 = x_156; x_110 = x_155; -x_111 = x_156; -x_112 = x_163; -x_113 = x_157; -x_114 = x_162; -x_115 = lean_box(0); +x_111 = lean_box(0); +x_112 = x_157; +x_113 = x_162; +x_114 = x_158; +x_115 = x_163; x_116 = x_165; goto block_154; } } else { -x_109 = x_158; +x_109 = x_156; x_110 = x_155; -x_111 = x_156; -x_112 = x_163; -x_113 = x_157; -x_114 = x_162; -x_115 = lean_box(0); +x_111 = lean_box(0); +x_112 = x_157; +x_113 = x_162; +x_114 = x_158; +x_115 = x_163; x_116 = x_165; goto block_154; } diff --git a/stage0/stdlib/Lean/Meta/Instances.c b/stage0/stdlib/Lean/Meta/Instances.c index 32c788e4fd03..2f282fac8f1a 100644 --- a/stage0/stdlib/Lean/Meta/Instances.c +++ b/stage0/stdlib/Lean/Meta/Instances.c @@ -7456,14 +7456,14 @@ goto block_58; { lean_object* x_20; lean_object* x_21; lean_object* x_22; x_20 = l_Lean_instInhabitedExpr; -x_21 = lean_array_get_borrowed(x_20, x_1, x_16); -lean_dec(x_16); +x_21 = lean_array_get_borrowed(x_20, x_1, x_17); +lean_dec(x_17); lean_inc(x_14); lean_inc_ref(x_12); -lean_inc(x_15); +lean_inc(x_13); lean_inc_ref(x_18); lean_inc(x_21); -x_22 = lean_infer_type(x_21, x_18, x_15, x_12, x_14); +x_22 = lean_infer_type(x_21, x_18, x_13, x_12, x_14); if (lean_obj_tag(x_22) == 0) { lean_object* x_23; lean_object* x_24; @@ -7472,21 +7472,21 @@ lean_inc(x_23); lean_dec_ref(x_22); lean_inc(x_14); lean_inc_ref(x_12); -lean_inc(x_15); +lean_inc(x_13); lean_inc_ref(x_18); -x_24 = l___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder_assignMVarsIn(x_1, x_2, x_23, x_18, x_15, x_12, x_14); +x_24 = l___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder_assignMVarsIn(x_1, x_2, x_23, x_18, x_13, x_12, x_14); if (lean_obj_tag(x_24) == 0) { lean_object* x_25; lean_dec_ref(x_24); lean_inc(x_21); -x_25 = l___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder_assignMVarsIn(x_1, x_2, x_21, x_18, x_15, x_12, x_14); +x_25 = l___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder_assignMVarsIn(x_1, x_2, x_21, x_18, x_13, x_12, x_14); if (lean_obj_tag(x_25) == 0) { lean_object* x_26; lean_dec_ref(x_25); x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_13); +lean_ctor_set(x_26, 0, x_15); lean_ctor_set(x_26, 1, x_19); x_6 = x_26; goto _start; @@ -7495,7 +7495,7 @@ else { uint8_t x_28; lean_dec_ref(x_19); -lean_dec_ref(x_13); +lean_dec_ref(x_15); lean_dec(x_10); lean_dec_ref(x_9); lean_dec(x_8); @@ -7524,9 +7524,9 @@ else uint8_t x_31; lean_dec_ref(x_19); lean_dec_ref(x_18); -lean_dec(x_15); +lean_dec_ref(x_15); lean_dec(x_14); -lean_dec_ref(x_13); +lean_dec(x_13); lean_dec_ref(x_12); lean_dec(x_10); lean_dec_ref(x_9); @@ -7556,9 +7556,9 @@ else uint8_t x_34; lean_dec_ref(x_19); lean_dec_ref(x_18); -lean_dec(x_15); +lean_dec_ref(x_15); lean_dec(x_14); -lean_dec_ref(x_13); +lean_dec(x_13); lean_dec_ref(x_12); lean_dec(x_10); lean_dec_ref(x_9); @@ -7596,11 +7596,11 @@ if (x_50 == 0) { lean_dec_ref(x_39); x_12 = x_43; -x_13 = x_46; +x_13 = x_42; x_14 = x_44; -x_15 = x_42; -x_16 = x_40; -x_17 = lean_box(0); +x_15 = x_46; +x_16 = lean_box(0); +x_17 = x_40; x_18 = x_41; x_19 = x_49; goto block_37; @@ -7615,11 +7615,11 @@ if (x_50 == 0) { lean_dec_ref(x_39); x_12 = x_43; -x_13 = x_46; +x_13 = x_42; x_14 = x_44; -x_15 = x_42; -x_16 = x_40; -x_17 = lean_box(0); +x_15 = x_46; +x_16 = lean_box(0); +x_17 = x_40; x_18 = x_41; x_19 = x_49; goto block_37; @@ -7632,11 +7632,11 @@ x_53 = lean_usize_of_nat(x_48); x_54 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder_spec__0(x_40, x_39, x_52, x_53, x_49); lean_dec_ref(x_39); x_12 = x_43; -x_13 = x_46; +x_13 = x_42; x_14 = x_44; -x_15 = x_42; -x_16 = x_40; -x_17 = lean_box(0); +x_15 = x_46; +x_16 = lean_box(0); +x_17 = x_40; x_18 = x_41; x_19 = x_54; goto block_37; @@ -7650,11 +7650,11 @@ x_56 = lean_usize_of_nat(x_48); x_57 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder_spec__0(x_40, x_39, x_55, x_56, x_49); lean_dec_ref(x_39); x_12 = x_43; -x_13 = x_46; +x_13 = x_42; x_14 = x_44; -x_15 = x_42; -x_16 = x_40; -x_17 = lean_box(0); +x_15 = x_46; +x_16 = lean_box(0); +x_17 = x_40; x_18 = x_41; x_19 = x_57; goto block_37; @@ -8196,29 +8196,29 @@ lean_object* x_20; lean_object* x_21; lean_object* x_22; x_20 = l_Lean_instInhabitedExpr; x_21 = lean_array_get_borrowed(x_20, x_1, x_12); lean_dec(x_12); -lean_inc(x_14); -lean_inc_ref(x_18); -lean_inc(x_13); +lean_inc(x_18); lean_inc_ref(x_15); +lean_inc(x_17); +lean_inc_ref(x_14); lean_inc(x_21); -x_22 = lean_infer_type(x_21, x_15, x_13, x_18, x_14); +x_22 = lean_infer_type(x_21, x_14, x_17, x_15, x_18); if (lean_obj_tag(x_22) == 0) { lean_object* x_23; lean_object* x_24; x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); lean_dec_ref(x_22); -lean_inc(x_14); -lean_inc_ref(x_18); -lean_inc(x_13); +lean_inc(x_18); lean_inc_ref(x_15); -x_24 = l___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder_assignMVarsIn(x_1, x_2, x_23, x_15, x_13, x_18, x_14); +lean_inc(x_17); +lean_inc_ref(x_14); +x_24 = l___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder_assignMVarsIn(x_1, x_2, x_23, x_14, x_17, x_15, x_18); if (lean_obj_tag(x_24) == 0) { lean_object* x_25; lean_dec_ref(x_24); lean_inc(x_21); -x_25 = l___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder_assignMVarsIn(x_1, x_2, x_21, x_15, x_13, x_18, x_14); +x_25 = l___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder_assignMVarsIn(x_1, x_2, x_21, x_14, x_17, x_15, x_18); if (lean_obj_tag(x_25) == 0) { lean_object* x_26; lean_object* x_27; @@ -8261,11 +8261,11 @@ else { uint8_t x_31; lean_dec_ref(x_19); -lean_dec_ref(x_18); +lean_dec(x_18); +lean_dec(x_17); lean_dec_ref(x_16); lean_dec_ref(x_15); -lean_dec(x_14); -lean_dec(x_13); +lean_dec_ref(x_14); lean_dec(x_10); lean_dec_ref(x_9); lean_dec(x_8); @@ -8293,11 +8293,11 @@ else { uint8_t x_34; lean_dec_ref(x_19); -lean_dec_ref(x_18); +lean_dec(x_18); +lean_dec(x_17); lean_dec_ref(x_16); lean_dec_ref(x_15); -lean_dec(x_14); -lean_dec(x_13); +lean_dec_ref(x_14); lean_dec(x_10); lean_dec_ref(x_9); lean_dec(x_8); @@ -8334,12 +8334,12 @@ if (x_50 == 0) { lean_dec_ref(x_39); x_12 = x_40; -x_13 = x_42; -x_14 = x_44; -x_15 = x_41; +x_13 = lean_box(0); +x_14 = x_41; +x_15 = x_43; x_16 = x_46; -x_17 = lean_box(0); -x_18 = x_43; +x_17 = x_42; +x_18 = x_44; x_19 = x_49; goto block_37; } @@ -8353,12 +8353,12 @@ if (x_50 == 0) { lean_dec_ref(x_39); x_12 = x_40; -x_13 = x_42; -x_14 = x_44; -x_15 = x_41; +x_13 = lean_box(0); +x_14 = x_41; +x_15 = x_43; x_16 = x_46; -x_17 = lean_box(0); -x_18 = x_43; +x_17 = x_42; +x_18 = x_44; x_19 = x_49; goto block_37; } @@ -8370,12 +8370,12 @@ x_53 = lean_usize_of_nat(x_48); x_54 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder_spec__0(x_40, x_39, x_52, x_53, x_49); lean_dec_ref(x_39); x_12 = x_40; -x_13 = x_42; -x_14 = x_44; -x_15 = x_41; +x_13 = lean_box(0); +x_14 = x_41; +x_15 = x_43; x_16 = x_46; -x_17 = lean_box(0); -x_18 = x_43; +x_17 = x_42; +x_18 = x_44; x_19 = x_54; goto block_37; } @@ -8388,12 +8388,12 @@ x_56 = lean_usize_of_nat(x_48); x_57 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder_spec__0(x_40, x_39, x_55, x_56, x_49); lean_dec_ref(x_39); x_12 = x_40; -x_13 = x_42; -x_14 = x_44; -x_15 = x_41; +x_13 = lean_box(0); +x_14 = x_41; +x_15 = x_43; x_16 = x_46; -x_17 = lean_box(0); -x_18 = x_43; +x_17 = x_42; +x_18 = x_44; x_19 = x_57; goto block_37; } diff --git a/stage0/stdlib/Lean/Meta/Match/Rewrite.c b/stage0/stdlib/Lean/Meta/Match/Rewrite.c new file mode 100644 index 000000000000..daee877a294c --- /dev/null +++ b/stage0/stdlib/Lean/Meta/Match/Rewrite.c @@ -0,0 +1,8590 @@ +// Lean compiler output +// Module: Lean.Meta.Match.Rewrite +// Imports: public import Lean.Meta.Tactic.Simp.Types import Lean.Meta.Match.MatcherApp.Transform import Lean.Meta.Tactic.Assumption import Lean.Meta.Tactic.Refl import Lean.Meta.Tactic.Simp.Rewrite +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00Lean_Meta_rwMatcher_spec__5___redArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_rwMatcher___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_reduceRecMatcher_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_rwMatcher___closed__12; +LEAN_EXPORT lean_object* l_Lean_Meta_rwMatcher___lam__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__5; +static lean_object* l_Lean_Meta_rwIfWith___closed__5; +static lean_object* l_Lean_Meta_rwMatcher___closed__21; +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00Lean_Meta_rwMatcher_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_mk_empty_array_with_capacity(lean_object*); +lean_object* l_Lean_mkAppN(lean_object*, lean_object*); +size_t lean_usize_shift_right(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_Meta_rwMatcher_spec__6___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static size_t l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4___redArg___closed__1; +static lean_object* l_Lean_Meta_rwMatcher___closed__4; +static lean_object* l_Lean_Meta_rwIfWith___closed__22; +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00Lean_Meta_rwMatcher_spec__5___redArg___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_indentD(lean_object*); +double lean_float_div(double, double); +LEAN_EXPORT lean_object* l_Lean_Meta_rwMatcher___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Array_isEmpty___redArg(lean_object*); +uint8_t l_Lean_Exception_isInterrupt(lean_object*); +static lean_object* l_Lean_Meta_rwMatcher___lam__1___closed__26; +lean_object* l_Lean_Meta_isExprDefEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +size_t lean_uint64_to_usize(uint64_t); +static lean_object* l_Lean_Meta_rwIfWith___closed__7; +lean_object* l_Lean_Meta_forallMetaTelescope(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__15___redArg(lean_object*); +uint8_t l_Lean_Expr_isAppOfArity(lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_instBEqMVarId_beq(lean_object*, lean_object*); +uint8_t l_Lean_Expr_isApp(lean_object*); +static lean_object* l_Lean_Meta_rwIfWith___closed__3; +static lean_object* l_Lean_Meta_rwMatcher___closed__13; +lean_object* l_Lean_Expr_sort___override(lean_object*); +lean_object* l_Lean_MessageData_ofList(lean_object*); +lean_object* lean_array_push(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_rwIfWith___closed__15; +uint8_t lean_usize_dec_eq(size_t, size_t); +lean_object* l_Lean_KVMap_find(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_rwMatcher___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_rwIfWith___closed__17; +LEAN_EXPORT lean_object* l_List_mapTR_loop___at___00Lean_Meta_rwMatcher_spec__7(lean_object*, lean_object*); +lean_object* l_Lean_replaceRef(lean_object*, lean_object*); +lean_object* lean_mk_array(lean_object*, lean_object*); +static double l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__2; +static lean_object* l_Lean_Meta_rwIfWith___closed__6; +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__15___redArg___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_rwIfWith___closed__10; +LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_Meta_rwMatcher_spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_rwMatcher___lam__1___closed__20; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_float_decLt(double, double); +static lean_object* l_Lean_Meta_rwIfWith___closed__13; +static lean_object* l_Lean_Meta_rwMatcher___lam__1___closed__6; +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00Lean_Meta_rwMatcher_spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_Meta_rwMatcher_spec__6___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MVarId_refl(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_rwMatcher___lam__1___closed__7; +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_rwMatcher___lam__1___closed__14; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at___00Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4_spec__18___redArg(lean_object*, lean_object*, lean_object*); +lean_object* lean_array_fget_borrowed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12___redArg(lean_object*); +uint8_t l_Lean_Expr_isAppOf(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_rwMatcher___lam__1___closed__5; +static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__4; +lean_object* lean_io_get_num_heartbeats(); +lean_object* l_Lean_Expr_cleanupAnnotations(lean_object*); +extern lean_object* l_Lean_trace_profiler_useHeartbeats; +lean_object* l_Lean_stringToMessageData(lean_object*); +static lean_object* l_Lean_Meta_rwIfWith___closed__9; +static lean_object* l_Lean_Meta_rwIfWith___closed__23; +lean_object* l_Lean_Exception_toMessageData(lean_object*); +static lean_object* l_Lean_Meta_rwMatcher___closed__11; +static lean_object* l_Lean_Meta_rwMatcher___lam__1___closed__18; +static lean_object* l_Lean_Meta_rwMatcher___lam__0___closed__0; +lean_object* l_Lean_Expr_appArg_x21(lean_object*); +static lean_object* l_Lean_Meta_rwIfWith___closed__12; +uint8_t l_Lean_Expr_hasMVar(lean_object*); +lean_object* l_Lean_Meta_instantiateMVarsIfMVarApp___redArg(lean_object*, lean_object*); +lean_object* l_Nat_reprFast(lean_object*); +lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_rwMatcher___lam__1___closed__19; +static lean_object* l_Lean_Meta_rwMatcher___lam__1___closed__22; +static lean_object* l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__6; +size_t lean_usize_of_nat(lean_object*); +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4___redArg(lean_object*, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_isMatcherApp___at___00Lean_Meta_rwMatcher_spec__1___redArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at___00Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4_spec__18___redArg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00Lean_Meta_rwMatcher_spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_rwIfWith___closed__18; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_st_ref_take(lean_object*); +static lean_object* l_Lean_Meta_rwMatcher___closed__19; +uint8_t lean_expr_eqv(lean_object*, lean_object*); +static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__2; +static lean_object* l_Lean_Meta_rwIfWith___closed__24; +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at___00Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4_spec__18(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_rwMatcher___lam__1___closed__15; +lean_object* l_Lean_MVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_rwMatcher___closed__3; +uint8_t l_Lean_Expr_isEq(lean_object*); +LEAN_EXPORT uint8_t l_Lean_Option_get___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__13(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_rwMatcher___lam__1___closed__2; +static lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12___redArg___closed__0; +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12___redArg___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_rwIfWith___closed__25; +static lean_object* l_Lean_Meta_rwIfWith___closed__26; +static lean_object* l_Lean_Meta_rwIfWith___closed__11; +lean_object* l_Lean_MessageData_ofFormat(lean_object*); +static lean_object* l_Lean_Meta_rwIfWith___closed__20; +static lean_object* l_Lean_Meta_rwMatcher___lam__1___closed__0; +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4(lean_object*, lean_object*, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0___redArg___boxed(lean_object*, lean_object*); +double lean_float_of_nat(lean_object*); +lean_object* lean_st_ref_get(lean_object*); +static lean_object* l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__3; +static lean_object* l_Lean_Meta_rwMatcher___lam__1___closed__3; +static lean_object* l_Lean_Meta_rwMatcher___lam__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_addMessageContextFull___at___00Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3_spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_to_list(lean_object*); +lean_object* l_Lean_Meta_mkEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__0; +lean_object* l_Lean_mkConst(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_rwMatcher___closed__9; +LEAN_EXPORT lean_object* l_Lean_addMessageContextFull___at___00Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3_spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_io_mono_nanos_now(); +lean_object* l_Lean_Expr_constName_x21(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_rwIfWith___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_rwMatcher___closed__6; +static lean_object* l_Lean_Meta_rwMatcher___lam__1___closed__23; +LEAN_EXPORT lean_object* l___private_Init_While_0__Lean_Loop_forIn_loop___at___00Lean_Meta_rwMatcher_spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_PersistentArray_push___redArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_isMatcherApp___at___00Lean_Meta_rwMatcher_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_rwMatcher___closed__16; +static lean_object* l_Lean_Meta_rwMatcher___closed__2; +LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_Meta_rwMatcher_spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_rwIfWith___closed__8; +lean_object* l_Lean_PersistentArray_append___redArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8(lean_object*, uint8_t, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_rwMatcher___closed__0; +lean_object* lean_get_congr_match_equations_for(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_rwMatcher___lam__0___closed__3; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__14_spec__16___boxed(lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_trace_profiler_threshold; +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00Lean_Meta_rwMatcher_spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_rwMatcher___lam__0(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_mkNot(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_While_0__Lean_Loop_forIn_loop___at___00Lean_Meta_rwMatcher_spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0___redArg(lean_object*, lean_object*); +lean_object* l_Lean_MVarId_assumption(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12___redArg___closed__2; +static lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12___redArg___closed__1; +lean_object* l_Lean_Expr_appFn_x21(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_rwMatcher___lam__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Option_get___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__16(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_rwIfWith___closed__0; +static lean_object* l_Lean_Meta_rwMatcher___lam__1___closed__9; +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_rwMatcher___lam__0___closed__2; +lean_object* lean_usize_to_nat(size_t); +lean_object* l_Lean_MessageData_ofConstName(lean_object*, uint8_t); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MessageData_ofExpr(lean_object*); +static lean_object* l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___closed__2; +static lean_object* l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_rwMatcher(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_isMatcherApp___at___00Lean_Meta_rwMatcher_spec__1___redArg___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_getAppNumArgs(lean_object*); +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0___redArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_rwMatcher_spec__4(size_t, size_t, lean_object*); +static lean_object* l_Lean_Meta_rwMatcher___closed__14; +LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__4; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_constLevels_x21(lean_object*); +static lean_object* l_Lean_Meta_rwMatcher___lam__1___closed__13; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4___redArg___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_appFnCleanup___redArg(lean_object*); +static lean_object* l_Lean_Meta_rwIfWith___closed__2; +uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +uint64_t l_Lean_instHashableMVarId_hash(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__14_spec__16(size_t, size_t, lean_object*); +lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); +uint8_t l_Lean_Meta_isMatcherAppCore(lean_object*, lean_object*); +lean_object* l_Lean_PersistentArray_toArray___redArg(lean_object*); +lean_object* l_Lean_indentExpr(lean_object*); +uint8_t l_Lean_Expr_isHEq(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Meta_rwMatcher_spec__9(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_rwIfWith___closed__1; +uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_SavedState_restore___redArg(lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Meta_Simp_isEqnThmHypothesis(lean_object*); +static lean_object* l_Lean_Meta_rwMatcher___lam__1___closed__12; +static lean_object* l_Lean_Meta_rwMatcher___lam__1___closed__21; +static lean_object* l_Lean_Meta_rwMatcher___closed__1; +lean_object* l_Lean_Meta_saveState___redArg(lean_object*, lean_object*); +static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__3; +static lean_object* l_Lean_Meta_rwMatcher___lam__1___closed__11; +static lean_object* l_Lean_Meta_rwMatcher___lam__1___closed__17; +static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__1; +lean_object* l_List_reverse___redArg(lean_object*); +lean_object* lean_nat_sub(lean_object*, lean_object*); +lean_object* l_Lean_Expr_getAppFn(lean_object*); +static lean_object* l_Lean_Meta_rwMatcher___closed__5; +static lean_object* l_Lean_Meta_rwMatcher___closed__15; +LEAN_EXPORT lean_object* l_Lean_Meta_isMatcherApp___at___00Lean_Meta_rwMatcher_spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static size_t l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4___redArg___closed__0; +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00Lean_Meta_rwMatcher_spec__2___redArg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at___00Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4_spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_rwIfWith___closed__19; +LEAN_EXPORT lean_object* l_Lean_Option_get___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__16___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_rwIfWith___closed__21; +static lean_object* l_Lean_Meta_rwMatcher___lam__1___closed__24; +size_t lean_usize_sub(size_t, size_t); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_rwMatcher___lam__1___closed__10; +static lean_object* l_Lean_Meta_rwIfWith___closed__16; +LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_instantiateMVarsCore(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0(lean_object*, lean_object*, lean_object*); +size_t lean_usize_add(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_Meta_rwMatcher___lam__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_mkEqOfHEq(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_uget(lean_object*, size_t); +static lean_object* l_Lean_Meta_rwMatcher___closed__7; +size_t lean_array_size(lean_object*); +static double l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__5; +static lean_object* l_Lean_Meta_rwMatcher___closed__18; +extern lean_object* l_Lean_trace_profiler; +lean_object* lean_st_ref_set(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Name_mkStr1(lean_object*); +size_t lean_usize_shift_left(size_t, size_t); +static lean_object* l_Lean_Meta_rwMatcher___lam__1___closed__4; +lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Option_get___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__13___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_rwMatcher_spec__4___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_headBeta(lean_object*); +lean_object* lean_array_get_size(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_rwIfWith(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_rwMatcher___closed__17; +lean_object* lean_infer_type(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_object* lean_nat_add(lean_object*, lean_object*); +lean_object* l_Lean_exceptEmoji___redArg(lean_object*); +static lean_object* l_Lean_Meta_rwIfWith___closed__4; +uint8_t l_Lean_Exception_isRuntime(lean_object*); +lean_object* l_Lean_Expr_beta(lean_object*, lean_object*); +static double l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___closed__0; +static lean_object* l_Lean_Meta_rwMatcher___lam__0___closed__1; +lean_object* l_Lean_Expr_mvarId_x21(lean_object*); +static lean_object* l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___closed__1; +static lean_object* l_Lean_Meta_rwMatcher___lam__1___closed__16; +static lean_object* l_Lean_Meta_rwMatcher___lam__1___closed__8; +static lean_object* l_Lean_Meta_rwMatcher___lam__1___closed__25; +lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +static lean_object* l_Lean_Meta_rwMatcher___closed__8; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_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_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00Lean_Meta_rwMatcher_spec__2___redArg(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_rwMatcher___closed__10; +static lean_object* l_Lean_Meta_rwMatcher___closed__20; +static lean_object* l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__0; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Meta_rwMatcher_spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MVarId_hrefl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +size_t lean_usize_land(size_t, size_t); +static lean_object* l_Lean_Meta_rwIfWith___closed__27; +double lean_float_sub(double, double); +static lean_object* l_Lean_Meta_rwIfWith___closed__14; +static lean_object* _init_l_Lean_Meta_rwIfWith___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("cond", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwIfWith___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_rwIfWith___closed__0; +x_2 = l_Lean_Name_mkStr1(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_rwIfWith___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("dite", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwIfWith___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_rwIfWith___closed__2; +x_2 = l_Lean_Name_mkStr1(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_rwIfWith___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ite", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwIfWith___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_rwIfWith___closed__4; +x_2 = l_Lean_Name_mkStr1(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_rwIfWith___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("if_neg", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwIfWith___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_rwIfWith___closed__6; +x_2 = l_Lean_Name_mkStr1(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_rwIfWith___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(6u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_rwIfWith___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("if_pos", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwIfWith___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_rwIfWith___closed__9; +x_2 = l_Lean_Name_mkStr1(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_rwIfWith___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_rwIfWith___closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("dif_neg", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwIfWith___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_rwIfWith___closed__12; +x_2 = l_Lean_Name_mkStr1(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_rwIfWith___closed__14() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("dif_pos", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwIfWith___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_rwIfWith___closed__14; +x_2 = l_Lean_Name_mkStr1(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_rwIfWith___closed__16() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Bool", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwIfWith___closed__17() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("true", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwIfWith___closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_rwIfWith___closed__17; +x_2 = l_Lean_Meta_rwIfWith___closed__16; +x_3 = l_Lean_Name_mkStr2(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_rwIfWith___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_rwIfWith___closed__18; +x_3 = l_Lean_mkConst(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_rwIfWith___closed__20() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("false", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwIfWith___closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_rwIfWith___closed__20; +x_2 = l_Lean_Meta_rwIfWith___closed__16; +x_3 = l_Lean_Name_mkStr2(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_rwIfWith___closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_rwIfWith___closed__21; +x_3 = l_Lean_mkConst(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_rwIfWith___closed__23() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("cond_neg", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwIfWith___closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_rwIfWith___closed__23; +x_2 = l_Lean_Meta_rwIfWith___closed__16; +x_3 = l_Lean_Name_mkStr2(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_rwIfWith___closed__25() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(5u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_rwIfWith___closed__26() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("cond_pos", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwIfWith___closed__27() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_rwIfWith___closed__26; +x_2 = l_Lean_Meta_rwIfWith___closed__16; +x_3 = l_Lean_Name_mkStr2(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_rwIfWith(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_8; lean_object* x_14; +lean_inc_ref(x_2); +x_14 = l_Lean_Meta_instantiateMVarsIfMVarApp___redArg(x_2, x_4); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +lean_dec_ref(x_14); +x_16 = l_Lean_Expr_cleanupAnnotations(x_15); +x_17 = l_Lean_Expr_isApp(x_16); +if (x_17 == 0) +{ +lean_dec_ref(x_16); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_1); +x_8 = lean_box(0); +goto block_13; +} +else +{ +lean_object* x_18; uint8_t x_19; +lean_inc_ref(x_16); +x_18 = l_Lean_Expr_appFnCleanup___redArg(x_16); +x_19 = l_Lean_Expr_isApp(x_18); +if (x_19 == 0) +{ +lean_dec_ref(x_18); +lean_dec_ref(x_16); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_1); +x_8 = lean_box(0); +goto block_13; +} +else +{ +lean_object* x_20; uint8_t x_21; +lean_inc_ref(x_18); +x_20 = l_Lean_Expr_appFnCleanup___redArg(x_18); +x_21 = l_Lean_Expr_isApp(x_20); +if (x_21 == 0) +{ +lean_dec_ref(x_20); +lean_dec_ref(x_18); +lean_dec_ref(x_16); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_1); +x_8 = lean_box(0); +goto block_13; +} +else +{ +lean_object* x_22; uint8_t x_23; +lean_inc_ref(x_20); +x_22 = l_Lean_Expr_appFnCleanup___redArg(x_20); +x_23 = l_Lean_Expr_isApp(x_22); +if (x_23 == 0) +{ +lean_dec_ref(x_22); +lean_dec_ref(x_20); +lean_dec_ref(x_18); +lean_dec_ref(x_16); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_1); +x_8 = lean_box(0); +goto block_13; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_24 = lean_ctor_get(x_16, 1); +lean_inc_ref(x_24); +lean_dec_ref(x_16); +x_25 = lean_ctor_get(x_18, 1); +lean_inc_ref(x_25); +lean_dec_ref(x_18); +x_26 = lean_ctor_get(x_20, 1); +lean_inc_ref(x_26); +lean_dec_ref(x_20); +x_27 = lean_ctor_get(x_22, 1); +lean_inc_ref(x_27); +x_28 = l_Lean_Expr_appFnCleanup___redArg(x_22); +x_29 = l_Lean_Meta_rwIfWith___closed__1; +x_30 = l_Lean_Expr_isConstOf(x_28, x_29); +if (x_30 == 0) +{ +uint8_t x_31; +x_31 = l_Lean_Expr_isApp(x_28); +if (x_31 == 0) +{ +lean_dec_ref(x_28); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_1); +x_8 = lean_box(0); +goto block_13; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_32 = lean_ctor_get(x_28, 1); +lean_inc_ref(x_32); +x_33 = l_Lean_Expr_appFnCleanup___redArg(x_28); +x_34 = l_Lean_Meta_rwIfWith___closed__3; +x_35 = l_Lean_Expr_isConstOf(x_33, x_34); +if (x_35 == 0) +{ +lean_object* x_36; uint8_t x_37; +x_36 = l_Lean_Meta_rwIfWith___closed__5; +x_37 = l_Lean_Expr_isConstOf(x_33, x_36); +if (x_37 == 0) +{ +lean_dec_ref(x_33); +lean_dec_ref(x_32); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_1); +x_8 = lean_box(0); +goto block_13; +} +else +{ +lean_object* x_38; +lean_inc(x_6); +lean_inc_ref(x_5); +lean_inc(x_4); +lean_inc_ref(x_3); +lean_inc_ref(x_1); +x_38 = lean_infer_type(x_1, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_38) == 0) +{ +lean_object* x_39; lean_object* x_40; +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +lean_dec_ref(x_38); +lean_inc(x_6); +lean_inc_ref(x_5); +lean_inc(x_4); +lean_inc_ref(x_3); +lean_inc_ref(x_27); +x_40 = l_Lean_Meta_isExprDefEq(x_27, x_39, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_40) == 0) +{ +uint8_t x_41; +x_41 = !lean_is_exclusive(x_40); +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_42 = lean_ctor_get(x_40, 0); +x_43 = l_Lean_Expr_constLevels_x21(x_33); +lean_dec_ref(x_33); +x_44 = lean_unbox(x_42); +lean_dec(x_42); +if (x_44 == 0) +{ +lean_object* x_45; +lean_free_object(x_40); +lean_inc(x_6); +lean_inc_ref(x_5); +lean_inc(x_4); +lean_inc_ref(x_3); +lean_inc_ref(x_1); +x_45 = lean_infer_type(x_1, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_45) == 0) +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +lean_dec_ref(x_45); +lean_inc_ref(x_27); +x_47 = l_Lean_mkNot(x_27); +x_48 = l_Lean_Meta_isExprDefEq(x_47, x_46, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_48) == 0) +{ +uint8_t x_49; +x_49 = !lean_is_exclusive(x_48); +if (x_49 == 0) +{ +lean_object* x_50; uint8_t x_51; +x_50 = lean_ctor_get(x_48, 0); +x_51 = lean_unbox(x_50); +lean_dec(x_50); +if (x_51 == 0) +{ +lean_free_object(x_48); +lean_dec(x_43); +lean_dec_ref(x_32); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec_ref(x_1); +x_8 = lean_box(0); +goto block_13; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +lean_dec_ref(x_2); +x_52 = l_Lean_Meta_rwIfWith___closed__7; +x_53 = l_Lean_mkConst(x_52, x_43); +x_54 = l_Lean_Meta_rwIfWith___closed__8; +x_55 = lean_array_push(x_54, x_27); +x_56 = lean_array_push(x_55, x_26); +x_57 = lean_array_push(x_56, x_1); +x_58 = lean_array_push(x_57, x_32); +x_59 = lean_array_push(x_58, x_25); +lean_inc_ref(x_24); +x_60 = lean_array_push(x_59, x_24); +x_61 = l_Lean_mkAppN(x_53, x_60); +lean_dec_ref(x_60); +x_62 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_62, 0, x_61); +x_63 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_63, 0, x_24); +lean_ctor_set(x_63, 1, x_62); +lean_ctor_set_uint8(x_63, sizeof(void*)*2, x_37); +lean_ctor_set(x_48, 0, x_63); +return x_48; +} +} +else +{ +lean_object* x_64; uint8_t x_65; +x_64 = lean_ctor_get(x_48, 0); +lean_inc(x_64); +lean_dec(x_48); +x_65 = lean_unbox(x_64); +lean_dec(x_64); +if (x_65 == 0) +{ +lean_dec(x_43); +lean_dec_ref(x_32); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec_ref(x_1); +x_8 = lean_box(0); +goto block_13; +} +else +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec_ref(x_2); +x_66 = l_Lean_Meta_rwIfWith___closed__7; +x_67 = l_Lean_mkConst(x_66, x_43); +x_68 = l_Lean_Meta_rwIfWith___closed__8; +x_69 = lean_array_push(x_68, x_27); +x_70 = lean_array_push(x_69, x_26); +x_71 = lean_array_push(x_70, x_1); +x_72 = lean_array_push(x_71, x_32); +x_73 = lean_array_push(x_72, x_25); +lean_inc_ref(x_24); +x_74 = lean_array_push(x_73, x_24); +x_75 = l_Lean_mkAppN(x_67, x_74); +lean_dec_ref(x_74); +x_76 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_76, 0, x_75); +x_77 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_77, 0, x_24); +lean_ctor_set(x_77, 1, x_76); +lean_ctor_set_uint8(x_77, sizeof(void*)*2, x_37); +x_78 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_78, 0, x_77); +return x_78; +} +} +} +else +{ +uint8_t x_79; +lean_dec(x_43); +lean_dec_ref(x_32); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_79 = !lean_is_exclusive(x_48); +if (x_79 == 0) +{ +return x_48; +} +else +{ +lean_object* x_80; lean_object* x_81; +x_80 = lean_ctor_get(x_48, 0); +lean_inc(x_80); +lean_dec(x_48); +x_81 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_81, 0, x_80); +return x_81; +} +} +} +else +{ +uint8_t x_82; +lean_dec(x_43); +lean_dec_ref(x_32); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_82 = !lean_is_exclusive(x_45); +if (x_82 == 0) +{ +return x_45; +} +else +{ +lean_object* x_83; lean_object* x_84; +x_83 = lean_ctor_get(x_45, 0); +lean_inc(x_83); +lean_dec(x_45); +x_84 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_84, 0, x_83); +return x_84; +} +} +} +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; lean_object* x_94; lean_object* x_95; lean_object* x_96; +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +x_85 = l_Lean_Meta_rwIfWith___closed__10; +x_86 = l_Lean_mkConst(x_85, x_43); +x_87 = l_Lean_Meta_rwIfWith___closed__8; +x_88 = lean_array_push(x_87, x_27); +x_89 = lean_array_push(x_88, x_26); +x_90 = lean_array_push(x_89, x_1); +x_91 = lean_array_push(x_90, x_32); +lean_inc_ref(x_25); +x_92 = lean_array_push(x_91, x_25); +x_93 = lean_array_push(x_92, x_24); +x_94 = l_Lean_mkAppN(x_86, x_93); +lean_dec_ref(x_93); +x_95 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_95, 0, x_94); +x_96 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_96, 0, x_25); +lean_ctor_set(x_96, 1, x_95); +lean_ctor_set_uint8(x_96, sizeof(void*)*2, x_37); +lean_ctor_set(x_40, 0, x_96); +return x_40; +} +} +else +{ +lean_object* x_97; lean_object* x_98; uint8_t x_99; +x_97 = lean_ctor_get(x_40, 0); +lean_inc(x_97); +lean_dec(x_40); +x_98 = l_Lean_Expr_constLevels_x21(x_33); +lean_dec_ref(x_33); +x_99 = lean_unbox(x_97); +lean_dec(x_97); +if (x_99 == 0) +{ +lean_object* x_100; +lean_inc(x_6); +lean_inc_ref(x_5); +lean_inc(x_4); +lean_inc_ref(x_3); +lean_inc_ref(x_1); +x_100 = lean_infer_type(x_1, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_100) == 0) +{ +lean_object* x_101; lean_object* x_102; lean_object* x_103; +x_101 = lean_ctor_get(x_100, 0); +lean_inc(x_101); +lean_dec_ref(x_100); +lean_inc_ref(x_27); +x_102 = l_Lean_mkNot(x_27); +x_103 = l_Lean_Meta_isExprDefEq(x_102, x_101, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_103) == 0) +{ +lean_object* x_104; lean_object* x_105; uint8_t x_106; +x_104 = lean_ctor_get(x_103, 0); +lean_inc(x_104); +if (lean_is_exclusive(x_103)) { + lean_ctor_release(x_103, 0); + x_105 = x_103; +} else { + lean_dec_ref(x_103); + x_105 = lean_box(0); +} +x_106 = lean_unbox(x_104); +lean_dec(x_104); +if (x_106 == 0) +{ +lean_dec(x_105); +lean_dec(x_98); +lean_dec_ref(x_32); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec_ref(x_1); +x_8 = lean_box(0); +goto block_13; +} +else +{ +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; +lean_dec_ref(x_2); +x_107 = l_Lean_Meta_rwIfWith___closed__7; +x_108 = l_Lean_mkConst(x_107, x_98); +x_109 = l_Lean_Meta_rwIfWith___closed__8; +x_110 = lean_array_push(x_109, x_27); +x_111 = lean_array_push(x_110, x_26); +x_112 = lean_array_push(x_111, x_1); +x_113 = lean_array_push(x_112, x_32); +x_114 = lean_array_push(x_113, x_25); +lean_inc_ref(x_24); +x_115 = lean_array_push(x_114, x_24); +x_116 = l_Lean_mkAppN(x_108, x_115); +lean_dec_ref(x_115); +x_117 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_117, 0, x_116); +x_118 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_118, 0, x_24); +lean_ctor_set(x_118, 1, x_117); +lean_ctor_set_uint8(x_118, sizeof(void*)*2, x_37); +if (lean_is_scalar(x_105)) { + x_119 = lean_alloc_ctor(0, 1, 0); +} else { + x_119 = x_105; +} +lean_ctor_set(x_119, 0, x_118); +return x_119; +} +} +else +{ +lean_object* x_120; lean_object* x_121; lean_object* x_122; +lean_dec(x_98); +lean_dec_ref(x_32); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_120 = lean_ctor_get(x_103, 0); +lean_inc(x_120); +if (lean_is_exclusive(x_103)) { + lean_ctor_release(x_103, 0); + x_121 = x_103; +} else { + lean_dec_ref(x_103); + x_121 = lean_box(0); +} +if (lean_is_scalar(x_121)) { + x_122 = lean_alloc_ctor(1, 1, 0); +} else { + x_122 = x_121; +} +lean_ctor_set(x_122, 0, x_120); +return x_122; +} +} +else +{ +lean_object* x_123; lean_object* x_124; lean_object* x_125; +lean_dec(x_98); +lean_dec_ref(x_32); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_123 = lean_ctor_get(x_100, 0); +lean_inc(x_123); +if (lean_is_exclusive(x_100)) { + lean_ctor_release(x_100, 0); + x_124 = x_100; +} else { + lean_dec_ref(x_100); + x_124 = lean_box(0); +} +if (lean_is_scalar(x_124)) { + x_125 = lean_alloc_ctor(1, 1, 0); +} else { + x_125 = x_124; +} +lean_ctor_set(x_125, 0, x_123); +return x_125; +} +} +else +{ +lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +x_126 = l_Lean_Meta_rwIfWith___closed__10; +x_127 = l_Lean_mkConst(x_126, x_98); +x_128 = l_Lean_Meta_rwIfWith___closed__8; +x_129 = lean_array_push(x_128, x_27); +x_130 = lean_array_push(x_129, x_26); +x_131 = lean_array_push(x_130, x_1); +x_132 = lean_array_push(x_131, x_32); +lean_inc_ref(x_25); +x_133 = lean_array_push(x_132, x_25); +x_134 = lean_array_push(x_133, x_24); +x_135 = l_Lean_mkAppN(x_127, x_134); +lean_dec_ref(x_134); +x_136 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_136, 0, x_135); +x_137 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_137, 0, x_25); +lean_ctor_set(x_137, 1, x_136); +lean_ctor_set_uint8(x_137, sizeof(void*)*2, x_37); +x_138 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_138, 0, x_137); +return x_138; +} +} +} +else +{ +uint8_t x_139; +lean_dec_ref(x_33); +lean_dec_ref(x_32); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_139 = !lean_is_exclusive(x_40); +if (x_139 == 0) +{ +return x_40; +} +else +{ +lean_object* x_140; lean_object* x_141; +x_140 = lean_ctor_get(x_40, 0); +lean_inc(x_140); +lean_dec(x_40); +x_141 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_141, 0, x_140); +return x_141; +} +} +} +else +{ +uint8_t x_142; +lean_dec_ref(x_33); +lean_dec_ref(x_32); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_142 = !lean_is_exclusive(x_38); +if (x_142 == 0) +{ +return x_38; +} +else +{ +lean_object* x_143; lean_object* x_144; +x_143 = lean_ctor_get(x_38, 0); +lean_inc(x_143); +lean_dec(x_38); +x_144 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_144, 0, x_143); +return x_144; +} +} +} +} +else +{ +lean_object* x_145; +lean_inc(x_6); +lean_inc_ref(x_5); +lean_inc(x_4); +lean_inc_ref(x_3); +lean_inc_ref(x_1); +x_145 = lean_infer_type(x_1, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_145) == 0) +{ +lean_object* x_146; lean_object* x_147; +x_146 = lean_ctor_get(x_145, 0); +lean_inc(x_146); +lean_dec_ref(x_145); +lean_inc(x_6); +lean_inc_ref(x_5); +lean_inc(x_4); +lean_inc_ref(x_3); +lean_inc_ref(x_27); +x_147 = l_Lean_Meta_isExprDefEq(x_27, x_146, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_147) == 0) +{ +uint8_t x_148; +x_148 = !lean_is_exclusive(x_147); +if (x_148 == 0) +{ +lean_object* x_149; lean_object* x_150; uint8_t x_151; +x_149 = lean_ctor_get(x_147, 0); +x_150 = l_Lean_Expr_constLevels_x21(x_33); +lean_dec_ref(x_33); +x_151 = lean_unbox(x_149); +lean_dec(x_149); +if (x_151 == 0) +{ +lean_object* x_152; +lean_free_object(x_147); +lean_inc(x_6); +lean_inc_ref(x_5); +lean_inc(x_4); +lean_inc_ref(x_3); +lean_inc_ref(x_1); +x_152 = lean_infer_type(x_1, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_152) == 0) +{ +lean_object* x_153; lean_object* x_154; lean_object* x_155; +x_153 = lean_ctor_get(x_152, 0); +lean_inc(x_153); +lean_dec_ref(x_152); +lean_inc_ref(x_27); +x_154 = l_Lean_mkNot(x_27); +x_155 = l_Lean_Meta_isExprDefEq(x_154, x_153, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_155) == 0) +{ +uint8_t x_156; +x_156 = !lean_is_exclusive(x_155); +if (x_156 == 0) +{ +lean_object* x_157; uint8_t x_158; +x_157 = lean_ctor_get(x_155, 0); +x_158 = lean_unbox(x_157); +lean_dec(x_157); +if (x_158 == 0) +{ +lean_free_object(x_155); +lean_dec(x_150); +lean_dec_ref(x_32); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec_ref(x_1); +x_8 = lean_box(0); +goto block_13; +} +else +{ +lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; +lean_dec_ref(x_2); +x_159 = l_Lean_Meta_rwIfWith___closed__11; +lean_inc_ref(x_1); +x_160 = lean_array_push(x_159, x_1); +lean_inc_ref(x_24); +x_161 = l_Lean_Expr_beta(x_24, x_160); +x_162 = l_Lean_Meta_rwIfWith___closed__13; +x_163 = l_Lean_mkConst(x_162, x_150); +x_164 = l_Lean_Meta_rwIfWith___closed__8; +x_165 = lean_array_push(x_164, x_27); +x_166 = lean_array_push(x_165, x_26); +x_167 = lean_array_push(x_166, x_1); +x_168 = lean_array_push(x_167, x_32); +x_169 = lean_array_push(x_168, x_25); +x_170 = lean_array_push(x_169, x_24); +x_171 = l_Lean_mkAppN(x_163, x_170); +lean_dec_ref(x_170); +x_172 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_172, 0, x_171); +x_173 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_173, 0, x_161); +lean_ctor_set(x_173, 1, x_172); +lean_ctor_set_uint8(x_173, sizeof(void*)*2, x_35); +lean_ctor_set(x_155, 0, x_173); +return x_155; +} +} +else +{ +lean_object* x_174; uint8_t x_175; +x_174 = lean_ctor_get(x_155, 0); +lean_inc(x_174); +lean_dec(x_155); +x_175 = lean_unbox(x_174); +lean_dec(x_174); +if (x_175 == 0) +{ +lean_dec(x_150); +lean_dec_ref(x_32); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec_ref(x_1); +x_8 = lean_box(0); +goto block_13; +} +else +{ +lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; +lean_dec_ref(x_2); +x_176 = l_Lean_Meta_rwIfWith___closed__11; +lean_inc_ref(x_1); +x_177 = lean_array_push(x_176, x_1); +lean_inc_ref(x_24); +x_178 = l_Lean_Expr_beta(x_24, x_177); +x_179 = l_Lean_Meta_rwIfWith___closed__13; +x_180 = l_Lean_mkConst(x_179, x_150); +x_181 = l_Lean_Meta_rwIfWith___closed__8; +x_182 = lean_array_push(x_181, x_27); +x_183 = lean_array_push(x_182, x_26); +x_184 = lean_array_push(x_183, x_1); +x_185 = lean_array_push(x_184, x_32); +x_186 = lean_array_push(x_185, x_25); +x_187 = lean_array_push(x_186, x_24); +x_188 = l_Lean_mkAppN(x_180, x_187); +lean_dec_ref(x_187); +x_189 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_189, 0, x_188); +x_190 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_190, 0, x_178); +lean_ctor_set(x_190, 1, x_189); +lean_ctor_set_uint8(x_190, sizeof(void*)*2, x_35); +x_191 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_191, 0, x_190); +return x_191; +} +} +} +else +{ +uint8_t x_192; +lean_dec(x_150); +lean_dec_ref(x_32); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_192 = !lean_is_exclusive(x_155); +if (x_192 == 0) +{ +return x_155; +} +else +{ +lean_object* x_193; lean_object* x_194; +x_193 = lean_ctor_get(x_155, 0); +lean_inc(x_193); +lean_dec(x_155); +x_194 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_194, 0, x_193); +return x_194; +} +} +} +else +{ +uint8_t x_195; +lean_dec(x_150); +lean_dec_ref(x_32); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_195 = !lean_is_exclusive(x_152); +if (x_195 == 0) +{ +return x_152; +} +else +{ +lean_object* x_196; lean_object* x_197; +x_196 = lean_ctor_get(x_152, 0); +lean_inc(x_196); +lean_dec(x_152); +x_197 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_197, 0, x_196); +return x_197; +} +} +} +else +{ +lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +x_198 = l_Lean_Meta_rwIfWith___closed__11; +lean_inc_ref(x_1); +x_199 = lean_array_push(x_198, x_1); +lean_inc_ref(x_25); +x_200 = l_Lean_Expr_beta(x_25, x_199); +x_201 = l_Lean_Meta_rwIfWith___closed__15; +x_202 = l_Lean_mkConst(x_201, x_150); +x_203 = l_Lean_Meta_rwIfWith___closed__8; +x_204 = lean_array_push(x_203, x_27); +x_205 = lean_array_push(x_204, x_26); +x_206 = lean_array_push(x_205, x_1); +x_207 = lean_array_push(x_206, x_32); +x_208 = lean_array_push(x_207, x_25); +x_209 = lean_array_push(x_208, x_24); +x_210 = l_Lean_mkAppN(x_202, x_209); +lean_dec_ref(x_209); +x_211 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_211, 0, x_210); +x_212 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_212, 0, x_200); +lean_ctor_set(x_212, 1, x_211); +lean_ctor_set_uint8(x_212, sizeof(void*)*2, x_35); +lean_ctor_set(x_147, 0, x_212); +return x_147; +} +} +else +{ +lean_object* x_213; lean_object* x_214; uint8_t x_215; +x_213 = lean_ctor_get(x_147, 0); +lean_inc(x_213); +lean_dec(x_147); +x_214 = l_Lean_Expr_constLevels_x21(x_33); +lean_dec_ref(x_33); +x_215 = lean_unbox(x_213); +lean_dec(x_213); +if (x_215 == 0) +{ +lean_object* x_216; +lean_inc(x_6); +lean_inc_ref(x_5); +lean_inc(x_4); +lean_inc_ref(x_3); +lean_inc_ref(x_1); +x_216 = lean_infer_type(x_1, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_216) == 0) +{ +lean_object* x_217; lean_object* x_218; lean_object* x_219; +x_217 = lean_ctor_get(x_216, 0); +lean_inc(x_217); +lean_dec_ref(x_216); +lean_inc_ref(x_27); +x_218 = l_Lean_mkNot(x_27); +x_219 = l_Lean_Meta_isExprDefEq(x_218, x_217, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_219) == 0) +{ +lean_object* x_220; lean_object* x_221; uint8_t x_222; +x_220 = lean_ctor_get(x_219, 0); +lean_inc(x_220); +if (lean_is_exclusive(x_219)) { + lean_ctor_release(x_219, 0); + x_221 = x_219; +} else { + lean_dec_ref(x_219); + x_221 = lean_box(0); +} +x_222 = lean_unbox(x_220); +lean_dec(x_220); +if (x_222 == 0) +{ +lean_dec(x_221); +lean_dec(x_214); +lean_dec_ref(x_32); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec_ref(x_1); +x_8 = lean_box(0); +goto block_13; +} +else +{ +lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; +lean_dec_ref(x_2); +x_223 = l_Lean_Meta_rwIfWith___closed__11; +lean_inc_ref(x_1); +x_224 = lean_array_push(x_223, x_1); +lean_inc_ref(x_24); +x_225 = l_Lean_Expr_beta(x_24, x_224); +x_226 = l_Lean_Meta_rwIfWith___closed__13; +x_227 = l_Lean_mkConst(x_226, x_214); +x_228 = l_Lean_Meta_rwIfWith___closed__8; +x_229 = lean_array_push(x_228, x_27); +x_230 = lean_array_push(x_229, x_26); +x_231 = lean_array_push(x_230, x_1); +x_232 = lean_array_push(x_231, x_32); +x_233 = lean_array_push(x_232, x_25); +x_234 = lean_array_push(x_233, x_24); +x_235 = l_Lean_mkAppN(x_227, x_234); +lean_dec_ref(x_234); +x_236 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_236, 0, x_235); +x_237 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_237, 0, x_225); +lean_ctor_set(x_237, 1, x_236); +lean_ctor_set_uint8(x_237, sizeof(void*)*2, x_35); +if (lean_is_scalar(x_221)) { + x_238 = lean_alloc_ctor(0, 1, 0); +} else { + x_238 = x_221; +} +lean_ctor_set(x_238, 0, x_237); +return x_238; +} +} +else +{ +lean_object* x_239; lean_object* x_240; lean_object* x_241; +lean_dec(x_214); +lean_dec_ref(x_32); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_239 = lean_ctor_get(x_219, 0); +lean_inc(x_239); +if (lean_is_exclusive(x_219)) { + lean_ctor_release(x_219, 0); + x_240 = x_219; +} else { + lean_dec_ref(x_219); + x_240 = lean_box(0); +} +if (lean_is_scalar(x_240)) { + x_241 = lean_alloc_ctor(1, 1, 0); +} else { + x_241 = x_240; +} +lean_ctor_set(x_241, 0, x_239); +return x_241; +} +} +else +{ +lean_object* x_242; lean_object* x_243; lean_object* x_244; +lean_dec(x_214); +lean_dec_ref(x_32); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_242 = lean_ctor_get(x_216, 0); +lean_inc(x_242); +if (lean_is_exclusive(x_216)) { + lean_ctor_release(x_216, 0); + x_243 = x_216; +} else { + lean_dec_ref(x_216); + x_243 = lean_box(0); +} +if (lean_is_scalar(x_243)) { + x_244 = lean_alloc_ctor(1, 1, 0); +} else { + x_244 = x_243; +} +lean_ctor_set(x_244, 0, x_242); +return x_244; +} +} +else +{ +lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +x_245 = l_Lean_Meta_rwIfWith___closed__11; +lean_inc_ref(x_1); +x_246 = lean_array_push(x_245, x_1); +lean_inc_ref(x_25); +x_247 = l_Lean_Expr_beta(x_25, x_246); +x_248 = l_Lean_Meta_rwIfWith___closed__15; +x_249 = l_Lean_mkConst(x_248, x_214); +x_250 = l_Lean_Meta_rwIfWith___closed__8; +x_251 = lean_array_push(x_250, x_27); +x_252 = lean_array_push(x_251, x_26); +x_253 = lean_array_push(x_252, x_1); +x_254 = lean_array_push(x_253, x_32); +x_255 = lean_array_push(x_254, x_25); +x_256 = lean_array_push(x_255, x_24); +x_257 = l_Lean_mkAppN(x_249, x_256); +lean_dec_ref(x_256); +x_258 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_258, 0, x_257); +x_259 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_259, 0, x_247); +lean_ctor_set(x_259, 1, x_258); +lean_ctor_set_uint8(x_259, sizeof(void*)*2, x_35); +x_260 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_260, 0, x_259); +return x_260; +} +} +} +else +{ +uint8_t x_261; +lean_dec_ref(x_33); +lean_dec_ref(x_32); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_261 = !lean_is_exclusive(x_147); +if (x_261 == 0) +{ +return x_147; +} +else +{ +lean_object* x_262; lean_object* x_263; +x_262 = lean_ctor_get(x_147, 0); +lean_inc(x_262); +lean_dec(x_147); +x_263 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_263, 0, x_262); +return x_263; +} +} +} +else +{ +uint8_t x_264; +lean_dec_ref(x_33); +lean_dec_ref(x_32); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_264 = !lean_is_exclusive(x_145); +if (x_264 == 0) +{ +return x_145; +} +else +{ +lean_object* x_265; lean_object* x_266; +x_265 = lean_ctor_get(x_145, 0); +lean_inc(x_265); +lean_dec(x_145); +x_266 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_266, 0, x_265); +return x_266; +} +} +} +} +} +else +{ +lean_object* x_267; +lean_inc(x_6); +lean_inc_ref(x_5); +lean_inc(x_4); +lean_inc_ref(x_3); +lean_inc_ref(x_1); +x_267 = lean_infer_type(x_1, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_267) == 0) +{ +lean_object* x_268; lean_object* x_269; lean_object* x_270; +x_268 = lean_ctor_get(x_267, 0); +lean_inc(x_268); +lean_dec_ref(x_267); +x_269 = l_Lean_Meta_rwIfWith___closed__19; +lean_inc(x_6); +lean_inc_ref(x_5); +lean_inc(x_4); +lean_inc_ref(x_3); +lean_inc_ref(x_26); +x_270 = l_Lean_Meta_mkEq(x_26, x_269, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_270) == 0) +{ +lean_object* x_271; lean_object* x_272; +x_271 = lean_ctor_get(x_270, 0); +lean_inc(x_271); +lean_dec_ref(x_270); +lean_inc(x_6); +lean_inc_ref(x_5); +lean_inc(x_4); +lean_inc_ref(x_3); +x_272 = l_Lean_Meta_isExprDefEq(x_268, x_271, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_272) == 0) +{ +uint8_t x_273; +x_273 = !lean_is_exclusive(x_272); +if (x_273 == 0) +{ +lean_object* x_274; lean_object* x_275; uint8_t x_276; +x_274 = lean_ctor_get(x_272, 0); +x_275 = l_Lean_Expr_constLevels_x21(x_28); +lean_dec_ref(x_28); +x_276 = lean_unbox(x_274); +lean_dec(x_274); +if (x_276 == 0) +{ +lean_object* x_277; +lean_free_object(x_272); +lean_inc(x_6); +lean_inc_ref(x_5); +lean_inc(x_4); +lean_inc_ref(x_3); +lean_inc_ref(x_1); +x_277 = lean_infer_type(x_1, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_277) == 0) +{ +lean_object* x_278; lean_object* x_279; lean_object* x_280; +x_278 = lean_ctor_get(x_277, 0); +lean_inc(x_278); +lean_dec_ref(x_277); +x_279 = l_Lean_Meta_rwIfWith___closed__22; +lean_inc(x_6); +lean_inc_ref(x_5); +lean_inc(x_4); +lean_inc_ref(x_3); +lean_inc_ref(x_26); +x_280 = l_Lean_Meta_mkEq(x_26, x_279, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_280) == 0) +{ +lean_object* x_281; lean_object* x_282; +x_281 = lean_ctor_get(x_280, 0); +lean_inc(x_281); +lean_dec_ref(x_280); +x_282 = l_Lean_Meta_isExprDefEq(x_278, x_281, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_282) == 0) +{ +uint8_t x_283; +x_283 = !lean_is_exclusive(x_282); +if (x_283 == 0) +{ +lean_object* x_284; uint8_t x_285; +x_284 = lean_ctor_get(x_282, 0); +x_285 = lean_unbox(x_284); +lean_dec(x_284); +if (x_285 == 0) +{ +lean_free_object(x_282); +lean_dec(x_275); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec_ref(x_1); +x_8 = lean_box(0); +goto block_13; +} +else +{ +lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; +lean_dec_ref(x_2); +x_286 = l_Lean_Meta_rwIfWith___closed__24; +x_287 = l_Lean_mkConst(x_286, x_275); +x_288 = l_Lean_Meta_rwIfWith___closed__25; +x_289 = lean_array_push(x_288, x_27); +x_290 = lean_array_push(x_289, x_26); +x_291 = lean_array_push(x_290, x_25); +lean_inc_ref(x_24); +x_292 = lean_array_push(x_291, x_24); +x_293 = lean_array_push(x_292, x_1); +x_294 = l_Lean_mkAppN(x_287, x_293); +lean_dec_ref(x_293); +x_295 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_295, 0, x_294); +x_296 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_296, 0, x_24); +lean_ctor_set(x_296, 1, x_295); +lean_ctor_set_uint8(x_296, sizeof(void*)*2, x_30); +lean_ctor_set(x_282, 0, x_296); +return x_282; +} +} +else +{ +lean_object* x_297; uint8_t x_298; +x_297 = lean_ctor_get(x_282, 0); +lean_inc(x_297); +lean_dec(x_282); +x_298 = lean_unbox(x_297); +lean_dec(x_297); +if (x_298 == 0) +{ +lean_dec(x_275); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec_ref(x_1); +x_8 = lean_box(0); +goto block_13; +} +else +{ +lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; +lean_dec_ref(x_2); +x_299 = l_Lean_Meta_rwIfWith___closed__24; +x_300 = l_Lean_mkConst(x_299, x_275); +x_301 = l_Lean_Meta_rwIfWith___closed__25; +x_302 = lean_array_push(x_301, x_27); +x_303 = lean_array_push(x_302, x_26); +x_304 = lean_array_push(x_303, x_25); +lean_inc_ref(x_24); +x_305 = lean_array_push(x_304, x_24); +x_306 = lean_array_push(x_305, x_1); +x_307 = l_Lean_mkAppN(x_300, x_306); +lean_dec_ref(x_306); +x_308 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_308, 0, x_307); +x_309 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_309, 0, x_24); +lean_ctor_set(x_309, 1, x_308); +lean_ctor_set_uint8(x_309, sizeof(void*)*2, x_30); +x_310 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_310, 0, x_309); +return x_310; +} +} +} +else +{ +uint8_t x_311; +lean_dec(x_275); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_311 = !lean_is_exclusive(x_282); +if (x_311 == 0) +{ +return x_282; +} +else +{ +lean_object* x_312; lean_object* x_313; +x_312 = lean_ctor_get(x_282, 0); +lean_inc(x_312); +lean_dec(x_282); +x_313 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_313, 0, x_312); +return x_313; +} +} +} +else +{ +uint8_t x_314; +lean_dec(x_278); +lean_dec(x_275); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_314 = !lean_is_exclusive(x_280); +if (x_314 == 0) +{ +return x_280; +} +else +{ +lean_object* x_315; lean_object* x_316; +x_315 = lean_ctor_get(x_280, 0); +lean_inc(x_315); +lean_dec(x_280); +x_316 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_316, 0, x_315); +return x_316; +} +} +} +else +{ +uint8_t x_317; +lean_dec(x_275); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_317 = !lean_is_exclusive(x_277); +if (x_317 == 0) +{ +return x_277; +} +else +{ +lean_object* x_318; lean_object* x_319; +x_318 = lean_ctor_get(x_277, 0); +lean_inc(x_318); +lean_dec(x_277); +x_319 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_319, 0, x_318); +return x_319; +} +} +} +else +{ +lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +x_320 = l_Lean_Meta_rwIfWith___closed__27; +x_321 = l_Lean_mkConst(x_320, x_275); +x_322 = l_Lean_Meta_rwIfWith___closed__25; +x_323 = lean_array_push(x_322, x_27); +x_324 = lean_array_push(x_323, x_26); +lean_inc_ref(x_25); +x_325 = lean_array_push(x_324, x_25); +x_326 = lean_array_push(x_325, x_24); +x_327 = lean_array_push(x_326, x_1); +x_328 = l_Lean_mkAppN(x_321, x_327); +lean_dec_ref(x_327); +x_329 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_329, 0, x_328); +x_330 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_330, 0, x_25); +lean_ctor_set(x_330, 1, x_329); +lean_ctor_set_uint8(x_330, sizeof(void*)*2, x_30); +lean_ctor_set(x_272, 0, x_330); +return x_272; +} +} +else +{ +lean_object* x_331; lean_object* x_332; uint8_t x_333; +x_331 = lean_ctor_get(x_272, 0); +lean_inc(x_331); +lean_dec(x_272); +x_332 = l_Lean_Expr_constLevels_x21(x_28); +lean_dec_ref(x_28); +x_333 = lean_unbox(x_331); +lean_dec(x_331); +if (x_333 == 0) +{ +lean_object* x_334; +lean_inc(x_6); +lean_inc_ref(x_5); +lean_inc(x_4); +lean_inc_ref(x_3); +lean_inc_ref(x_1); +x_334 = lean_infer_type(x_1, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_334) == 0) +{ +lean_object* x_335; lean_object* x_336; lean_object* x_337; +x_335 = lean_ctor_get(x_334, 0); +lean_inc(x_335); +lean_dec_ref(x_334); +x_336 = l_Lean_Meta_rwIfWith___closed__22; +lean_inc(x_6); +lean_inc_ref(x_5); +lean_inc(x_4); +lean_inc_ref(x_3); +lean_inc_ref(x_26); +x_337 = l_Lean_Meta_mkEq(x_26, x_336, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_337) == 0) +{ +lean_object* x_338; lean_object* x_339; +x_338 = lean_ctor_get(x_337, 0); +lean_inc(x_338); +lean_dec_ref(x_337); +x_339 = l_Lean_Meta_isExprDefEq(x_335, x_338, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_339) == 0) +{ +lean_object* x_340; lean_object* x_341; uint8_t x_342; +x_340 = lean_ctor_get(x_339, 0); +lean_inc(x_340); +if (lean_is_exclusive(x_339)) { + lean_ctor_release(x_339, 0); + x_341 = x_339; +} else { + lean_dec_ref(x_339); + x_341 = lean_box(0); +} +x_342 = lean_unbox(x_340); +lean_dec(x_340); +if (x_342 == 0) +{ +lean_dec(x_341); +lean_dec(x_332); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec_ref(x_1); +x_8 = lean_box(0); +goto block_13; +} +else +{ +lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; +lean_dec_ref(x_2); +x_343 = l_Lean_Meta_rwIfWith___closed__24; +x_344 = l_Lean_mkConst(x_343, x_332); +x_345 = l_Lean_Meta_rwIfWith___closed__25; +x_346 = lean_array_push(x_345, x_27); +x_347 = lean_array_push(x_346, x_26); +x_348 = lean_array_push(x_347, x_25); +lean_inc_ref(x_24); +x_349 = lean_array_push(x_348, x_24); +x_350 = lean_array_push(x_349, x_1); +x_351 = l_Lean_mkAppN(x_344, x_350); +lean_dec_ref(x_350); +x_352 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_352, 0, x_351); +x_353 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_353, 0, x_24); +lean_ctor_set(x_353, 1, x_352); +lean_ctor_set_uint8(x_353, sizeof(void*)*2, x_30); +if (lean_is_scalar(x_341)) { + x_354 = lean_alloc_ctor(0, 1, 0); +} else { + x_354 = x_341; +} +lean_ctor_set(x_354, 0, x_353); +return x_354; +} +} +else +{ +lean_object* x_355; lean_object* x_356; lean_object* x_357; +lean_dec(x_332); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_355 = lean_ctor_get(x_339, 0); +lean_inc(x_355); +if (lean_is_exclusive(x_339)) { + lean_ctor_release(x_339, 0); + x_356 = x_339; +} else { + lean_dec_ref(x_339); + x_356 = lean_box(0); +} +if (lean_is_scalar(x_356)) { + x_357 = lean_alloc_ctor(1, 1, 0); +} else { + x_357 = x_356; +} +lean_ctor_set(x_357, 0, x_355); +return x_357; +} +} +else +{ +lean_object* x_358; lean_object* x_359; lean_object* x_360; +lean_dec(x_335); +lean_dec(x_332); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_358 = lean_ctor_get(x_337, 0); +lean_inc(x_358); +if (lean_is_exclusive(x_337)) { + lean_ctor_release(x_337, 0); + x_359 = x_337; +} else { + lean_dec_ref(x_337); + x_359 = lean_box(0); +} +if (lean_is_scalar(x_359)) { + x_360 = lean_alloc_ctor(1, 1, 0); +} else { + x_360 = x_359; +} +lean_ctor_set(x_360, 0, x_358); +return x_360; +} +} +else +{ +lean_object* x_361; lean_object* x_362; lean_object* x_363; +lean_dec(x_332); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_361 = lean_ctor_get(x_334, 0); +lean_inc(x_361); +if (lean_is_exclusive(x_334)) { + lean_ctor_release(x_334, 0); + x_362 = x_334; +} else { + lean_dec_ref(x_334); + x_362 = lean_box(0); +} +if (lean_is_scalar(x_362)) { + x_363 = lean_alloc_ctor(1, 1, 0); +} else { + x_363 = x_362; +} +lean_ctor_set(x_363, 0, x_361); +return x_363; +} +} +else +{ +lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +x_364 = l_Lean_Meta_rwIfWith___closed__27; +x_365 = l_Lean_mkConst(x_364, x_332); +x_366 = l_Lean_Meta_rwIfWith___closed__25; +x_367 = lean_array_push(x_366, x_27); +x_368 = lean_array_push(x_367, x_26); +lean_inc_ref(x_25); +x_369 = lean_array_push(x_368, x_25); +x_370 = lean_array_push(x_369, x_24); +x_371 = lean_array_push(x_370, x_1); +x_372 = l_Lean_mkAppN(x_365, x_371); +lean_dec_ref(x_371); +x_373 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_373, 0, x_372); +x_374 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_374, 0, x_25); +lean_ctor_set(x_374, 1, x_373); +lean_ctor_set_uint8(x_374, sizeof(void*)*2, x_30); +x_375 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_375, 0, x_374); +return x_375; +} +} +} +else +{ +uint8_t x_376; +lean_dec_ref(x_28); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_376 = !lean_is_exclusive(x_272); +if (x_376 == 0) +{ +return x_272; +} +else +{ +lean_object* x_377; lean_object* x_378; +x_377 = lean_ctor_get(x_272, 0); +lean_inc(x_377); +lean_dec(x_272); +x_378 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_378, 0, x_377); +return x_378; +} +} +} +else +{ +uint8_t x_379; +lean_dec(x_268); +lean_dec_ref(x_28); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_379 = !lean_is_exclusive(x_270); +if (x_379 == 0) +{ +return x_270; +} +else +{ +lean_object* x_380; lean_object* x_381; +x_380 = lean_ctor_get(x_270, 0); +lean_inc(x_380); +lean_dec(x_270); +x_381 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_381, 0, x_380); +return x_381; +} +} +} +else +{ +uint8_t x_382; +lean_dec_ref(x_28); +lean_dec_ref(x_27); +lean_dec_ref(x_26); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_382 = !lean_is_exclusive(x_267); +if (x_382 == 0) +{ +return x_267; +} +else +{ +lean_object* x_383; lean_object* x_384; +x_383 = lean_ctor_get(x_267, 0); +lean_inc(x_383); +lean_dec(x_267); +x_384 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_384, 0, x_383); +return x_384; +} +} +} +} +} +} +} +} +else +{ +uint8_t x_385; +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_385 = !lean_is_exclusive(x_14); +if (x_385 == 0) +{ +return x_14; +} +else +{ +lean_object* x_386; lean_object* x_387; +x_386 = lean_ctor_get(x_14, 0); +lean_inc(x_386); +lean_dec(x_14); +x_387 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_387, 0, x_386); +return x_387; +} +} +block_13: +{ +lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_box(0); +x_10 = 1; +x_11 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_11, 0, x_2); +lean_ctor_set(x_11, 1, x_9); +lean_ctor_set_uint8(x_11, sizeof(void*)*2, x_10); +x_12 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_12, 0, x_11); +return x_12; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_rwIfWith___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Meta_rwIfWith(x_1, x_2, x_3, x_4, x_5, x_6); +return x_8; +} +} +static size_t _init_l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4___redArg___closed__0() { +_start: +{ +size_t x_1; size_t x_2; size_t x_3; +x_1 = 5; +x_2 = 1; +x_3 = lean_usize_shift_left(x_2, x_1); +return x_3; +} +} +static size_t _init_l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4___redArg___closed__1() { +_start: +{ +size_t x_1; size_t x_2; size_t x_3; +x_1 = 1; +x_2 = l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4___redArg___closed__0; +x_3 = lean_usize_sub(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at___00Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4_spec__18___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_array_get_size(x_1); +x_5 = lean_nat_dec_lt(x_2, x_4); +if (x_5 == 0) +{ +lean_dec(x_2); +return x_5; +} +else +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_array_fget_borrowed(x_1, x_2); +x_7 = l_Lean_instBEqMVarId_beq(x_3, x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_unsigned_to_nat(1u); +x_9 = lean_nat_add(x_2, x_8); +lean_dec(x_2); +x_2 = x_9; +goto _start; +} +else +{ +lean_dec(x_2); +return x_7; +} +} +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4___redArg(lean_object* x_1, size_t x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_4; lean_object* x_5; size_t x_6; size_t x_7; size_t x_8; lean_object* x_9; lean_object* x_10; +x_4 = lean_ctor_get(x_1, 0); +lean_inc_ref(x_4); +lean_dec_ref(x_1); +x_5 = lean_box(2); +x_6 = 5; +x_7 = l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4___redArg___closed__1; +x_8 = lean_usize_land(x_2, x_7); +x_9 = lean_usize_to_nat(x_8); +x_10 = lean_array_get(x_5, x_4, x_9); +lean_dec(x_9); +lean_dec_ref(x_4); +switch (lean_obj_tag(x_10)) { +case 0: +{ +lean_object* x_11; uint8_t x_12; +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +lean_dec_ref(x_10); +x_12 = l_Lean_instBEqMVarId_beq(x_3, x_11); +lean_dec(x_11); +return x_12; +} +case 1: +{ +lean_object* x_13; size_t x_14; +x_13 = lean_ctor_get(x_10, 0); +lean_inc(x_13); +lean_dec_ref(x_10); +x_14 = lean_usize_shift_right(x_2, x_6); +x_1 = x_13; +x_2 = x_14; +goto _start; +} +default: +{ +uint8_t x_16; +x_16 = 0; +return x_16; +} +} +} +else +{ +lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_17 = lean_ctor_get(x_1, 0); +lean_inc_ref(x_17); +lean_dec_ref(x_1); +x_18 = lean_unsigned_to_nat(0u); +x_19 = l_Lean_PersistentHashMap_containsAtAux___at___00Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4_spec__18___redArg(x_17, x_18, x_3); +lean_dec_ref(x_17); +return x_19; +} +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0___redArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint64_t x_3; size_t x_4; uint8_t x_5; +x_3 = l_Lean_instHashableMVarId_hash(x_2); +x_4 = lean_uint64_to_usize(x_3); +x_5 = l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4___redArg(x_1, x_4, x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0___redArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; lean_object* x_9; +x_4 = lean_st_ref_get(x_2); +x_5 = lean_ctor_get(x_4, 0); +lean_inc_ref(x_5); +lean_dec(x_4); +x_6 = lean_ctor_get(x_5, 7); +lean_inc_ref(x_6); +lean_dec_ref(x_5); +x_7 = l_Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0___redArg(x_6, x_1); +x_8 = lean_box(x_7); +x_9 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_9, 0, x_8); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0___redArg(x_1, x_3); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_isMatcherApp___at___00Lean_Meta_rwMatcher_spec__1___redArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; +x_4 = lean_st_ref_get(x_2); +x_5 = lean_ctor_get(x_4, 0); +lean_inc_ref(x_5); +lean_dec(x_4); +x_6 = l_Lean_Meta_isMatcherAppCore(x_5, x_1); +x_7 = lean_box(x_6); +x_8 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_8, 0, x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_isMatcherApp___at___00Lean_Meta_rwMatcher_spec__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_7; +x_7 = l_Lean_Meta_isMatcherApp___at___00Lean_Meta_rwMatcher_spec__1___redArg(x_1, x_5); +return x_7; +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = l_Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0___redArg(x_2, x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at___00Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4_spec__18(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; +x_7 = l_Lean_PersistentHashMap_containsAtAux___at___00Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4_spec__18___redArg(x_2, x_5, x_6); +return x_7; +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4(lean_object* x_1, lean_object* x_2, size_t x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; +x_5 = l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4___redArg(x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00Lean_Meta_rwMatcher_spec__5___redArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_4; +x_4 = l_Lean_Expr_hasMVar(x_1); +if (x_4 == 0) +{ +lean_object* x_5; +x_5 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_5, 0, x_1); +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_6 = lean_st_ref_get(x_2); +x_7 = lean_ctor_get(x_6, 0); +lean_inc_ref(x_7); +lean_dec(x_6); +x_8 = l_Lean_instantiateMVarsCore(x_7, x_1); +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_ref(x_8); +x_11 = lean_st_ref_take(x_2); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_11, 0); +lean_dec(x_13); +lean_ctor_set(x_11, 0, x_10); +x_14 = lean_st_ref_set(x_2, x_11); +x_15 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_15, 0, x_9); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_16 = lean_ctor_get(x_11, 1); +x_17 = lean_ctor_get(x_11, 2); +x_18 = lean_ctor_get(x_11, 3); +x_19 = lean_ctor_get(x_11, 4); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_11); +x_20 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_20, 0, x_10); +lean_ctor_set(x_20, 1, x_16); +lean_ctor_set(x_20, 2, x_17); +lean_ctor_set(x_20, 3, x_18); +lean_ctor_set(x_20, 4, x_19); +x_21 = lean_st_ref_set(x_2, x_20); +x_22 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_22, 0, x_9); +return x_22; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00Lean_Meta_rwMatcher_spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_instantiateMVars___at___00Lean_Meta_rwMatcher_spec__5___redArg(x_1, x_3); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00Lean_Meta_rwMatcher_spec__2___redArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; +x_4 = lean_ctor_get(x_2, 2); +x_5 = lean_ctor_get(x_2, 13); +x_6 = l_Lean_checkTraceOption(x_5, x_4, x_1); +x_7 = lean_box(x_6); +x_8 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_8, 0, x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00Lean_Meta_rwMatcher_spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_isTracingEnabledFor___at___00Lean_Meta_rwMatcher_spec__2___redArg(x_1, x_4); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_addMessageContextFull___at___00Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3_spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_7 = lean_st_ref_get(x_5); +x_8 = lean_ctor_get(x_7, 0); +lean_inc_ref(x_8); +lean_dec(x_7); +x_9 = lean_st_ref_get(x_3); +x_10 = lean_ctor_get(x_9, 0); +lean_inc_ref(x_10); +lean_dec(x_9); +x_11 = lean_ctor_get(x_2, 2); +x_12 = lean_ctor_get(x_4, 2); +lean_inc(x_12); +lean_inc_ref(x_11); +x_13 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_13, 0, x_8); +lean_ctor_set(x_13, 1, x_10); +lean_ctor_set(x_13, 2, x_11); +lean_ctor_set(x_13, 3, x_12); +x_14 = lean_alloc_ctor(3, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_1); +x_15 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_15, 0, x_14); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_Meta_rwMatcher_spec__6___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_7 = lean_ctor_get(x_4, 5); +x_8 = l_Lean_addMessageContextFull___at___00Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3_spec__4(x_1, x_2, x_3, x_4, x_5); +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_8, 0); +lean_inc(x_7); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_7); +lean_ctor_set(x_11, 1, x_10); +lean_ctor_set_tag(x_8, 1); +lean_ctor_set(x_8, 0, x_11); +return x_8; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_8, 0); +lean_inc(x_12); +lean_dec(x_8); +lean_inc(x_7); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_7); +lean_ctor_set(x_13, 1, x_12); +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_13); +return x_14; +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_Meta_rwMatcher_spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_throwError___at___00Lean_Meta_rwMatcher_spec__6___redArg(x_2, x_3, x_4, x_5, x_6); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__14_spec__16(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_ctor_get(x_5, 1); +lean_inc_ref(x_6); +lean_dec(x_5); +x_7 = lean_unsigned_to_nat(0u); +x_8 = lean_array_uset(x_3, x_2, x_7); +x_9 = 1; +x_10 = lean_usize_add(x_2, x_9); +x_11 = lean_array_uset(x_8, x_2, x_6); +x_2 = x_10; +x_3 = x_11; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__14(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_7); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_11 = lean_ctor_get(x_7, 5); +x_12 = lean_st_ref_get(x_8); +x_13 = lean_ctor_get(x_12, 4); +lean_inc_ref(x_13); +lean_dec(x_12); +x_14 = lean_ctor_get(x_13, 0); +lean_inc_ref(x_14); +lean_dec_ref(x_13); +x_15 = l_Lean_replaceRef(x_3, x_11); +lean_dec(x_11); +lean_ctor_set(x_7, 5, x_15); +x_16 = l_Lean_PersistentArray_toArray___redArg(x_14); +lean_dec_ref(x_14); +x_17 = lean_array_size(x_16); +x_18 = 0; +x_19 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__14_spec__16(x_17, x_18, x_16); +x_20 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_20, 0, x_2); +lean_ctor_set(x_20, 1, x_4); +lean_ctor_set(x_20, 2, x_19); +x_21 = l_Lean_addMessageContextFull___at___00Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3_spec__4(x_20, x_5, x_6, x_7, x_8); +lean_dec_ref(x_7); +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_23 = lean_ctor_get(x_21, 0); +x_24 = lean_st_ref_take(x_8); +x_25 = !lean_is_exclusive(x_24); +if (x_25 == 0) +{ +lean_object* x_26; uint8_t x_27; +x_26 = lean_ctor_get(x_24, 4); +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_28 = lean_ctor_get(x_26, 0); +lean_dec(x_28); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_3); +lean_ctor_set(x_29, 1, x_23); +x_30 = l_Lean_PersistentArray_push___redArg(x_1, x_29); +lean_ctor_set(x_26, 0, x_30); +x_31 = lean_st_ref_set(x_8, x_24); +x_32 = lean_box(0); +lean_ctor_set(x_21, 0, x_32); +return x_21; +} +else +{ +uint64_t x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_33 = lean_ctor_get_uint64(x_26, sizeof(void*)*1); +lean_dec(x_26); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_3); +lean_ctor_set(x_34, 1, x_23); +x_35 = l_Lean_PersistentArray_push___redArg(x_1, x_34); +x_36 = lean_alloc_ctor(0, 1, 8); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set_uint64(x_36, sizeof(void*)*1, x_33); +lean_ctor_set(x_24, 4, x_36); +x_37 = lean_st_ref_set(x_8, x_24); +x_38 = lean_box(0); +lean_ctor_set(x_21, 0, x_38); +return x_21; +} +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint64_t x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_39 = lean_ctor_get(x_24, 4); +x_40 = lean_ctor_get(x_24, 0); +x_41 = lean_ctor_get(x_24, 1); +x_42 = lean_ctor_get(x_24, 2); +x_43 = lean_ctor_get(x_24, 3); +x_44 = lean_ctor_get(x_24, 5); +x_45 = lean_ctor_get(x_24, 6); +x_46 = lean_ctor_get(x_24, 7); +x_47 = lean_ctor_get(x_24, 8); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_39); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_24); +x_48 = lean_ctor_get_uint64(x_39, sizeof(void*)*1); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + x_49 = x_39; +} else { + lean_dec_ref(x_39); + x_49 = lean_box(0); +} +x_50 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_50, 0, x_3); +lean_ctor_set(x_50, 1, x_23); +x_51 = l_Lean_PersistentArray_push___redArg(x_1, x_50); +if (lean_is_scalar(x_49)) { + x_52 = lean_alloc_ctor(0, 1, 8); +} else { + x_52 = x_49; +} +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set_uint64(x_52, sizeof(void*)*1, x_48); +x_53 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_53, 0, x_40); +lean_ctor_set(x_53, 1, x_41); +lean_ctor_set(x_53, 2, x_42); +lean_ctor_set(x_53, 3, x_43); +lean_ctor_set(x_53, 4, x_52); +lean_ctor_set(x_53, 5, x_44); +lean_ctor_set(x_53, 6, x_45); +lean_ctor_set(x_53, 7, x_46); +lean_ctor_set(x_53, 8, x_47); +x_54 = lean_st_ref_set(x_8, x_53); +x_55 = lean_box(0); +lean_ctor_set(x_21, 0, x_55); +return x_21; +} +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; uint64_t x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_56 = lean_ctor_get(x_21, 0); +lean_inc(x_56); +lean_dec(x_21); +x_57 = lean_st_ref_take(x_8); +x_58 = lean_ctor_get(x_57, 4); +lean_inc_ref(x_58); +x_59 = lean_ctor_get(x_57, 0); +lean_inc_ref(x_59); +x_60 = lean_ctor_get(x_57, 1); +lean_inc(x_60); +x_61 = lean_ctor_get(x_57, 2); +lean_inc_ref(x_61); +x_62 = lean_ctor_get(x_57, 3); +lean_inc_ref(x_62); +x_63 = lean_ctor_get(x_57, 5); +lean_inc_ref(x_63); +x_64 = lean_ctor_get(x_57, 6); +lean_inc_ref(x_64); +x_65 = lean_ctor_get(x_57, 7); +lean_inc_ref(x_65); +x_66 = lean_ctor_get(x_57, 8); +lean_inc_ref(x_66); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + lean_ctor_release(x_57, 1); + lean_ctor_release(x_57, 2); + lean_ctor_release(x_57, 3); + lean_ctor_release(x_57, 4); + lean_ctor_release(x_57, 5); + lean_ctor_release(x_57, 6); + lean_ctor_release(x_57, 7); + lean_ctor_release(x_57, 8); + x_67 = x_57; +} else { + lean_dec_ref(x_57); + x_67 = lean_box(0); +} +x_68 = lean_ctor_get_uint64(x_58, sizeof(void*)*1); +if (lean_is_exclusive(x_58)) { + lean_ctor_release(x_58, 0); + x_69 = x_58; +} else { + lean_dec_ref(x_58); + x_69 = lean_box(0); +} +x_70 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_70, 0, x_3); +lean_ctor_set(x_70, 1, x_56); +x_71 = l_Lean_PersistentArray_push___redArg(x_1, x_70); +if (lean_is_scalar(x_69)) { + x_72 = lean_alloc_ctor(0, 1, 8); +} else { + x_72 = x_69; +} +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set_uint64(x_72, sizeof(void*)*1, x_68); +if (lean_is_scalar(x_67)) { + x_73 = lean_alloc_ctor(0, 9, 0); +} else { + x_73 = x_67; +} +lean_ctor_set(x_73, 0, x_59); +lean_ctor_set(x_73, 1, x_60); +lean_ctor_set(x_73, 2, x_61); +lean_ctor_set(x_73, 3, x_62); +lean_ctor_set(x_73, 4, x_72); +lean_ctor_set(x_73, 5, x_63); +lean_ctor_set(x_73, 6, x_64); +lean_ctor_set(x_73, 7, x_65); +lean_ctor_set(x_73, 8, x_66); +x_74 = lean_st_ref_set(x_8, x_73); +x_75 = lean_box(0); +x_76 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_76, 0, x_75); +return x_76; +} +} +else +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; lean_object* x_90; uint8_t x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; size_t x_99; size_t x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; uint64_t x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; +x_77 = lean_ctor_get(x_7, 0); +x_78 = lean_ctor_get(x_7, 1); +x_79 = lean_ctor_get(x_7, 2); +x_80 = lean_ctor_get(x_7, 3); +x_81 = lean_ctor_get(x_7, 4); +x_82 = lean_ctor_get(x_7, 5); +x_83 = lean_ctor_get(x_7, 6); +x_84 = lean_ctor_get(x_7, 7); +x_85 = lean_ctor_get(x_7, 8); +x_86 = lean_ctor_get(x_7, 9); +x_87 = lean_ctor_get(x_7, 10); +x_88 = lean_ctor_get(x_7, 11); +x_89 = lean_ctor_get_uint8(x_7, sizeof(void*)*14); +x_90 = lean_ctor_get(x_7, 12); +x_91 = lean_ctor_get_uint8(x_7, sizeof(void*)*14 + 1); +x_92 = lean_ctor_get(x_7, 13); +lean_inc(x_92); +lean_inc(x_90); +lean_inc(x_88); +lean_inc(x_87); +lean_inc(x_86); +lean_inc(x_85); +lean_inc(x_84); +lean_inc(x_83); +lean_inc(x_82); +lean_inc(x_81); +lean_inc(x_80); +lean_inc(x_79); +lean_inc(x_78); +lean_inc(x_77); +lean_dec(x_7); +x_93 = lean_st_ref_get(x_8); +x_94 = lean_ctor_get(x_93, 4); +lean_inc_ref(x_94); +lean_dec(x_93); +x_95 = lean_ctor_get(x_94, 0); +lean_inc_ref(x_95); +lean_dec_ref(x_94); +x_96 = l_Lean_replaceRef(x_3, x_82); +lean_dec(x_82); +x_97 = lean_alloc_ctor(0, 14, 2); +lean_ctor_set(x_97, 0, x_77); +lean_ctor_set(x_97, 1, x_78); +lean_ctor_set(x_97, 2, x_79); +lean_ctor_set(x_97, 3, x_80); +lean_ctor_set(x_97, 4, x_81); +lean_ctor_set(x_97, 5, x_96); +lean_ctor_set(x_97, 6, x_83); +lean_ctor_set(x_97, 7, x_84); +lean_ctor_set(x_97, 8, x_85); +lean_ctor_set(x_97, 9, x_86); +lean_ctor_set(x_97, 10, x_87); +lean_ctor_set(x_97, 11, x_88); +lean_ctor_set(x_97, 12, x_90); +lean_ctor_set(x_97, 13, x_92); +lean_ctor_set_uint8(x_97, sizeof(void*)*14, x_89); +lean_ctor_set_uint8(x_97, sizeof(void*)*14 + 1, x_91); +x_98 = l_Lean_PersistentArray_toArray___redArg(x_95); +lean_dec_ref(x_95); +x_99 = lean_array_size(x_98); +x_100 = 0; +x_101 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__14_spec__16(x_99, x_100, x_98); +x_102 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_102, 0, x_2); +lean_ctor_set(x_102, 1, x_4); +lean_ctor_set(x_102, 2, x_101); +x_103 = l_Lean_addMessageContextFull___at___00Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3_spec__4(x_102, x_5, x_6, x_97, x_8); +lean_dec_ref(x_97); +x_104 = lean_ctor_get(x_103, 0); +lean_inc(x_104); +if (lean_is_exclusive(x_103)) { + lean_ctor_release(x_103, 0); + x_105 = x_103; +} else { + lean_dec_ref(x_103); + x_105 = lean_box(0); +} +x_106 = lean_st_ref_take(x_8); +x_107 = lean_ctor_get(x_106, 4); +lean_inc_ref(x_107); +x_108 = lean_ctor_get(x_106, 0); +lean_inc_ref(x_108); +x_109 = lean_ctor_get(x_106, 1); +lean_inc(x_109); +x_110 = lean_ctor_get(x_106, 2); +lean_inc_ref(x_110); +x_111 = lean_ctor_get(x_106, 3); +lean_inc_ref(x_111); +x_112 = lean_ctor_get(x_106, 5); +lean_inc_ref(x_112); +x_113 = lean_ctor_get(x_106, 6); +lean_inc_ref(x_113); +x_114 = lean_ctor_get(x_106, 7); +lean_inc_ref(x_114); +x_115 = lean_ctor_get(x_106, 8); +lean_inc_ref(x_115); +if (lean_is_exclusive(x_106)) { + lean_ctor_release(x_106, 0); + lean_ctor_release(x_106, 1); + lean_ctor_release(x_106, 2); + lean_ctor_release(x_106, 3); + lean_ctor_release(x_106, 4); + lean_ctor_release(x_106, 5); + lean_ctor_release(x_106, 6); + lean_ctor_release(x_106, 7); + lean_ctor_release(x_106, 8); + x_116 = x_106; +} else { + lean_dec_ref(x_106); + x_116 = lean_box(0); +} +x_117 = lean_ctor_get_uint64(x_107, sizeof(void*)*1); +if (lean_is_exclusive(x_107)) { + lean_ctor_release(x_107, 0); + x_118 = x_107; +} else { + lean_dec_ref(x_107); + x_118 = lean_box(0); +} +x_119 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_119, 0, x_3); +lean_ctor_set(x_119, 1, x_104); +x_120 = l_Lean_PersistentArray_push___redArg(x_1, x_119); +if (lean_is_scalar(x_118)) { + x_121 = lean_alloc_ctor(0, 1, 8); +} else { + x_121 = x_118; +} +lean_ctor_set(x_121, 0, x_120); +lean_ctor_set_uint64(x_121, sizeof(void*)*1, x_117); +if (lean_is_scalar(x_116)) { + x_122 = lean_alloc_ctor(0, 9, 0); +} else { + x_122 = x_116; +} +lean_ctor_set(x_122, 0, x_108); +lean_ctor_set(x_122, 1, x_109); +lean_ctor_set(x_122, 2, x_110); +lean_ctor_set(x_122, 3, x_111); +lean_ctor_set(x_122, 4, x_121); +lean_ctor_set(x_122, 5, x_112); +lean_ctor_set(x_122, 6, x_113); +lean_ctor_set(x_122, 7, x_114); +lean_ctor_set(x_122, 8, x_115); +x_123 = lean_st_ref_set(x_8, x_122); +x_124 = lean_box(0); +if (lean_is_scalar(x_105)) { + x_125 = lean_alloc_ctor(0, 1, 0); +} else { + x_125 = x_105; +} +lean_ctor_set(x_125, 0, x_124); +return x_125; +} +} +} +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__15___redArg(lean_object* x_1) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_ctor_set_tag(x_1, 1); +return x_1; +} +else +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +lean_dec(x_1); +x_5 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_5, 0, x_4); +return x_5; +} +} +else +{ +uint8_t x_6; +x_6 = !lean_is_exclusive(x_1); +if (x_6 == 0) +{ +lean_ctor_set_tag(x_1, 0); +return x_1; +} +else +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_1, 0); +lean_inc(x_7); +lean_dec(x_1); +x_8 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_8, 0, x_7); +return x_8; +} +} +} +} +static double _init_l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___closed__0() { +_start: +{ +lean_object* x_1; double x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_float_of_nat(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 53, 53); +return x_1; +} +} +static lean_object* _init_l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__0; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static double _init_l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__2() { +_start: +{ +lean_object* x_1; double x_2; +x_1 = lean_unsigned_to_nat(1000000000u); +x_2 = lean_float_of_nat(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_trace_profiler; +return x_1; +} +} +LEAN_EXPORT uint8_t l_Lean_Option_get___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__13(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_ctor_get(x_2, 0); +x_4 = lean_ctor_get(x_2, 1); +x_5 = l_Lean_KVMap_find(x_1, x_3); +if (lean_obj_tag(x_5) == 0) +{ +uint8_t x_6; +x_6 = lean_unbox(x_4); +return x_6; +} +else +{ +lean_object* x_7; +x_7 = lean_ctor_get(x_5, 0); +lean_inc(x_7); +lean_dec_ref(x_5); +if (lean_obj_tag(x_7) == 1) +{ +uint8_t x_8; +x_8 = lean_ctor_get_uint8(x_7, 0); +lean_dec_ref(x_7); +return x_8; +} +else +{ +uint8_t x_9; +lean_dec(x_7); +x_9 = lean_unbox(x_4); +return x_9; +} +} +} +} +static lean_object* _init_l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_trace_profiler_threshold; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Option_get___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__16(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_ctor_get(x_2, 0); +x_4 = lean_ctor_get(x_2, 1); +x_5 = l_Lean_KVMap_find(x_1, x_3); +if (lean_obj_tag(x_5) == 0) +{ +lean_inc(x_4); +return x_4; +} +else +{ +lean_object* x_6; +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +lean_dec_ref(x_5); +if (lean_obj_tag(x_6) == 3) +{ +lean_object* x_7; +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +lean_dec_ref(x_6); +return x_7; +} +else +{ +lean_dec(x_6); +lean_inc(x_4); +return x_4; +} +} +} +} +static double _init_l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__5() { +_start: +{ +lean_object* x_1; double x_2; +x_1 = lean_unsigned_to_nat(1000u); +x_2 = lean_float_of_nat(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12___redArg___closed__0() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(32u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12___redArg___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12___redArg___closed__0; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12___redArg___closed__2() { +_start: +{ +size_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = 5; +x_2 = lean_unsigned_to_nat(0u); +x_3 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12___redArg___closed__0; +x_4 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12___redArg___closed__1; +x_5 = lean_alloc_ctor(0, 4, sizeof(size_t)*1); +lean_ctor_set(x_5, 0, x_4); +lean_ctor_set(x_5, 1, x_3); +lean_ctor_set(x_5, 2, x_2); +lean_ctor_set(x_5, 3, x_2); +lean_ctor_set_usize(x_5, 4, x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12___redArg(lean_object* x_1) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_3 = lean_st_ref_get(x_1); +x_4 = lean_ctor_get(x_3, 4); +lean_inc_ref(x_4); +lean_dec(x_3); +x_5 = lean_ctor_get(x_4, 0); +lean_inc_ref(x_5); +lean_dec_ref(x_4); +x_6 = lean_st_ref_take(x_1); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; uint8_t x_9; +x_8 = lean_ctor_get(x_6, 4); +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_ctor_get(x_8, 0); +lean_dec(x_10); +x_11 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12___redArg___closed__2; +lean_ctor_set(x_8, 0, x_11); +x_12 = lean_st_ref_set(x_1, x_6); +x_13 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_13, 0, x_5); +return x_13; +} +else +{ +uint64_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_14 = lean_ctor_get_uint64(x_8, sizeof(void*)*1); +lean_dec(x_8); +x_15 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12___redArg___closed__2; +x_16 = lean_alloc_ctor(0, 1, 8); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set_uint64(x_16, sizeof(void*)*1, x_14); +lean_ctor_set(x_6, 4, x_16); +x_17 = lean_st_ref_set(x_1, x_6); +x_18 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_18, 0, x_5); +return x_18; +} +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint64_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_19 = lean_ctor_get(x_6, 4); +x_20 = lean_ctor_get(x_6, 0); +x_21 = lean_ctor_get(x_6, 1); +x_22 = lean_ctor_get(x_6, 2); +x_23 = lean_ctor_get(x_6, 3); +x_24 = lean_ctor_get(x_6, 5); +x_25 = lean_ctor_get(x_6, 6); +x_26 = lean_ctor_get(x_6, 7); +x_27 = lean_ctor_get(x_6, 8); +lean_inc(x_27); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); +lean_inc(x_19); +lean_inc(x_23); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_6); +x_28 = lean_ctor_get_uint64(x_19, sizeof(void*)*1); +if (lean_is_exclusive(x_19)) { + lean_ctor_release(x_19, 0); + x_29 = x_19; +} else { + lean_dec_ref(x_19); + x_29 = lean_box(0); +} +x_30 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12___redArg___closed__2; +if (lean_is_scalar(x_29)) { + x_31 = lean_alloc_ctor(0, 1, 8); +} else { + x_31 = x_29; +} +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set_uint64(x_31, sizeof(void*)*1, x_28); +x_32 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_32, 0, x_20); +lean_ctor_set(x_32, 1, x_21); +lean_ctor_set(x_32, 2, x_22); +lean_ctor_set(x_32, 3, x_23); +lean_ctor_set(x_32, 4, x_31); +lean_ctor_set(x_32, 5, x_24); +lean_ctor_set(x_32, 6, x_25); +lean_ctor_set(x_32, 7, x_26); +lean_ctor_set(x_32, 8, x_27); +x_33 = lean_st_ref_set(x_1, x_32); +x_34 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_34, 0, x_5); +return x_34; +} +} +} +static lean_object* _init_l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_trace_profiler_useHeartbeats; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_28; lean_object* x_29; double x_30; double x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_39; lean_object* x_40; lean_object* x_41; double x_42; double x_43; uint8_t x_44; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; uint8_t x_64; lean_object* x_65; double x_66; double x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_75; lean_object* x_76; double x_77; double x_78; lean_object* x_79; lean_object* x_80; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; double x_90; double x_91; uint8_t x_92; uint8_t x_93; lean_object* x_127; lean_object* x_128; lean_object* x_129; double x_130; double x_131; uint8_t x_132; double x_133; uint8_t x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; uint8_t x_159; lean_object* x_160; double x_161; double x_162; lean_object* x_163; lean_object* x_164; uint8_t x_165; uint8_t x_199; lean_object* x_200; double x_201; double x_202; lean_object* x_203; lean_object* x_204; double x_205; uint8_t x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; uint8_t x_249; +x_11 = lean_ctor_get(x_8, 2); +x_12 = lean_ctor_get(x_8, 5); +lean_inc(x_1); +x_85 = l_Lean_isTracingEnabledFor___at___00Lean_Meta_rwMatcher_spec__2___redArg(x_1, x_8); +x_86 = lean_ctor_get(x_85, 0); +lean_inc(x_86); +lean_dec_ref(x_85); +x_249 = lean_unbox(x_86); +if (x_249 == 0) +{ +lean_object* x_250; uint8_t x_251; +x_250 = l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__3; +x_251 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__13(x_11, x_250); +if (x_251 == 0) +{ +lean_object* x_252; +lean_dec(x_86); +lean_dec_ref(x_5); +lean_dec_ref(x_2); +lean_dec(x_1); +x_252 = lean_apply_5(x_3, x_6, x_7, x_8, x_9, lean_box(0)); +return x_252; +} +else +{ +lean_inc(x_12); +goto block_248; +} +} +else +{ +lean_inc(x_12); +goto block_248; +} +block_27: +{ +lean_object* x_22; +x_22 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__14(x_14, x_16, x_12, x_15, x_17, x_18, x_19, x_20); +lean_dec(x_20); +lean_dec(x_18); +lean_dec_ref(x_17); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; +lean_dec_ref(x_22); +x_23 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__15___redArg(x_13); +return x_23; +} +else +{ +uint8_t x_24; +lean_dec_ref(x_13); +x_24 = !lean_is_exclusive(x_22); +if (x_24 == 0) +{ +return x_22; +} +else +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_22, 0); +lean_inc(x_25); +lean_dec(x_22); +x_26 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_26, 0, x_25); +return x_26; +} +} +} +block_38: +{ +if (x_32 == 0) +{ +double x_35; lean_object* x_36; +x_35 = l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___closed__0; +x_36 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_36, 0, x_1); +lean_ctor_set(x_36, 1, x_5); +lean_ctor_set_float(x_36, sizeof(void*)*2, x_35); +lean_ctor_set_float(x_36, sizeof(void*)*2 + 8, x_35); +lean_ctor_set_uint8(x_36, sizeof(void*)*2 + 16, x_4); +x_13 = x_28; +x_14 = x_29; +x_15 = x_33; +x_16 = x_36; +x_17 = x_6; +x_18 = x_7; +x_19 = x_8; +x_20 = x_9; +x_21 = lean_box(0); +goto block_27; +} +else +{ +lean_object* x_37; +x_37 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_37, 0, x_1); +lean_ctor_set(x_37, 1, x_5); +lean_ctor_set_float(x_37, sizeof(void*)*2, x_30); +lean_ctor_set_float(x_37, sizeof(void*)*2 + 8, x_31); +lean_ctor_set_uint8(x_37, sizeof(void*)*2 + 16, x_4); +x_13 = x_28; +x_14 = x_29; +x_15 = x_33; +x_16 = x_37; +x_17 = x_6; +x_18 = x_7; +x_19 = x_8; +x_20 = x_9; +x_21 = lean_box(0); +goto block_27; +} +} +block_48: +{ +lean_object* x_45; +lean_inc(x_9); +lean_inc_ref(x_8); +lean_inc(x_7); +lean_inc_ref(x_6); +lean_inc_ref(x_40); +x_45 = lean_apply_6(x_2, x_40, x_6, x_7, x_8, x_9, lean_box(0)); +if (lean_obj_tag(x_45) == 0) +{ +lean_object* x_46; +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +lean_dec_ref(x_45); +x_28 = x_40; +x_29 = x_41; +x_30 = x_42; +x_31 = x_43; +x_32 = x_44; +x_33 = x_46; +x_34 = lean_box(0); +goto block_38; +} +else +{ +lean_object* x_47; +lean_dec_ref(x_45); +x_47 = l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__1; +x_28 = x_40; +x_29 = x_41; +x_30 = x_42; +x_31 = x_43; +x_32 = x_44; +x_33 = x_47; +x_34 = lean_box(0); +goto block_38; +} +} +block_63: +{ +lean_object* x_58; +x_58 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__14(x_49, x_52, x_12, x_50, x_53, x_54, x_55, x_56); +lean_dec(x_56); +lean_dec(x_54); +lean_dec_ref(x_53); +if (lean_obj_tag(x_58) == 0) +{ +lean_object* x_59; +lean_dec_ref(x_58); +x_59 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__15___redArg(x_51); +return x_59; +} +else +{ +uint8_t x_60; +lean_dec_ref(x_51); +x_60 = !lean_is_exclusive(x_58); +if (x_60 == 0) +{ +return x_58; +} +else +{ +lean_object* x_61; lean_object* x_62; +x_61 = lean_ctor_get(x_58, 0); +lean_inc(x_61); +lean_dec(x_58); +x_62 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_62, 0, x_61); +return x_62; +} +} +} +block_74: +{ +if (x_64 == 0) +{ +double x_71; lean_object* x_72; +x_71 = l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___closed__0; +x_72 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_72, 0, x_1); +lean_ctor_set(x_72, 1, x_5); +lean_ctor_set_float(x_72, sizeof(void*)*2, x_71); +lean_ctor_set_float(x_72, sizeof(void*)*2 + 8, x_71); +lean_ctor_set_uint8(x_72, sizeof(void*)*2 + 16, x_4); +x_49 = x_65; +x_50 = x_69; +x_51 = x_68; +x_52 = x_72; +x_53 = x_6; +x_54 = x_7; +x_55 = x_8; +x_56 = x_9; +x_57 = lean_box(0); +goto block_63; +} +else +{ +lean_object* x_73; +x_73 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_73, 0, x_1); +lean_ctor_set(x_73, 1, x_5); +lean_ctor_set_float(x_73, sizeof(void*)*2, x_67); +lean_ctor_set_float(x_73, sizeof(void*)*2 + 8, x_66); +lean_ctor_set_uint8(x_73, sizeof(void*)*2 + 16, x_4); +x_49 = x_65; +x_50 = x_69; +x_51 = x_68; +x_52 = x_73; +x_53 = x_6; +x_54 = x_7; +x_55 = x_8; +x_56 = x_9; +x_57 = lean_box(0); +goto block_63; +} +} +block_84: +{ +lean_object* x_81; +lean_inc(x_9); +lean_inc_ref(x_8); +lean_inc(x_7); +lean_inc_ref(x_6); +lean_inc_ref(x_80); +x_81 = lean_apply_6(x_2, x_80, x_6, x_7, x_8, x_9, lean_box(0)); +if (lean_obj_tag(x_81) == 0) +{ +lean_object* x_82; +x_82 = lean_ctor_get(x_81, 0); +lean_inc(x_82); +lean_dec_ref(x_81); +x_64 = x_75; +x_65 = x_76; +x_66 = x_77; +x_67 = x_78; +x_68 = x_80; +x_69 = x_82; +x_70 = lean_box(0); +goto block_74; +} +else +{ +lean_object* x_83; +lean_dec_ref(x_81); +x_83 = l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__1; +x_64 = x_75; +x_65 = x_76; +x_66 = x_77; +x_67 = x_78; +x_68 = x_80; +x_69 = x_83; +x_70 = lean_box(0); +goto block_74; +} +} +block_126: +{ +uint8_t x_94; +x_94 = lean_unbox(x_86); +lean_dec(x_86); +if (x_94 == 0) +{ +if (x_93 == 0) +{ +lean_object* x_95; uint8_t x_96; +lean_dec(x_12); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec_ref(x_5); +lean_dec_ref(x_2); +lean_dec(x_1); +x_95 = lean_st_ref_take(x_9); +x_96 = !lean_is_exclusive(x_95); +if (x_96 == 0) +{ +lean_object* x_97; uint8_t x_98; +x_97 = lean_ctor_get(x_95, 4); +x_98 = !lean_is_exclusive(x_97); +if (x_98 == 0) +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_99 = lean_ctor_get(x_97, 0); +x_100 = l_Lean_PersistentArray_append___redArg(x_89, x_99); +lean_dec_ref(x_99); +lean_ctor_set(x_97, 0, x_100); +x_101 = lean_st_ref_set(x_9, x_95); +lean_dec(x_9); +x_102 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__15___redArg(x_87); +return x_102; +} +else +{ +uint64_t x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; +x_103 = lean_ctor_get_uint64(x_97, sizeof(void*)*1); +x_104 = lean_ctor_get(x_97, 0); +lean_inc(x_104); +lean_dec(x_97); +x_105 = l_Lean_PersistentArray_append___redArg(x_89, x_104); +lean_dec_ref(x_104); +x_106 = lean_alloc_ctor(0, 1, 8); +lean_ctor_set(x_106, 0, x_105); +lean_ctor_set_uint64(x_106, sizeof(void*)*1, x_103); +lean_ctor_set(x_95, 4, x_106); +x_107 = lean_st_ref_set(x_9, x_95); +lean_dec(x_9); +x_108 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__15___redArg(x_87); +return x_108; +} +} +else +{ +lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; uint64_t x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; +x_109 = lean_ctor_get(x_95, 4); +x_110 = lean_ctor_get(x_95, 0); +x_111 = lean_ctor_get(x_95, 1); +x_112 = lean_ctor_get(x_95, 2); +x_113 = lean_ctor_get(x_95, 3); +x_114 = lean_ctor_get(x_95, 5); +x_115 = lean_ctor_get(x_95, 6); +x_116 = lean_ctor_get(x_95, 7); +x_117 = lean_ctor_get(x_95, 8); +lean_inc(x_117); +lean_inc(x_116); +lean_inc(x_115); +lean_inc(x_114); +lean_inc(x_109); +lean_inc(x_113); +lean_inc(x_112); +lean_inc(x_111); +lean_inc(x_110); +lean_dec(x_95); +x_118 = lean_ctor_get_uint64(x_109, sizeof(void*)*1); +x_119 = lean_ctor_get(x_109, 0); +lean_inc_ref(x_119); +if (lean_is_exclusive(x_109)) { + lean_ctor_release(x_109, 0); + x_120 = x_109; +} else { + lean_dec_ref(x_109); + x_120 = lean_box(0); +} +x_121 = l_Lean_PersistentArray_append___redArg(x_89, x_119); +lean_dec_ref(x_119); +if (lean_is_scalar(x_120)) { + x_122 = lean_alloc_ctor(0, 1, 8); +} else { + x_122 = x_120; +} +lean_ctor_set(x_122, 0, x_121); +lean_ctor_set_uint64(x_122, sizeof(void*)*1, x_118); +x_123 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_123, 0, x_110); +lean_ctor_set(x_123, 1, x_111); +lean_ctor_set(x_123, 2, x_112); +lean_ctor_set(x_123, 3, x_113); +lean_ctor_set(x_123, 4, x_122); +lean_ctor_set(x_123, 5, x_114); +lean_ctor_set(x_123, 6, x_115); +lean_ctor_set(x_123, 7, x_116); +lean_ctor_set(x_123, 8, x_117); +x_124 = lean_st_ref_set(x_9, x_123); +lean_dec(x_9); +x_125 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__15___redArg(x_87); +return x_125; +} +} +else +{ +x_39 = lean_box(0); +x_40 = x_87; +x_41 = x_89; +x_42 = x_90; +x_43 = x_91; +x_44 = x_92; +goto block_48; +} +} +else +{ +x_39 = lean_box(0); +x_40 = x_87; +x_41 = x_89; +x_42 = x_90; +x_43 = x_91; +x_44 = x_92; +goto block_48; +} +} +block_136: +{ +double x_134; uint8_t x_135; +x_134 = lean_float_sub(x_131, x_130); +x_135 = lean_float_decLt(x_133, x_134); +x_87 = x_128; +x_88 = lean_box(0); +x_89 = x_129; +x_90 = x_130; +x_91 = x_131; +x_92 = x_132; +x_93 = x_135; +goto block_126; +} +block_158: +{ +lean_object* x_142; double x_143; double x_144; double x_145; double x_146; double x_147; lean_object* x_148; uint8_t x_149; +x_142 = lean_io_mono_nanos_now(); +x_143 = lean_float_of_nat(x_139); +x_144 = l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__2; +x_145 = lean_float_div(x_143, x_144); +x_146 = lean_float_of_nat(x_142); +x_147 = lean_float_div(x_146, x_144); +x_148 = l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__3; +x_149 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__13(x_11, x_148); +if (x_149 == 0) +{ +x_87 = x_140; +x_88 = lean_box(0); +x_89 = x_138; +x_90 = x_145; +x_91 = x_147; +x_92 = x_149; +x_93 = x_149; +goto block_126; +} +else +{ +if (x_137 == 0) +{ +lean_object* x_150; lean_object* x_151; double x_152; double x_153; double x_154; +x_150 = l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__4; +x_151 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__16(x_11, x_150); +x_152 = lean_float_of_nat(x_151); +x_153 = l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__5; +x_154 = lean_float_div(x_152, x_153); +x_127 = lean_box(0); +x_128 = x_140; +x_129 = x_138; +x_130 = x_145; +x_131 = x_147; +x_132 = x_149; +x_133 = x_154; +goto block_136; +} +else +{ +lean_object* x_155; lean_object* x_156; double x_157; +x_155 = l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__4; +x_156 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__16(x_11, x_155); +x_157 = lean_float_of_nat(x_156); +x_127 = lean_box(0); +x_128 = x_140; +x_129 = x_138; +x_130 = x_145; +x_131 = x_147; +x_132 = x_149; +x_133 = x_157; +goto block_136; +} +} +} +block_198: +{ +uint8_t x_166; +x_166 = lean_unbox(x_86); +lean_dec(x_86); +if (x_166 == 0) +{ +if (x_165 == 0) +{ +lean_object* x_167; uint8_t x_168; +lean_dec(x_12); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec_ref(x_5); +lean_dec_ref(x_2); +lean_dec(x_1); +x_167 = lean_st_ref_take(x_9); +x_168 = !lean_is_exclusive(x_167); +if (x_168 == 0) +{ +lean_object* x_169; uint8_t x_170; +x_169 = lean_ctor_get(x_167, 4); +x_170 = !lean_is_exclusive(x_169); +if (x_170 == 0) +{ +lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; +x_171 = lean_ctor_get(x_169, 0); +x_172 = l_Lean_PersistentArray_append___redArg(x_160, x_171); +lean_dec_ref(x_171); +lean_ctor_set(x_169, 0, x_172); +x_173 = lean_st_ref_set(x_9, x_167); +lean_dec(x_9); +x_174 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__15___redArg(x_164); +return x_174; +} +else +{ +uint64_t x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; +x_175 = lean_ctor_get_uint64(x_169, sizeof(void*)*1); +x_176 = lean_ctor_get(x_169, 0); +lean_inc(x_176); +lean_dec(x_169); +x_177 = l_Lean_PersistentArray_append___redArg(x_160, x_176); +lean_dec_ref(x_176); +x_178 = lean_alloc_ctor(0, 1, 8); +lean_ctor_set(x_178, 0, x_177); +lean_ctor_set_uint64(x_178, sizeof(void*)*1, x_175); +lean_ctor_set(x_167, 4, x_178); +x_179 = lean_st_ref_set(x_9, x_167); +lean_dec(x_9); +x_180 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__15___redArg(x_164); +return x_180; +} +} +else +{ +lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; uint64_t x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; +x_181 = lean_ctor_get(x_167, 4); +x_182 = lean_ctor_get(x_167, 0); +x_183 = lean_ctor_get(x_167, 1); +x_184 = lean_ctor_get(x_167, 2); +x_185 = lean_ctor_get(x_167, 3); +x_186 = lean_ctor_get(x_167, 5); +x_187 = lean_ctor_get(x_167, 6); +x_188 = lean_ctor_get(x_167, 7); +x_189 = lean_ctor_get(x_167, 8); +lean_inc(x_189); +lean_inc(x_188); +lean_inc(x_187); +lean_inc(x_186); +lean_inc(x_181); +lean_inc(x_185); +lean_inc(x_184); +lean_inc(x_183); +lean_inc(x_182); +lean_dec(x_167); +x_190 = lean_ctor_get_uint64(x_181, sizeof(void*)*1); +x_191 = lean_ctor_get(x_181, 0); +lean_inc_ref(x_191); +if (lean_is_exclusive(x_181)) { + lean_ctor_release(x_181, 0); + x_192 = x_181; +} else { + lean_dec_ref(x_181); + x_192 = lean_box(0); +} +x_193 = l_Lean_PersistentArray_append___redArg(x_160, x_191); +lean_dec_ref(x_191); +if (lean_is_scalar(x_192)) { + x_194 = lean_alloc_ctor(0, 1, 8); +} else { + x_194 = x_192; +} +lean_ctor_set(x_194, 0, x_193); +lean_ctor_set_uint64(x_194, sizeof(void*)*1, x_190); +x_195 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_195, 0, x_182); +lean_ctor_set(x_195, 1, x_183); +lean_ctor_set(x_195, 2, x_184); +lean_ctor_set(x_195, 3, x_185); +lean_ctor_set(x_195, 4, x_194); +lean_ctor_set(x_195, 5, x_186); +lean_ctor_set(x_195, 6, x_187); +lean_ctor_set(x_195, 7, x_188); +lean_ctor_set(x_195, 8, x_189); +x_196 = lean_st_ref_set(x_9, x_195); +lean_dec(x_9); +x_197 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__15___redArg(x_164); +return x_197; +} +} +else +{ +x_75 = x_159; +x_76 = x_160; +x_77 = x_161; +x_78 = x_162; +x_79 = lean_box(0); +x_80 = x_164; +goto block_84; +} +} +else +{ +x_75 = x_159; +x_76 = x_160; +x_77 = x_161; +x_78 = x_162; +x_79 = lean_box(0); +x_80 = x_164; +goto block_84; +} +} +block_208: +{ +double x_206; uint8_t x_207; +x_206 = lean_float_sub(x_201, x_202); +x_207 = lean_float_decLt(x_205, x_206); +x_159 = x_199; +x_160 = x_200; +x_161 = x_201; +x_162 = x_202; +x_163 = lean_box(0); +x_164 = x_204; +x_165 = x_207; +goto block_198; +} +block_227: +{ +lean_object* x_214; double x_215; double x_216; lean_object* x_217; uint8_t x_218; +x_214 = lean_io_get_num_heartbeats(); +x_215 = lean_float_of_nat(x_211); +x_216 = lean_float_of_nat(x_214); +x_217 = l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__3; +x_218 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__13(x_11, x_217); +if (x_218 == 0) +{ +x_159 = x_218; +x_160 = x_210; +x_161 = x_216; +x_162 = x_215; +x_163 = lean_box(0); +x_164 = x_212; +x_165 = x_218; +goto block_198; +} +else +{ +if (x_209 == 0) +{ +lean_object* x_219; lean_object* x_220; double x_221; double x_222; double x_223; +x_219 = l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__4; +x_220 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__16(x_11, x_219); +x_221 = lean_float_of_nat(x_220); +x_222 = l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__5; +x_223 = lean_float_div(x_221, x_222); +x_199 = x_218; +x_200 = x_210; +x_201 = x_216; +x_202 = x_215; +x_203 = lean_box(0); +x_204 = x_212; +x_205 = x_223; +goto block_208; +} +else +{ +lean_object* x_224; lean_object* x_225; double x_226; +x_224 = l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__4; +x_225 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__16(x_11, x_224); +x_226 = lean_float_of_nat(x_225); +x_199 = x_218; +x_200 = x_210; +x_201 = x_216; +x_202 = x_215; +x_203 = lean_box(0); +x_204 = x_212; +x_205 = x_226; +goto block_208; +} +} +} +block_248: +{ +lean_object* x_228; lean_object* x_229; lean_object* x_230; uint8_t x_231; +x_228 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12___redArg(x_9); +x_229 = lean_ctor_get(x_228, 0); +lean_inc(x_229); +lean_dec_ref(x_228); +x_230 = l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__6; +x_231 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__13(x_11, x_230); +if (x_231 == 0) +{ +lean_object* x_232; lean_object* x_233; +x_232 = lean_io_mono_nanos_now(); +lean_inc(x_9); +lean_inc_ref(x_8); +lean_inc(x_7); +lean_inc_ref(x_6); +x_233 = lean_apply_5(x_3, x_6, x_7, x_8, x_9, lean_box(0)); +if (lean_obj_tag(x_233) == 0) +{ +uint8_t x_234; +x_234 = !lean_is_exclusive(x_233); +if (x_234 == 0) +{ +lean_ctor_set_tag(x_233, 1); +x_137 = x_231; +x_138 = x_229; +x_139 = x_232; +x_140 = x_233; +x_141 = lean_box(0); +goto block_158; +} +else +{ +lean_object* x_235; lean_object* x_236; +x_235 = lean_ctor_get(x_233, 0); +lean_inc(x_235); +lean_dec(x_233); +x_236 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_236, 0, x_235); +x_137 = x_231; +x_138 = x_229; +x_139 = x_232; +x_140 = x_236; +x_141 = lean_box(0); +goto block_158; +} +} +else +{ +uint8_t x_237; +x_237 = !lean_is_exclusive(x_233); +if (x_237 == 0) +{ +lean_ctor_set_tag(x_233, 0); +x_137 = x_231; +x_138 = x_229; +x_139 = x_232; +x_140 = x_233; +x_141 = lean_box(0); +goto block_158; +} +else +{ +lean_object* x_238; lean_object* x_239; +x_238 = lean_ctor_get(x_233, 0); +lean_inc(x_238); +lean_dec(x_233); +x_239 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_239, 0, x_238); +x_137 = x_231; +x_138 = x_229; +x_139 = x_232; +x_140 = x_239; +x_141 = lean_box(0); +goto block_158; +} +} +} +else +{ +lean_object* x_240; lean_object* x_241; +x_240 = lean_io_get_num_heartbeats(); +lean_inc(x_9); +lean_inc_ref(x_8); +lean_inc(x_7); +lean_inc_ref(x_6); +x_241 = lean_apply_5(x_3, x_6, x_7, x_8, x_9, lean_box(0)); +if (lean_obj_tag(x_241) == 0) +{ +uint8_t x_242; +x_242 = !lean_is_exclusive(x_241); +if (x_242 == 0) +{ +lean_ctor_set_tag(x_241, 1); +x_209 = x_231; +x_210 = x_229; +x_211 = x_240; +x_212 = x_241; +x_213 = lean_box(0); +goto block_227; +} +else +{ +lean_object* x_243; lean_object* x_244; +x_243 = lean_ctor_get(x_241, 0); +lean_inc(x_243); +lean_dec(x_241); +x_244 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_244, 0, x_243); +x_209 = x_231; +x_210 = x_229; +x_211 = x_240; +x_212 = x_244; +x_213 = lean_box(0); +goto block_227; +} +} +else +{ +uint8_t x_245; +x_245 = !lean_is_exclusive(x_241); +if (x_245 == 0) +{ +lean_ctor_set_tag(x_241, 0); +x_209 = x_231; +x_210 = x_229; +x_211 = x_240; +x_212 = x_241; +x_213 = lean_box(0); +goto block_227; +} +else +{ +lean_object* x_246; lean_object* x_247; +x_246 = lean_ctor_get(x_241, 0); +lean_inc(x_246); +lean_dec(x_241); +x_247 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_247, 0, x_246); +x_209 = x_231; +x_210 = x_229; +x_211 = x_240; +x_212 = x_247; +x_213 = lean_box(0); +goto block_227; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_12; +} +} +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__15(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_8; +x_8 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__15___redArg(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_6; +x_6 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12___redArg(x_4); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_rwMatcher___lam__2(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_9 = lean_box(0); +x_10 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_10, 0, x_1); +lean_ctor_set(x_10, 1, x_9); +lean_ctor_set_uint8(x_10, sizeof(void*)*2, x_2); +x_11 = lean_box(0); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set(x_12, 1, x_11); +x_13 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_13, 0, x_12); +return x_13; +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at___00Lean_Meta_rwMatcher_spec__7(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; +x_3 = l_List_reverse___redArg(x_2); +return x_3; +} +else +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_1, 0); +x_6 = lean_ctor_get(x_1, 1); +x_7 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_7, 0, x_5); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set(x_1, 0, x_7); +{ +lean_object* _tmp_0 = x_6; +lean_object* _tmp_1 = x_1; +x_1 = _tmp_0; +x_2 = _tmp_1; +} +goto _start; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_1, 0); +x_10 = lean_ctor_get(x_1, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_1); +x_11 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, x_9); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_2); +x_1 = x_10; +x_2 = x_12; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_rwMatcher_spec__4(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = l_Lean_Expr_mvarId_x21(x_5); +lean_dec(x_5); +x_9 = 1; +x_10 = lean_usize_add(x_2, x_9); +x_11 = lean_array_uset(x_7, x_2, x_8); +x_2 = x_10; +x_3 = x_11; +goto _start; +} +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___lam__0___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" rewriting with ", 16, 16); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___lam__0___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_rwMatcher___lam__0___closed__0; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___lam__0___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" in", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___lam__0___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_rwMatcher___lam__0___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_rwMatcher___lam__0(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_10 = l_Lean_exceptEmoji___redArg(x_4); +x_11 = l_Lean_stringToMessageData(x_10); +x_12 = l_Lean_Meta_rwMatcher___lam__0___closed__1; +x_13 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_13, 0, x_11); +lean_ctor_set(x_13, 1, x_12); +x_14 = l_Lean_MessageData_ofConstName(x_1, x_2); +x_15 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_14); +x_16 = l_Lean_Meta_rwMatcher___lam__0___closed__3; +x_17 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +x_18 = l_Lean_indentExpr(x_3); +x_19 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +x_20 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_20, 0, x_19); +return x_20; +} +} +LEAN_EXPORT lean_object* l___private_Init_While_0__Lean_Loop_forIn_loop___at___00Lean_Meta_rwMatcher_spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_7; +lean_inc(x_5); +lean_inc_ref(x_4); +lean_inc(x_3); +lean_inc_ref(x_2); +lean_inc_ref(x_1); +x_7 = l_Lean_Meta_reduceRecMatcher_x3f(x_1, x_2, x_3, x_4, x_5); +if (lean_obj_tag(x_7) == 0) +{ +uint8_t x_8; +x_8 = !lean_is_exclusive(x_7); +if (x_8 == 0) +{ +lean_object* x_9; +x_9 = lean_ctor_get(x_7, 0); +if (lean_obj_tag(x_9) == 1) +{ +lean_object* x_10; lean_object* x_11; +lean_free_object(x_7); +lean_dec_ref(x_1); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +lean_dec_ref(x_9); +x_11 = l_Lean_Expr_headBeta(x_10); +x_1 = x_11; +goto _start; +} +else +{ +lean_object* x_13; uint8_t x_14; +lean_dec(x_9); +lean_inc_ref(x_1); +x_13 = l_Lean_Expr_headBeta(x_1); +x_14 = lean_expr_eqv(x_1, x_13); +if (x_14 == 0) +{ +lean_free_object(x_7); +lean_dec_ref(x_1); +x_1 = x_13; +goto _start; +} +else +{ +lean_dec_ref(x_13); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec(x_3); +lean_dec_ref(x_2); +lean_ctor_set(x_7, 0, x_1); +return x_7; +} +} +} +else +{ +lean_object* x_16; +x_16 = lean_ctor_get(x_7, 0); +lean_inc(x_16); +lean_dec(x_7); +if (lean_obj_tag(x_16) == 1) +{ +lean_object* x_17; lean_object* x_18; +lean_dec_ref(x_1); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec_ref(x_16); +x_18 = l_Lean_Expr_headBeta(x_17); +x_1 = x_18; +goto _start; +} +else +{ +lean_object* x_20; uint8_t x_21; +lean_dec(x_16); +lean_inc_ref(x_1); +x_20 = l_Lean_Expr_headBeta(x_1); +x_21 = lean_expr_eqv(x_1, x_20); +if (x_21 == 0) +{ +lean_dec_ref(x_1); +x_1 = x_20; +goto _start; +} +else +{ +lean_object* x_23; +lean_dec_ref(x_20); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec(x_3); +lean_dec_ref(x_2); +x_23 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_23, 0, x_1); +return x_23; +} +} +} +} +else +{ +uint8_t x_24; +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_24 = !lean_is_exclusive(x_7); +if (x_24 == 0) +{ +return x_7; +} +else +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_7, 0); +lean_inc(x_25); +lean_dec(x_7); +x_26 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_26, 0, x_25); +return x_26; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Meta_rwMatcher_spec__9(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_10; lean_object* x_11; uint8_t x_16; +x_16 = lean_usize_dec_eq(x_2, x_3); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_21; +x_17 = lean_array_uget(x_1, x_2); +x_21 = l_Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0___redArg(x_17, x_6); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_22; uint8_t x_23; +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +lean_dec_ref(x_21); +x_23 = lean_unbox(x_22); +lean_dec(x_22); +if (x_23 == 0) +{ +x_18 = lean_box(0); +goto block_20; +} +else +{ +lean_dec(x_17); +x_10 = x_4; +x_11 = lean_box(0); +goto block_15; +} +} +else +{ +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_24; uint8_t x_25; +x_24 = lean_ctor_get(x_21, 0); +lean_inc(x_24); +lean_dec_ref(x_21); +x_25 = lean_unbox(x_24); +lean_dec(x_24); +if (x_25 == 0) +{ +lean_dec(x_17); +x_10 = x_4; +x_11 = lean_box(0); +goto block_15; +} +else +{ +x_18 = lean_box(0); +goto block_20; +} +} +else +{ +uint8_t x_26; +lean_dec(x_17); +lean_dec_ref(x_4); +x_26 = !lean_is_exclusive(x_21); +if (x_26 == 0) +{ +return x_21; +} +else +{ +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_21, 0); +lean_inc(x_27); +lean_dec(x_21); +x_28 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_28, 0, x_27); +return x_28; +} +} +} +block_20: +{ +lean_object* x_19; +x_19 = lean_array_push(x_4, x_17); +x_10 = x_19; +x_11 = lean_box(0); +goto block_15; +} +} +else +{ +lean_object* x_29; +x_29 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_29, 0, x_4); +return x_29; +} +block_15: +{ +size_t x_12; size_t x_13; +x_12 = 1; +x_13 = lean_usize_add(x_2, x_12); +x_2 = x_13; +x_4 = x_10; +goto _start; +} +} +} +static lean_object* _init_l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___closed__2() { +_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_addTrace___at___00Lean_Meta_rwMatcher_spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_8 = lean_ctor_get(x_5, 5); +x_9 = l_Lean_addMessageContextFull___at___00Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3_spec__4(x_2, x_3, x_4, x_5, x_6); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_st_ref_take(x_6); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_ctor_get(x_12, 4); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) +{ +lean_object* x_16; double x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_16 = lean_ctor_get(x_14, 0); +x_17 = l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___closed__0; +x_18 = 0; +x_19 = l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___closed__1; +x_20 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_20, 0, x_1); +lean_ctor_set(x_20, 1, x_19); +lean_ctor_set_float(x_20, sizeof(void*)*2, x_17); +lean_ctor_set_float(x_20, sizeof(void*)*2 + 8, x_17); +lean_ctor_set_uint8(x_20, sizeof(void*)*2 + 16, x_18); +x_21 = l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___closed__2; +x_22 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_11); +lean_ctor_set(x_22, 2, x_21); +lean_inc(x_8); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_8); +lean_ctor_set(x_23, 1, x_22); +x_24 = l_Lean_PersistentArray_push___redArg(x_16, x_23); +lean_ctor_set(x_14, 0, x_24); +x_25 = lean_st_ref_set(x_6, x_12); +x_26 = lean_box(0); +lean_ctor_set(x_9, 0, x_26); +return x_9; +} +else +{ +uint64_t x_27; lean_object* x_28; double x_29; uint8_t x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_27 = lean_ctor_get_uint64(x_14, sizeof(void*)*1); +x_28 = lean_ctor_get(x_14, 0); +lean_inc(x_28); +lean_dec(x_14); +x_29 = l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___closed__0; +x_30 = 0; +x_31 = l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___closed__1; +x_32 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_32, 0, x_1); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set_float(x_32, sizeof(void*)*2, x_29); +lean_ctor_set_float(x_32, sizeof(void*)*2 + 8, x_29); +lean_ctor_set_uint8(x_32, sizeof(void*)*2 + 16, x_30); +x_33 = l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___closed__2; +x_34 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_11); +lean_ctor_set(x_34, 2, x_33); +lean_inc(x_8); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_8); +lean_ctor_set(x_35, 1, x_34); +x_36 = l_Lean_PersistentArray_push___redArg(x_28, x_35); +x_37 = lean_alloc_ctor(0, 1, 8); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set_uint64(x_37, sizeof(void*)*1, x_27); +lean_ctor_set(x_12, 4, x_37); +x_38 = lean_st_ref_set(x_6, x_12); +x_39 = lean_box(0); +lean_ctor_set(x_9, 0, x_39); +return x_9; +} +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint64_t x_49; lean_object* x_50; lean_object* x_51; double x_52; uint8_t x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_40 = lean_ctor_get(x_12, 4); +x_41 = lean_ctor_get(x_12, 0); +x_42 = lean_ctor_get(x_12, 1); +x_43 = lean_ctor_get(x_12, 2); +x_44 = lean_ctor_get(x_12, 3); +x_45 = lean_ctor_get(x_12, 5); +x_46 = lean_ctor_get(x_12, 6); +x_47 = lean_ctor_get(x_12, 7); +x_48 = lean_ctor_get(x_12, 8); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_40); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_12); +x_49 = lean_ctor_get_uint64(x_40, sizeof(void*)*1); +x_50 = lean_ctor_get(x_40, 0); +lean_inc_ref(x_50); +if (lean_is_exclusive(x_40)) { + lean_ctor_release(x_40, 0); + x_51 = x_40; +} else { + lean_dec_ref(x_40); + x_51 = lean_box(0); +} +x_52 = l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___closed__0; +x_53 = 0; +x_54 = l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___closed__1; +x_55 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_55, 0, x_1); +lean_ctor_set(x_55, 1, x_54); +lean_ctor_set_float(x_55, sizeof(void*)*2, x_52); +lean_ctor_set_float(x_55, sizeof(void*)*2 + 8, x_52); +lean_ctor_set_uint8(x_55, sizeof(void*)*2 + 16, x_53); +x_56 = l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___closed__2; +x_57 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_11); +lean_ctor_set(x_57, 2, x_56); +lean_inc(x_8); +x_58 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_58, 0, x_8); +lean_ctor_set(x_58, 1, x_57); +x_59 = l_Lean_PersistentArray_push___redArg(x_50, x_58); +if (lean_is_scalar(x_51)) { + x_60 = lean_alloc_ctor(0, 1, 8); +} else { + x_60 = x_51; +} +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set_uint64(x_60, sizeof(void*)*1, x_49); +x_61 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_61, 0, x_41); +lean_ctor_set(x_61, 1, x_42); +lean_ctor_set(x_61, 2, x_43); +lean_ctor_set(x_61, 3, x_44); +lean_ctor_set(x_61, 4, x_60); +lean_ctor_set(x_61, 5, x_45); +lean_ctor_set(x_61, 6, x_46); +lean_ctor_set(x_61, 7, x_47); +lean_ctor_set(x_61, 8, x_48); +x_62 = lean_st_ref_set(x_6, x_61); +x_63 = lean_box(0); +lean_ctor_set(x_9, 0, x_63); +return x_9; +} +} +else +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; uint64_t x_76; lean_object* x_77; lean_object* x_78; double x_79; uint8_t x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_64 = lean_ctor_get(x_9, 0); +lean_inc(x_64); +lean_dec(x_9); +x_65 = lean_st_ref_take(x_6); +x_66 = lean_ctor_get(x_65, 4); +lean_inc_ref(x_66); +x_67 = lean_ctor_get(x_65, 0); +lean_inc_ref(x_67); +x_68 = lean_ctor_get(x_65, 1); +lean_inc(x_68); +x_69 = lean_ctor_get(x_65, 2); +lean_inc_ref(x_69); +x_70 = lean_ctor_get(x_65, 3); +lean_inc_ref(x_70); +x_71 = lean_ctor_get(x_65, 5); +lean_inc_ref(x_71); +x_72 = lean_ctor_get(x_65, 6); +lean_inc_ref(x_72); +x_73 = lean_ctor_get(x_65, 7); +lean_inc_ref(x_73); +x_74 = lean_ctor_get(x_65, 8); +lean_inc_ref(x_74); +if (lean_is_exclusive(x_65)) { + lean_ctor_release(x_65, 0); + lean_ctor_release(x_65, 1); + lean_ctor_release(x_65, 2); + lean_ctor_release(x_65, 3); + lean_ctor_release(x_65, 4); + lean_ctor_release(x_65, 5); + lean_ctor_release(x_65, 6); + lean_ctor_release(x_65, 7); + lean_ctor_release(x_65, 8); + x_75 = x_65; +} else { + lean_dec_ref(x_65); + x_75 = lean_box(0); +} +x_76 = lean_ctor_get_uint64(x_66, sizeof(void*)*1); +x_77 = lean_ctor_get(x_66, 0); +lean_inc_ref(x_77); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + x_78 = x_66; +} else { + lean_dec_ref(x_66); + x_78 = lean_box(0); +} +x_79 = l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___closed__0; +x_80 = 0; +x_81 = l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___closed__1; +x_82 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_82, 0, x_1); +lean_ctor_set(x_82, 1, x_81); +lean_ctor_set_float(x_82, sizeof(void*)*2, x_79); +lean_ctor_set_float(x_82, sizeof(void*)*2 + 8, x_79); +lean_ctor_set_uint8(x_82, sizeof(void*)*2 + 16, x_80); +x_83 = l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___closed__2; +x_84 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_84, 0, x_82); +lean_ctor_set(x_84, 1, x_64); +lean_ctor_set(x_84, 2, x_83); +lean_inc(x_8); +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_8); +lean_ctor_set(x_85, 1, x_84); +x_86 = l_Lean_PersistentArray_push___redArg(x_77, x_85); +if (lean_is_scalar(x_78)) { + x_87 = lean_alloc_ctor(0, 1, 8); +} else { + x_87 = x_78; +} +lean_ctor_set(x_87, 0, x_86); +lean_ctor_set_uint64(x_87, sizeof(void*)*1, x_76); +if (lean_is_scalar(x_75)) { + x_88 = lean_alloc_ctor(0, 9, 0); +} else { + x_88 = x_75; +} +lean_ctor_set(x_88, 0, x_67); +lean_ctor_set(x_88, 1, x_68); +lean_ctor_set(x_88, 2, x_69); +lean_ctor_set(x_88, 3, x_70); +lean_ctor_set(x_88, 4, x_87); +lean_ctor_set(x_88, 5, x_71); +lean_ctor_set(x_88, 6, x_72); +lean_ctor_set(x_88, 7, x_73); +lean_ctor_set(x_88, 8, x_74); +x_89 = lean_st_ref_set(x_6, x_88); +x_90 = lean_box(0); +x_91 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_91, 0, x_90); +return x_91; +} +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Failed to apply ", 16, 16); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_rwMatcher___closed__0; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(":", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_rwMatcher___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Meta", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Match", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("debug", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_rwMatcher___closed__6; +x_2 = l_Lean_Meta_rwMatcher___closed__5; +x_3 = l_Lean_Meta_rwMatcher___closed__4; +x_4 = l_Lean_Name_mkStr3(x_3, x_2, x_1); +return x_4; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Not a matcher application:", 26, 26); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_rwMatcher___closed__8; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("When trying to reduce arm ", 26, 26); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_rwMatcher___closed__10; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(", only ", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_rwMatcher___closed__12; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___closed__14() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" equations for ", 15, 15); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_rwMatcher___closed__14; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = l_Lean_Expr_sort___override(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___closed__17() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("PSum", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___closed__18() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("casesOn", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_rwMatcher___closed__18; +x_2 = l_Lean_Meta_rwMatcher___closed__17; +x_3 = l_Lean_Name_mkStr2(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___closed__20() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("PSigma", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_rwMatcher___closed__18; +x_2 = l_Lean_Meta_rwMatcher___closed__20; +x_3 = l_Lean_Name_mkStr2(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_rwMatcher___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_2); +x_11 = l_Lean_Meta_rwMatcher___lam__0(x_1, x_10, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_8); +lean_dec_ref(x_7); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec_ref(x_4); +return x_11; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___lam__1___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Could not un-HEq `", 18, 18); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___lam__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_rwMatcher___lam__1___closed__0; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___lam__1___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("`:", 2, 2); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___lam__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_rwMatcher___lam__1___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___lam__1___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" ", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___lam__1___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_rwMatcher___lam__1___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___lam__1___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Not all hypotheses of `", 23, 23); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___lam__1___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_rwMatcher___lam__1___closed__6; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___lam__1___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("` could be discharged: ", 23, 23); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___lam__1___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_rwMatcher___lam__1___closed__8; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Failed to resolve `", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__0; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("`", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Failed to discharge `", 21, 21); +return x_1; +} +} +static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8(lean_object* x_1, uint8_t x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_12; lean_object* x_13; lean_object* x_18; lean_object* x_20; lean_object* x_22; uint8_t x_24; +x_24 = lean_usize_dec_lt(x_5, x_4); +if (x_24 == 0) +{ +lean_object* x_25; +lean_dec(x_10); +lean_dec_ref(x_9); +lean_dec(x_8); +lean_dec_ref(x_7); +x_25 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_25, 0, x_6); +return x_25; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; lean_object* x_65; +x_26 = lean_array_uget(x_3, x_5); +x_65 = l_Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0___redArg(x_26, x_8); +if (lean_obj_tag(x_65) == 0) +{ +lean_object* x_66; uint8_t x_67; +x_66 = lean_ctor_get(x_65, 0); +lean_inc(x_66); +lean_dec_ref(x_65); +x_67 = lean_unbox(x_66); +lean_dec(x_66); +if (x_67 == 0) +{ +lean_object* x_68; +lean_inc(x_26); +x_68 = l_Lean_MVarId_getType(x_26, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_68) == 0) +{ +lean_object* x_69; uint8_t x_70; +x_69 = lean_ctor_get(x_68, 0); +lean_inc(x_69); +lean_dec_ref(x_68); +lean_inc(x_69); +x_70 = l_Lean_Meta_Simp_isEqnThmHypothesis(x_69); +if (x_70 == 0) +{ +uint8_t x_71; +x_71 = l_Lean_Expr_isEq(x_69); +if (x_71 == 0) +{ +uint8_t x_72; +x_72 = l_Lean_Expr_isHEq(x_69); +lean_dec(x_69); +if (x_72 == 0) +{ +lean_dec(x_26); +x_12 = x_1; +x_13 = lean_box(0); +goto block_17; +} +else +{ +lean_object* x_73; +x_73 = l_Lean_Meta_saveState___redArg(x_8, x_10); +if (lean_obj_tag(x_73) == 0) +{ +lean_object* x_74; lean_object* x_75; +x_74 = lean_ctor_get(x_73, 0); +lean_inc(x_74); +lean_dec_ref(x_73); +lean_inc(x_10); +lean_inc_ref(x_9); +lean_inc(x_8); +lean_inc_ref(x_7); +lean_inc(x_26); +x_75 = l_Lean_MVarId_assumption(x_26, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_75) == 0) +{ +lean_dec(x_74); +lean_dec(x_26); +x_20 = x_75; +goto block_21; +} +else +{ +lean_object* x_76; uint8_t x_77; uint8_t x_89; +x_76 = lean_ctor_get(x_75, 0); +lean_inc(x_76); +x_89 = l_Lean_Exception_isInterrupt(x_76); +if (x_89 == 0) +{ +uint8_t x_90; +x_90 = l_Lean_Exception_isRuntime(x_76); +x_77 = x_90; +goto block_88; +} +else +{ +lean_dec(x_76); +x_77 = x_89; +goto block_88; +} +block_88: +{ +if (x_77 == 0) +{ +lean_object* x_78; +lean_dec_ref(x_75); +x_78 = l_Lean_Meta_SavedState_restore___redArg(x_74, x_8, x_10); +lean_dec(x_74); +if (lean_obj_tag(x_78) == 0) +{ +lean_object* x_79; +lean_dec_ref(x_78); +x_79 = l_Lean_Meta_saveState___redArg(x_8, x_10); +if (lean_obj_tag(x_79) == 0) +{ +lean_object* x_80; lean_object* x_81; +x_80 = lean_ctor_get(x_79, 0); +lean_inc(x_80); +lean_dec_ref(x_79); +lean_inc(x_10); +lean_inc_ref(x_9); +lean_inc(x_8); +lean_inc_ref(x_7); +lean_inc(x_26); +x_81 = l_Lean_MVarId_hrefl(x_26, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_81) == 0) +{ +lean_dec(x_80); +lean_dec(x_26); +x_20 = x_81; +goto block_21; +} +else +{ +lean_object* x_82; uint8_t x_83; +x_82 = lean_ctor_get(x_81, 0); +lean_inc(x_82); +x_83 = l_Lean_Exception_isInterrupt(x_82); +if (x_83 == 0) +{ +uint8_t x_84; +x_84 = l_Lean_Exception_isRuntime(x_82); +x_46 = x_81; +x_47 = lean_box(0); +x_48 = x_80; +x_49 = x_84; +goto block_64; +} +else +{ +lean_dec(x_82); +x_46 = x_81; +x_47 = lean_box(0); +x_48 = x_80; +x_49 = x_83; +goto block_64; +} +} +} +else +{ +uint8_t x_85; +lean_dec(x_26); +lean_dec(x_10); +lean_dec_ref(x_9); +lean_dec(x_8); +lean_dec_ref(x_7); +x_85 = !lean_is_exclusive(x_79); +if (x_85 == 0) +{ +return x_79; +} +else +{ +lean_object* x_86; lean_object* x_87; +x_86 = lean_ctor_get(x_79, 0); +lean_inc(x_86); +lean_dec(x_79); +x_87 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_87, 0, x_86); +return x_87; +} +} +} +else +{ +lean_dec(x_26); +x_20 = x_78; +goto block_21; +} +} +else +{ +lean_dec(x_74); +lean_dec(x_26); +x_20 = x_75; +goto block_21; +} +} +} +} +else +{ +uint8_t x_91; +lean_dec(x_26); +lean_dec(x_10); +lean_dec_ref(x_9); +lean_dec(x_8); +lean_dec_ref(x_7); +x_91 = !lean_is_exclusive(x_73); +if (x_91 == 0) +{ +return x_73; +} +else +{ +lean_object* x_92; lean_object* x_93; +x_92 = lean_ctor_get(x_73, 0); +lean_inc(x_92); +lean_dec(x_73); +x_93 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_93, 0, x_92); +return x_93; +} +} +} +} +else +{ +lean_object* x_94; +lean_dec(x_69); +x_94 = l_Lean_Meta_saveState___redArg(x_8, x_10); +if (lean_obj_tag(x_94) == 0) +{ +lean_object* x_95; lean_object* x_96; +x_95 = lean_ctor_get(x_94, 0); +lean_inc(x_95); +lean_dec_ref(x_94); +lean_inc(x_10); +lean_inc_ref(x_9); +lean_inc(x_8); +lean_inc_ref(x_7); +lean_inc(x_26); +x_96 = l_Lean_MVarId_assumption(x_26, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_96) == 0) +{ +lean_dec(x_95); +lean_dec(x_26); +x_18 = x_96; +goto block_19; +} +else +{ +lean_object* x_97; uint8_t x_98; uint8_t x_110; +x_97 = lean_ctor_get(x_96, 0); +lean_inc(x_97); +x_110 = l_Lean_Exception_isInterrupt(x_97); +if (x_110 == 0) +{ +uint8_t x_111; +x_111 = l_Lean_Exception_isRuntime(x_97); +x_98 = x_111; +goto block_109; +} +else +{ +lean_dec(x_97); +x_98 = x_110; +goto block_109; +} +block_109: +{ +if (x_98 == 0) +{ +lean_object* x_99; +lean_dec_ref(x_96); +x_99 = l_Lean_Meta_SavedState_restore___redArg(x_95, x_8, x_10); +lean_dec(x_95); +if (lean_obj_tag(x_99) == 0) +{ +lean_object* x_100; +lean_dec_ref(x_99); +x_100 = l_Lean_Meta_saveState___redArg(x_8, x_10); +if (lean_obj_tag(x_100) == 0) +{ +lean_object* x_101; lean_object* x_102; +x_101 = lean_ctor_get(x_100, 0); +lean_inc(x_101); +lean_dec_ref(x_100); +lean_inc(x_10); +lean_inc_ref(x_9); +lean_inc(x_8); +lean_inc_ref(x_7); +lean_inc(x_26); +x_102 = l_Lean_MVarId_refl(x_26, x_2, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_102) == 0) +{ +lean_dec(x_101); +lean_dec(x_26); +x_18 = x_102; +goto block_19; +} +else +{ +lean_object* x_103; uint8_t x_104; +x_103 = lean_ctor_get(x_102, 0); +lean_inc(x_103); +x_104 = l_Lean_Exception_isInterrupt(x_103); +if (x_104 == 0) +{ +uint8_t x_105; +x_105 = l_Lean_Exception_isRuntime(x_103); +x_27 = x_101; +x_28 = lean_box(0); +x_29 = x_102; +x_30 = x_105; +goto block_45; +} +else +{ +lean_dec(x_103); +x_27 = x_101; +x_28 = lean_box(0); +x_29 = x_102; +x_30 = x_104; +goto block_45; +} +} +} +else +{ +uint8_t x_106; +lean_dec(x_26); +lean_dec(x_10); +lean_dec_ref(x_9); +lean_dec(x_8); +lean_dec_ref(x_7); +x_106 = !lean_is_exclusive(x_100); +if (x_106 == 0) +{ +return x_100; +} +else +{ +lean_object* x_107; lean_object* x_108; +x_107 = lean_ctor_get(x_100, 0); +lean_inc(x_107); +lean_dec(x_100); +x_108 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_108, 0, x_107); +return x_108; +} +} +} +else +{ +lean_dec(x_26); +x_18 = x_99; +goto block_19; +} +} +else +{ +lean_dec(x_95); +lean_dec(x_26); +x_18 = x_96; +goto block_19; +} +} +} +} +else +{ +uint8_t x_112; +lean_dec(x_26); +lean_dec(x_10); +lean_dec_ref(x_9); +lean_dec(x_8); +lean_dec_ref(x_7); +x_112 = !lean_is_exclusive(x_94); +if (x_112 == 0) +{ +return x_94; +} +else +{ +lean_object* x_113; lean_object* x_114; +x_113 = lean_ctor_get(x_94, 0); +lean_inc(x_113); +lean_dec(x_94); +x_114 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_114, 0, x_113); +return x_114; +} +} +} +} +else +{ +lean_object* x_115; +lean_dec(x_69); +x_115 = l_Lean_Meta_saveState___redArg(x_8, x_10); +if (lean_obj_tag(x_115) == 0) +{ +lean_object* x_116; lean_object* x_117; +x_116 = lean_ctor_get(x_115, 0); +lean_inc(x_116); +lean_dec_ref(x_115); +lean_inc(x_10); +lean_inc_ref(x_9); +lean_inc(x_8); +lean_inc_ref(x_7); +lean_inc(x_26); +x_117 = l_Lean_MVarId_assumption(x_26, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_117) == 0) +{ +lean_dec(x_116); +lean_dec(x_26); +x_22 = x_117; +goto block_23; +} +else +{ +lean_object* x_118; uint8_t x_119; uint8_t x_135; +x_118 = lean_ctor_get(x_117, 0); +lean_inc(x_118); +x_135 = l_Lean_Exception_isInterrupt(x_118); +if (x_135 == 0) +{ +uint8_t x_136; +x_136 = l_Lean_Exception_isRuntime(x_118); +x_119 = x_136; +goto block_134; +} +else +{ +lean_dec(x_118); +x_119 = x_135; +goto block_134; +} +block_134: +{ +if (x_119 == 0) +{ +lean_object* x_120; +lean_dec_ref(x_117); +x_120 = l_Lean_Meta_SavedState_restore___redArg(x_116, x_8, x_10); +lean_dec(x_116); +if (lean_obj_tag(x_120) == 0) +{ +uint8_t x_121; +x_121 = !lean_is_exclusive(x_120); +if (x_121 == 0) +{ +lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_122 = lean_ctor_get(x_120, 0); +lean_dec(x_122); +x_123 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__5; +lean_ctor_set_tag(x_120, 1); +lean_ctor_set(x_120, 0, x_26); +x_124 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_124, 0, x_123); +lean_ctor_set(x_124, 1, x_120); +x_125 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__3; +x_126 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_126, 0, x_124); +lean_ctor_set(x_126, 1, x_125); +x_127 = l_Lean_throwError___at___00Lean_Meta_rwMatcher_spec__6___redArg(x_126, x_7, x_8, x_9, x_10); +x_22 = x_127; +goto block_23; +} +else +{ +lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; +lean_dec(x_120); +x_128 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__5; +x_129 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_129, 0, x_26); +x_130 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_130, 0, x_128); +lean_ctor_set(x_130, 1, x_129); +x_131 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__3; +x_132 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_132, 0, x_130); +lean_ctor_set(x_132, 1, x_131); +x_133 = l_Lean_throwError___at___00Lean_Meta_rwMatcher_spec__6___redArg(x_132, x_7, x_8, x_9, x_10); +x_22 = x_133; +goto block_23; +} +} +else +{ +lean_dec(x_26); +x_22 = x_120; +goto block_23; +} +} +else +{ +lean_dec(x_116); +lean_dec(x_26); +x_22 = x_117; +goto block_23; +} +} +} +} +else +{ +uint8_t x_137; +lean_dec(x_26); +lean_dec(x_10); +lean_dec_ref(x_9); +lean_dec(x_8); +lean_dec_ref(x_7); +x_137 = !lean_is_exclusive(x_115); +if (x_137 == 0) +{ +return x_115; +} +else +{ +lean_object* x_138; lean_object* x_139; +x_138 = lean_ctor_get(x_115, 0); +lean_inc(x_138); +lean_dec(x_115); +x_139 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_139, 0, x_138); +return x_139; +} +} +} +} +else +{ +uint8_t x_140; +lean_dec(x_26); +lean_dec(x_10); +lean_dec_ref(x_9); +lean_dec(x_8); +lean_dec_ref(x_7); +x_140 = !lean_is_exclusive(x_68); +if (x_140 == 0) +{ +return x_68; +} +else +{ +lean_object* x_141; lean_object* x_142; +x_141 = lean_ctor_get(x_68, 0); +lean_inc(x_141); +lean_dec(x_68); +x_142 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_142, 0, x_141); +return x_142; +} +} +} +else +{ +lean_dec(x_26); +x_12 = x_1; +x_13 = lean_box(0); +goto block_17; +} +} +else +{ +uint8_t x_143; +lean_dec(x_26); +lean_dec(x_10); +lean_dec_ref(x_9); +lean_dec(x_8); +lean_dec_ref(x_7); +x_143 = !lean_is_exclusive(x_65); +if (x_143 == 0) +{ +return x_65; +} +else +{ +lean_object* x_144; lean_object* x_145; +x_144 = lean_ctor_get(x_65, 0); +lean_inc(x_144); +lean_dec(x_65); +x_145 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_145, 0, x_144); +return x_145; +} +} +block_45: +{ +if (x_30 == 0) +{ +lean_object* x_31; +lean_dec_ref(x_29); +x_31 = l_Lean_Meta_SavedState_restore___redArg(x_27, x_8, x_10); +lean_dec_ref(x_27); +if (lean_obj_tag(x_31) == 0) +{ +uint8_t x_32; +x_32 = !lean_is_exclusive(x_31); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_33 = lean_ctor_get(x_31, 0); +lean_dec(x_33); +x_34 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__1; +lean_ctor_set_tag(x_31, 1); +lean_ctor_set(x_31, 0, x_26); +x_35 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_31); +x_36 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__3; +x_37 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +x_38 = l_Lean_throwError___at___00Lean_Meta_rwMatcher_spec__6___redArg(x_37, x_7, x_8, x_9, x_10); +x_18 = x_38; +goto block_19; +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +lean_dec(x_31); +x_39 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__1; +x_40 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_40, 0, x_26); +x_41 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +x_42 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__3; +x_43 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +x_44 = l_Lean_throwError___at___00Lean_Meta_rwMatcher_spec__6___redArg(x_43, x_7, x_8, x_9, x_10); +x_18 = x_44; +goto block_19; +} +} +else +{ +lean_dec(x_26); +x_18 = x_31; +goto block_19; +} +} +else +{ +lean_dec_ref(x_27); +lean_dec(x_26); +x_18 = x_29; +goto block_19; +} +} +block_64: +{ +if (x_49 == 0) +{ +lean_object* x_50; +lean_dec_ref(x_46); +x_50 = l_Lean_Meta_SavedState_restore___redArg(x_48, x_8, x_10); +lean_dec_ref(x_48); +if (lean_obj_tag(x_50) == 0) +{ +uint8_t x_51; +x_51 = !lean_is_exclusive(x_50); +if (x_51 == 0) +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_52 = lean_ctor_get(x_50, 0); +lean_dec(x_52); +x_53 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__1; +lean_ctor_set_tag(x_50, 1); +lean_ctor_set(x_50, 0, x_26); +x_54 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set(x_54, 1, x_50); +x_55 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__3; +x_56 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_55); +x_57 = l_Lean_throwError___at___00Lean_Meta_rwMatcher_spec__6___redArg(x_56, x_7, x_8, x_9, x_10); +x_20 = x_57; +goto block_21; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +lean_dec(x_50); +x_58 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__1; +x_59 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_59, 0, x_26); +x_60 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +x_61 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__3; +x_62 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_62, 0, x_60); +lean_ctor_set(x_62, 1, x_61); +x_63 = l_Lean_throwError___at___00Lean_Meta_rwMatcher_spec__6___redArg(x_62, x_7, x_8, x_9, x_10); +x_20 = x_63; +goto block_21; +} +} +else +{ +lean_dec(x_26); +x_20 = x_50; +goto block_21; +} +} +else +{ +lean_dec_ref(x_48); +lean_dec(x_26); +x_20 = x_46; +goto block_21; +} +} +} +block_17: +{ +size_t x_14; size_t x_15; +x_14 = 1; +x_15 = lean_usize_add(x_5, x_14); +x_5 = x_15; +x_6 = x_12; +goto _start; +} +block_19: +{ +if (lean_obj_tag(x_18) == 0) +{ +lean_dec_ref(x_18); +x_12 = x_1; +x_13 = lean_box(0); +goto block_17; +} +else +{ +lean_dec(x_10); +lean_dec_ref(x_9); +lean_dec(x_8); +lean_dec_ref(x_7); +return x_18; +} +} +block_21: +{ +if (lean_obj_tag(x_20) == 0) +{ +lean_dec_ref(x_20); +x_12 = x_1; +x_13 = lean_box(0); +goto block_17; +} +else +{ +lean_dec(x_10); +lean_dec_ref(x_9); +lean_dec(x_8); +lean_dec_ref(x_7); +return x_20; +} +} +block_23: +{ +if (lean_obj_tag(x_22) == 0) +{ +lean_dec_ref(x_22); +x_12 = x_1; +x_13 = lean_box(0); +goto block_17; +} +else +{ +lean_dec(x_10); +lean_dec_ref(x_9); +lean_dec(x_8); +lean_dec_ref(x_7); +return x_22; +} +} +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___lam__1___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___lam__1___closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Left-hand side `", 16, 16); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___lam__1___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_rwMatcher___lam__1___closed__11; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___lam__1___closed__13() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("` of `", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___lam__1___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_rwMatcher___lam__1___closed__13; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___lam__1___closed__15() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("` does not apply to `", 21, 21); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___lam__1___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_rwMatcher___lam__1___closed__15; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___lam__1___closed__17() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("HEq", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___lam__1___closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_rwMatcher___lam__1___closed__17; +x_2 = l_Lean_Name_mkStr1(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___lam__1___closed__19() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Eq", 2, 2); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___lam__1___closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_rwMatcher___lam__1___closed__19; +x_2 = l_Lean_Name_mkStr1(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___lam__1___closed__21() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Type of `", 9, 9); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___lam__1___closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_rwMatcher___lam__1___closed__21; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___lam__1___closed__23() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("` is not an equality", 20, 20); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___lam__1___closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_rwMatcher___lam__1___closed__23; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___lam__1___closed__25() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("eqProof has type", 16, 16); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_rwMatcher___lam__1___closed__26() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_rwMatcher___lam__1___closed__25; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_rwMatcher___lam__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_19; lean_object* x_20; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; lean_object* x_48; uint8_t x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; uint8_t x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; uint8_t x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_104; lean_object* x_105; size_t x_106; uint8_t x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_130; size_t x_131; lean_object* x_132; uint8_t x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_165; +lean_inc(x_10); +lean_inc_ref(x_9); +lean_inc(x_8); +lean_inc_ref(x_7); +lean_inc_ref(x_1); +x_165 = lean_infer_type(x_1, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_165) == 0) +{ +lean_object* x_166; uint8_t x_167; lean_object* x_168; +x_166 = lean_ctor_get(x_165, 0); +lean_inc(x_166); +lean_dec_ref(x_165); +x_167 = 0; +lean_inc(x_10); +lean_inc_ref(x_9); +lean_inc(x_8); +lean_inc_ref(x_7); +x_168 = l_Lean_Meta_forallMetaTelescope(x_166, x_167, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_168) == 0) +{ +lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_207; lean_object* x_208; uint8_t x_209; +x_169 = lean_ctor_get(x_168, 0); +lean_inc(x_169); +lean_dec_ref(x_168); +x_170 = lean_ctor_get(x_169, 1); +lean_inc(x_170); +x_171 = lean_ctor_get(x_169, 0); +lean_inc(x_171); +if (lean_is_exclusive(x_169)) { + lean_ctor_release(x_169, 0); + lean_ctor_release(x_169, 1); + x_172 = x_169; +} else { + lean_dec_ref(x_169); + x_172 = lean_box(0); +} +x_173 = lean_ctor_get(x_170, 1); +lean_inc(x_173); +if (lean_is_exclusive(x_170)) { + lean_ctor_release(x_170, 0); + lean_ctor_release(x_170, 1); + x_174 = x_170; +} else { + lean_dec_ref(x_170); + x_174 = lean_box(0); +} +lean_inc(x_2); +x_207 = l_Lean_isTracingEnabledFor___at___00Lean_Meta_rwMatcher_spec__2___redArg(x_2, x_9); +x_208 = lean_ctor_get(x_207, 0); +lean_inc(x_208); +lean_dec_ref(x_207); +x_209 = lean_unbox(x_208); +lean_dec(x_208); +if (x_209 == 0) +{ +lean_dec(x_2); +x_175 = x_7; +x_176 = x_8; +x_177 = x_9; +x_178 = x_10; +x_179 = lean_box(0); +goto block_206; +} +else +{ +lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; +x_210 = l_Lean_Meta_rwMatcher___lam__1___closed__26; +lean_inc(x_173); +x_211 = l_Lean_indentExpr(x_173); +x_212 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_212, 0, x_210); +lean_ctor_set(x_212, 1, x_211); +x_213 = l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3(x_2, x_212, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_213) == 0) +{ +lean_dec_ref(x_213); +x_175 = x_7; +x_176 = x_8; +x_177 = x_9; +x_178 = x_10; +x_179 = lean_box(0); +goto block_206; +} +else +{ +uint8_t x_214; +lean_dec(x_174); +lean_dec(x_173); +lean_dec(x_172); +lean_dec(x_171); +lean_dec(x_10); +lean_dec_ref(x_9); +lean_dec(x_8); +lean_dec_ref(x_7); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_1); +x_214 = !lean_is_exclusive(x_213); +if (x_214 == 0) +{ +return x_213; +} +else +{ +lean_object* x_215; lean_object* x_216; +x_215 = lean_ctor_get(x_213, 0); +lean_inc(x_215); +lean_dec(x_213); +x_216 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_216, 0, x_215); +return x_216; +} +} +} +block_206: +{ +lean_object* x_180; size_t x_181; size_t x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; uint8_t x_186; +x_180 = l_Lean_mkAppN(x_1, x_171); +x_181 = lean_array_size(x_171); +x_182 = 0; +x_183 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_rwMatcher_spec__4(x_181, x_182, x_171); +x_184 = l_Lean_Meta_rwMatcher___lam__1___closed__18; +x_185 = lean_unsigned_to_nat(4u); +x_186 = l_Lean_Expr_isAppOfArity(x_173, x_184, x_185); +if (x_186 == 0) +{ +lean_object* x_187; lean_object* x_188; uint8_t x_189; +x_187 = l_Lean_Meta_rwMatcher___lam__1___closed__20; +x_188 = lean_unsigned_to_nat(3u); +x_189 = l_Lean_Expr_isAppOfArity(x_173, x_187, x_188); +if (x_189 == 0) +{ +lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; uint8_t x_196; +lean_dec_ref(x_183); +lean_dec_ref(x_180); +lean_dec(x_173); +lean_dec_ref(x_5); +x_190 = l_Lean_Meta_rwMatcher___lam__1___closed__22; +x_191 = l_Lean_MessageData_ofConstName(x_4, x_189); +if (lean_is_scalar(x_174)) { + x_192 = lean_alloc_ctor(7, 2, 0); +} else { + x_192 = x_174; + lean_ctor_set_tag(x_192, 7); +} +lean_ctor_set(x_192, 0, x_190); +lean_ctor_set(x_192, 1, x_191); +x_193 = l_Lean_Meta_rwMatcher___lam__1___closed__24; +if (lean_is_scalar(x_172)) { + x_194 = lean_alloc_ctor(7, 2, 0); +} else { + x_194 = x_172; + lean_ctor_set_tag(x_194, 7); +} +lean_ctor_set(x_194, 0, x_192); +lean_ctor_set(x_194, 1, x_193); +x_195 = l_Lean_throwError___at___00Lean_Meta_rwMatcher_spec__6___redArg(x_194, x_175, x_176, x_177, x_178); +lean_dec(x_178); +lean_dec_ref(x_177); +lean_dec(x_176); +lean_dec_ref(x_175); +x_196 = !lean_is_exclusive(x_195); +if (x_196 == 0) +{ +return x_195; +} +else +{ +lean_object* x_197; lean_object* x_198; +x_197 = lean_ctor_get(x_195, 0); +lean_inc(x_197); +lean_dec(x_195); +x_198 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_198, 0, x_197); +return x_198; +} +} +else +{ +lean_object* x_199; lean_object* x_200; lean_object* x_201; +lean_dec(x_174); +lean_dec(x_172); +x_199 = l_Lean_Expr_appFn_x21(x_173); +x_200 = l_Lean_Expr_appArg_x21(x_199); +lean_dec_ref(x_199); +x_201 = l_Lean_Expr_appArg_x21(x_173); +lean_dec(x_173); +x_130 = x_183; +x_131 = x_182; +x_132 = x_180; +x_133 = x_186; +x_134 = x_200; +x_135 = x_201; +x_136 = x_175; +x_137 = x_176; +x_138 = x_177; +x_139 = x_178; +x_140 = lean_box(0); +goto block_164; +} +} +else +{ +lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; +lean_dec(x_174); +lean_dec(x_172); +x_202 = l_Lean_Expr_appFn_x21(x_173); +x_203 = l_Lean_Expr_appFn_x21(x_202); +lean_dec_ref(x_202); +x_204 = l_Lean_Expr_appArg_x21(x_203); +lean_dec_ref(x_203); +x_205 = l_Lean_Expr_appArg_x21(x_173); +lean_dec(x_173); +x_130 = x_183; +x_131 = x_182; +x_132 = x_180; +x_133 = x_3; +x_134 = x_204; +x_135 = x_205; +x_136 = x_175; +x_137 = x_176; +x_138 = x_177; +x_139 = x_178; +x_140 = lean_box(0); +goto block_164; +} +} +} +else +{ +uint8_t x_217; +lean_dec(x_10); +lean_dec_ref(x_9); +lean_dec(x_8); +lean_dec_ref(x_7); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec_ref(x_1); +x_217 = !lean_is_exclusive(x_168); +if (x_217 == 0) +{ +return x_168; +} +else +{ +lean_object* x_218; lean_object* x_219; +x_218 = lean_ctor_get(x_168, 0); +lean_inc(x_218); +lean_dec(x_168); +x_219 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_219, 0, x_218); +return x_219; +} +} +} +else +{ +uint8_t x_220; +lean_dec(x_10); +lean_dec_ref(x_9); +lean_dec(x_8); +lean_dec_ref(x_7); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec_ref(x_1); +x_220 = !lean_is_exclusive(x_165); +if (x_220 == 0) +{ +return x_165; +} +else +{ +lean_object* x_221; lean_object* x_222; +x_221 = lean_ctor_get(x_165, 0); +lean_inc(x_221); +lean_dec(x_165); +x_222 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_222, 0, x_221); +return x_222; +} +} +block_18: +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_13); +x_16 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_16, 0, x_12); +lean_ctor_set(x_16, 1, x_15); +lean_ctor_set_uint8(x_16, sizeof(void*)*2, x_3); +x_17 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_17, 0, x_16); +return x_17; +} +block_25: +{ +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +lean_dec_ref(x_20); +x_12 = x_19; +x_13 = x_21; +x_14 = lean_box(0); +goto block_18; +} +else +{ +uint8_t x_22; +lean_dec_ref(x_19); +x_22 = !lean_is_exclusive(x_20); +if (x_22 == 0) +{ +return x_20; +} +else +{ +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_20, 0); +lean_inc(x_23); +lean_dec(x_20); +x_24 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_24, 0, x_23); +return x_24; +} +} +} +block_47: +{ +if (x_35 == 0) +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +lean_dec_ref(x_33); +x_36 = l_Lean_Meta_rwMatcher___lam__1___closed__1; +x_37 = l_Lean_MessageData_ofExpr(x_32); +x_38 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +x_39 = l_Lean_Meta_rwMatcher___lam__1___closed__3; +x_40 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set(x_40, 1, x_39); +x_41 = l_Lean_Exception_toMessageData(x_26); +x_42 = l_Lean_indentD(x_41); +x_43 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_43, 0, x_40); +lean_ctor_set(x_43, 1, x_42); +x_44 = l_Lean_Meta_rwMatcher___lam__1___closed__5; +x_45 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +x_46 = l_Lean_throwError___at___00Lean_Meta_rwMatcher_spec__6___redArg(x_45, x_30, x_34, x_29, x_31); +lean_dec(x_31); +lean_dec_ref(x_29); +lean_dec(x_34); +lean_dec_ref(x_30); +x_19 = x_27; +x_20 = x_46; +goto block_25; +} +else +{ +lean_dec(x_34); +lean_dec_ref(x_32); +lean_dec(x_31); +lean_dec_ref(x_30); +lean_dec_ref(x_29); +lean_dec_ref(x_26); +x_19 = x_27; +x_20 = x_33; +goto block_25; +} +} +block_65: +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = l_Lean_instantiateMVars___at___00Lean_Meta_rwMatcher_spec__5___redArg(x_48, x_52); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +lean_dec_ref(x_56); +x_58 = l_Lean_instantiateMVars___at___00Lean_Meta_rwMatcher_spec__5___redArg(x_50, x_52); +if (x_49 == 0) +{ +lean_object* x_59; +lean_dec(x_54); +lean_dec_ref(x_53); +lean_dec(x_52); +lean_dec_ref(x_51); +x_59 = lean_ctor_get(x_58, 0); +lean_inc(x_59); +lean_dec_ref(x_58); +x_12 = x_57; +x_13 = x_59; +x_14 = lean_box(0); +goto block_18; +} +else +{ +lean_object* x_60; lean_object* x_61; +x_60 = lean_ctor_get(x_58, 0); +lean_inc(x_60); +lean_dec_ref(x_58); +lean_inc(x_54); +lean_inc_ref(x_53); +lean_inc(x_52); +lean_inc_ref(x_51); +lean_inc(x_60); +x_61 = l_Lean_Meta_mkEqOfHEq(x_60, x_3, x_51, x_52, x_53, x_54); +if (lean_obj_tag(x_61) == 0) +{ +lean_dec(x_60); +lean_dec(x_54); +lean_dec_ref(x_53); +lean_dec(x_52); +lean_dec_ref(x_51); +x_19 = x_57; +x_20 = x_61; +goto block_25; +} +else +{ +lean_object* x_62; uint8_t x_63; +x_62 = lean_ctor_get(x_61, 0); +lean_inc(x_62); +x_63 = l_Lean_Exception_isInterrupt(x_62); +if (x_63 == 0) +{ +uint8_t x_64; +lean_inc(x_62); +x_64 = l_Lean_Exception_isRuntime(x_62); +x_26 = x_62; +x_27 = x_57; +x_28 = lean_box(0); +x_29 = x_53; +x_30 = x_51; +x_31 = x_54; +x_32 = x_60; +x_33 = x_61; +x_34 = x_52; +x_35 = x_64; +goto block_47; +} +else +{ +x_26 = x_62; +x_27 = x_57; +x_28 = lean_box(0); +x_29 = x_53; +x_30 = x_51; +x_31 = x_54; +x_32 = x_60; +x_33 = x_61; +x_34 = x_52; +x_35 = x_63; +goto block_47; +} +} +} +} +block_90: +{ +uint8_t x_75; +x_75 = l_Array_isEmpty___redArg(x_73); +if (x_75 == 0) +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; uint8_t x_87; +lean_dec_ref(x_71); +lean_dec_ref(x_66); +x_76 = l_Lean_Meta_rwMatcher___lam__1___closed__7; +x_77 = l_Lean_MessageData_ofConstName(x_4, x_75); +x_78 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_77); +x_79 = l_Lean_Meta_rwMatcher___lam__1___closed__9; +x_80 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_80, 0, x_78); +lean_ctor_set(x_80, 1, x_79); +x_81 = lean_array_to_list(x_73); +x_82 = lean_box(0); +x_83 = l_List_mapTR_loop___at___00Lean_Meta_rwMatcher_spec__7(x_81, x_82); +x_84 = l_Lean_MessageData_ofList(x_83); +x_85 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_85, 0, x_80); +lean_ctor_set(x_85, 1, x_84); +x_86 = l_Lean_throwError___at___00Lean_Meta_rwMatcher_spec__6___redArg(x_85, x_67, x_68, x_69, x_72); +lean_dec(x_72); +lean_dec_ref(x_69); +lean_dec(x_68); +lean_dec_ref(x_67); +x_87 = !lean_is_exclusive(x_86); +if (x_87 == 0) +{ +return x_86; +} +else +{ +lean_object* x_88; lean_object* x_89; +x_88 = lean_ctor_get(x_86, 0); +lean_inc(x_88); +lean_dec(x_86); +x_89 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_89, 0, x_88); +return x_89; +} +} +else +{ +lean_dec_ref(x_73); +lean_dec(x_4); +x_48 = x_66; +x_49 = x_70; +x_50 = x_71; +x_51 = x_67; +x_52 = x_68; +x_53 = x_69; +x_54 = x_72; +x_55 = lean_box(0); +goto block_65; +} +} +block_103: +{ +if (lean_obj_tag(x_98) == 0) +{ +lean_object* x_99; +x_99 = lean_ctor_get(x_98, 0); +lean_inc(x_99); +lean_dec_ref(x_98); +x_66 = x_91; +x_67 = x_92; +x_68 = x_93; +x_69 = x_94; +x_70 = x_95; +x_71 = x_97; +x_72 = x_96; +x_73 = x_99; +x_74 = lean_box(0); +goto block_90; +} +else +{ +uint8_t x_100; +lean_dec_ref(x_97); +lean_dec(x_96); +lean_dec_ref(x_94); +lean_dec(x_93); +lean_dec_ref(x_92); +lean_dec_ref(x_91); +lean_dec(x_4); +x_100 = !lean_is_exclusive(x_98); +if (x_100 == 0) +{ +return x_98; +} +else +{ +lean_object* x_101; lean_object* x_102; +x_101 = lean_ctor_get(x_98, 0); +lean_inc(x_101); +lean_dec(x_98); +x_102 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_102, 0, x_101); +return x_102; +} +} +} +block_129: +{ +lean_object* x_114; size_t x_115; lean_object* x_116; +x_114 = lean_box(0); +x_115 = lean_array_size(x_104); +lean_inc(x_112); +lean_inc_ref(x_111); +lean_inc(x_110); +lean_inc_ref(x_109); +x_116 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8(x_114, x_3, x_104, x_115, x_106, x_114, x_109, x_110, x_111, x_112); +if (lean_obj_tag(x_116) == 0) +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; +lean_dec_ref(x_116); +x_117 = lean_unsigned_to_nat(0u); +x_118 = lean_array_get_size(x_104); +x_119 = l_Lean_Meta_rwMatcher___lam__1___closed__10; +x_120 = lean_nat_dec_lt(x_117, x_118); +if (x_120 == 0) +{ +lean_dec_ref(x_104); +x_66 = x_105; +x_67 = x_109; +x_68 = x_110; +x_69 = x_111; +x_70 = x_107; +x_71 = x_108; +x_72 = x_112; +x_73 = x_119; +x_74 = lean_box(0); +goto block_90; +} +else +{ +uint8_t x_121; +x_121 = lean_nat_dec_le(x_118, x_118); +if (x_121 == 0) +{ +if (x_120 == 0) +{ +lean_dec_ref(x_104); +x_66 = x_105; +x_67 = x_109; +x_68 = x_110; +x_69 = x_111; +x_70 = x_107; +x_71 = x_108; +x_72 = x_112; +x_73 = x_119; +x_74 = lean_box(0); +goto block_90; +} +else +{ +size_t x_122; lean_object* x_123; +x_122 = lean_usize_of_nat(x_118); +x_123 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Meta_rwMatcher_spec__9(x_104, x_106, x_122, x_119, x_109, x_110, x_111, x_112); +lean_dec_ref(x_104); +x_91 = x_105; +x_92 = x_109; +x_93 = x_110; +x_94 = x_111; +x_95 = x_107; +x_96 = x_112; +x_97 = x_108; +x_98 = x_123; +goto block_103; +} +} +else +{ +size_t x_124; lean_object* x_125; +x_124 = lean_usize_of_nat(x_118); +x_125 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Meta_rwMatcher_spec__9(x_104, x_106, x_124, x_119, x_109, x_110, x_111, x_112); +lean_dec_ref(x_104); +x_91 = x_105; +x_92 = x_109; +x_93 = x_110; +x_94 = x_111; +x_95 = x_107; +x_96 = x_112; +x_97 = x_108; +x_98 = x_125; +goto block_103; +} +} +} +else +{ +uint8_t x_126; +lean_dec(x_112); +lean_dec_ref(x_111); +lean_dec(x_110); +lean_dec_ref(x_109); +lean_dec_ref(x_108); +lean_dec_ref(x_105); +lean_dec_ref(x_104); +lean_dec(x_4); +x_126 = !lean_is_exclusive(x_116); +if (x_126 == 0) +{ +return x_116; +} +else +{ +lean_object* x_127; lean_object* x_128; +x_127 = lean_ctor_get(x_116, 0); +lean_inc(x_127); +lean_dec(x_116); +x_128 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_128, 0, x_127); +return x_128; +} +} +} +block_164: +{ +lean_object* x_141; +lean_inc(x_139); +lean_inc_ref(x_138); +lean_inc(x_137); +lean_inc_ref(x_136); +lean_inc_ref(x_134); +lean_inc_ref(x_5); +x_141 = l_Lean_Meta_isExprDefEq(x_5, x_134, x_136, x_137, x_138, x_139); +if (lean_obj_tag(x_141) == 0) +{ +lean_object* x_142; uint8_t x_143; +x_142 = lean_ctor_get(x_141, 0); +lean_inc(x_142); +lean_dec_ref(x_141); +x_143 = lean_unbox(x_142); +lean_dec(x_142); +if (x_143 == 0) +{ +lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; uint8_t x_158; +lean_dec_ref(x_135); +lean_dec_ref(x_132); +lean_dec_ref(x_130); +x_144 = l_Lean_Meta_rwMatcher___lam__1___closed__12; +x_145 = l_Lean_MessageData_ofExpr(x_134); +x_146 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_146, 0, x_144); +lean_ctor_set(x_146, 1, x_145); +x_147 = l_Lean_Meta_rwMatcher___lam__1___closed__14; +x_148 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_148, 0, x_146); +lean_ctor_set(x_148, 1, x_147); +x_149 = l_Lean_MessageData_ofConstName(x_4, x_6); +x_150 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_150, 0, x_148); +lean_ctor_set(x_150, 1, x_149); +x_151 = l_Lean_Meta_rwMatcher___lam__1___closed__16; +x_152 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_152, 0, x_150); +lean_ctor_set(x_152, 1, x_151); +x_153 = l_Lean_MessageData_ofExpr(x_5); +x_154 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_154, 0, x_152); +lean_ctor_set(x_154, 1, x_153); +x_155 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__3; +x_156 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_156, 0, x_154); +lean_ctor_set(x_156, 1, x_155); +x_157 = l_Lean_throwError___at___00Lean_Meta_rwMatcher_spec__6___redArg(x_156, x_136, x_137, x_138, x_139); +lean_dec(x_139); +lean_dec_ref(x_138); +lean_dec(x_137); +lean_dec_ref(x_136); +x_158 = !lean_is_exclusive(x_157); +if (x_158 == 0) +{ +return x_157; +} +else +{ +lean_object* x_159; lean_object* x_160; +x_159 = lean_ctor_get(x_157, 0); +lean_inc(x_159); +lean_dec(x_157); +x_160 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_160, 0, x_159); +return x_160; +} +} +else +{ +lean_dec_ref(x_134); +lean_dec_ref(x_5); +x_104 = x_130; +x_105 = x_135; +x_106 = x_131; +x_107 = x_133; +x_108 = x_132; +x_109 = x_136; +x_110 = x_137; +x_111 = x_138; +x_112 = x_139; +x_113 = lean_box(0); +goto block_129; +} +} +else +{ +uint8_t x_161; +lean_dec(x_139); +lean_dec_ref(x_138); +lean_dec(x_137); +lean_dec_ref(x_136); +lean_dec_ref(x_135); +lean_dec_ref(x_134); +lean_dec_ref(x_132); +lean_dec_ref(x_130); +lean_dec_ref(x_5); +lean_dec(x_4); +x_161 = !lean_is_exclusive(x_141); +if (x_161 == 0) +{ +return x_141; +} +else +{ +lean_object* x_162; lean_object* x_163; +x_162 = lean_ctor_get(x_141, 0); +lean_inc(x_162); +lean_dec(x_141); +x_163 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_163, 0, x_162); +return x_163; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_rwMatcher___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; uint8_t x_13; lean_object* x_14; +x_12 = lean_unbox(x_3); +x_13 = lean_unbox(x_6); +x_14 = l_Lean_Meta_rwMatcher___lam__1(x_1, x_2, x_12, x_4, x_5, x_13, x_7, x_8, x_9, x_10); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_rwMatcher___lam__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_2); +x_10 = l_Lean_Meta_rwMatcher___lam__2(x_1, x_9, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec(x_5); +lean_dec_ref(x_4); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_rwMatcher(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_8; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; uint8_t x_46; lean_object* x_47; uint8_t x_52; lean_object* x_53; uint8_t x_58; lean_object* x_233; uint8_t x_234; +x_233 = l_Lean_Meta_rwMatcher___closed__19; +x_234 = l_Lean_Expr_isAppOf(x_2, x_233); +if (x_234 == 0) +{ +lean_object* x_235; uint8_t x_236; +x_235 = l_Lean_Meta_rwMatcher___closed__21; +x_236 = l_Lean_Expr_isAppOf(x_2, x_235); +x_58 = x_236; +goto block_232; +} +else +{ +x_58 = x_234; +goto block_232; +} +block_18: +{ +if (lean_obj_tag(x_8) == 0) +{ +uint8_t x_9; +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_8, 0); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +lean_dec(x_10); +lean_ctor_set(x_8, 0, x_11); +return x_8; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_8, 0); +lean_inc(x_12); +lean_dec(x_8); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +lean_dec(x_12); +x_14 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_14, 0, x_13); +return x_14; +} +} +else +{ +uint8_t x_15; +x_15 = !lean_is_exclusive(x_8); +if (x_15 == 0) +{ +return x_8; +} +else +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_8, 0); +lean_inc(x_16); +lean_dec(x_8); +x_17 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_17, 0, x_16); +return x_17; +} +} +} +block_45: +{ +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; uint8_t x_28; +lean_dec_ref(x_23); +lean_inc(x_19); +x_26 = l_Lean_isTracingEnabledFor___at___00Lean_Meta_rwMatcher_spec__2___redArg(x_19, x_5); +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +lean_dec_ref(x_26); +x_28 = lean_unbox(x_27); +lean_dec(x_27); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; +lean_dec_ref(x_24); +lean_dec(x_22); +lean_dec(x_19); +x_29 = lean_box(0); +x_30 = lean_apply_6(x_21, x_29, x_3, x_4, x_5, x_6, lean_box(0)); +x_8 = x_30; +goto block_18; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_31 = l_Lean_Meta_rwMatcher___closed__1; +x_32 = l_Lean_MessageData_ofConstName(x_22, x_25); +x_33 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +x_34 = l_Lean_Meta_rwMatcher___closed__3; +x_35 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +x_36 = l_Lean_Exception_toMessageData(x_24); +x_37 = l_Lean_indentD(x_36); +x_38 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3(x_19, x_38, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; lean_object* x_41; +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +lean_dec_ref(x_39); +x_41 = lean_apply_6(x_21, x_40, x_3, x_4, x_5, x_6, lean_box(0)); +x_8 = x_41; +goto block_18; +} +else +{ +uint8_t x_42; +lean_dec_ref(x_21); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +x_42 = !lean_is_exclusive(x_39); +if (x_42 == 0) +{ +return x_39; +} +else +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_39, 0); +lean_inc(x_43); +lean_dec(x_39); +x_44 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_44, 0, x_43); +return x_44; +} +} +} +} +else +{ +lean_dec_ref(x_24); +lean_dec(x_22); +lean_dec_ref(x_21); +lean_dec(x_19); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +return x_23; +} +} +block_51: +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_box(0); +x_49 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_49, 0, x_2); +lean_ctor_set(x_49, 1, x_48); +lean_ctor_set_uint8(x_49, sizeof(void*)*2, x_46); +x_50 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_50, 0, x_49); +return x_50; +} +block_57: +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_box(0); +x_55 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_55, 0, x_2); +lean_ctor_set(x_55, 1, x_54); +lean_ctor_set_uint8(x_55, sizeof(void*)*2, x_52); +x_56 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_56, 0, x_55); +return x_56; +} +block_232: +{ +uint8_t x_59; +x_59 = 1; +if (x_58 == 0) +{ +lean_object* x_60; uint8_t x_61; +x_60 = l_Lean_Meta_isMatcherApp___at___00Lean_Meta_rwMatcher_spec__1___redArg(x_2, x_6); +x_61 = !lean_is_exclusive(x_60); +if (x_61 == 0) +{ +lean_object* x_62; uint8_t x_63; +x_62 = lean_ctor_get(x_60, 0); +x_63 = lean_unbox(x_62); +lean_dec(x_62); +if (x_63 == 0) +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; +lean_free_object(x_60); +lean_dec(x_1); +x_64 = l_Lean_Meta_rwMatcher___closed__7; +x_65 = l_Lean_isTracingEnabledFor___at___00Lean_Meta_rwMatcher_spec__2___redArg(x_64, x_5); +x_66 = lean_ctor_get(x_65, 0); +lean_inc(x_66); +lean_dec_ref(x_65); +x_67 = lean_unbox(x_66); +lean_dec(x_66); +if (x_67 == 0) +{ +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +x_46 = x_59; +x_47 = lean_box(0); +goto block_51; +} +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_68 = l_Lean_Meta_rwMatcher___closed__9; +lean_inc_ref(x_2); +x_69 = l_Lean_indentExpr(x_2); +x_70 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set(x_70, 1, x_69); +x_71 = l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3(x_64, x_70, x_3, x_4, x_5, x_6); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +if (lean_obj_tag(x_71) == 0) +{ +lean_dec_ref(x_71); +x_46 = x_59; +x_47 = lean_box(0); +goto block_51; +} +else +{ +uint8_t x_72; +lean_dec_ref(x_2); +x_72 = !lean_is_exclusive(x_71); +if (x_72 == 0) +{ +return x_71; +} +else +{ +lean_object* x_73; lean_object* x_74; +x_73 = lean_ctor_get(x_71, 0); +lean_inc(x_73); +lean_dec(x_71); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_73); +return x_74; +} +} +} +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_75 = l_Lean_Expr_getAppFn(x_2); +x_76 = l_Lean_Expr_constName_x21(x_75); +lean_inc(x_6); +lean_inc_ref(x_5); +lean_inc(x_4); +lean_inc_ref(x_3); +lean_inc(x_76); +x_77 = lean_get_congr_match_equations_for(x_76, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_77) == 0) +{ +lean_object* x_78; lean_object* x_79; uint8_t x_80; +x_78 = lean_ctor_get(x_77, 0); +lean_inc(x_78); +lean_dec_ref(x_77); +x_79 = lean_array_get_size(x_78); +x_80 = lean_nat_dec_lt(x_1, x_79); +if (x_80 == 0) +{ +lean_object* x_81; lean_object* x_82; uint8_t x_83; +lean_dec(x_78); +lean_dec_ref(x_75); +x_81 = l_Lean_Meta_rwMatcher___closed__7; +x_82 = l_Lean_isTracingEnabledFor___at___00Lean_Meta_rwMatcher_spec__2___redArg(x_81, x_5); +x_83 = !lean_is_exclusive(x_82); +if (x_83 == 0) +{ +lean_object* x_84; uint8_t x_85; +x_84 = lean_ctor_get(x_82, 0); +x_85 = lean_unbox(x_84); +lean_dec(x_84); +if (x_85 == 0) +{ +lean_free_object(x_82); +lean_dec(x_76); +lean_free_object(x_60); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec(x_1); +x_52 = x_59; +x_53 = lean_box(0); +goto block_57; +} +else +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_86 = l_Lean_Meta_rwMatcher___closed__11; +x_87 = l_Nat_reprFast(x_1); +lean_ctor_set_tag(x_82, 3); +lean_ctor_set(x_82, 0, x_87); +x_88 = l_Lean_MessageData_ofFormat(x_82); +x_89 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_89, 0, x_86); +lean_ctor_set(x_89, 1, x_88); +x_90 = l_Lean_Meta_rwMatcher___closed__13; +x_91 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_91, 0, x_89); +lean_ctor_set(x_91, 1, x_90); +x_92 = l_Nat_reprFast(x_79); +lean_ctor_set_tag(x_60, 3); +lean_ctor_set(x_60, 0, x_92); +x_93 = l_Lean_MessageData_ofFormat(x_60); +x_94 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_94, 0, x_91); +lean_ctor_set(x_94, 1, x_93); +x_95 = l_Lean_Meta_rwMatcher___closed__15; +x_96 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_96, 0, x_94); +lean_ctor_set(x_96, 1, x_95); +x_97 = l_Lean_MessageData_ofConstName(x_76, x_58); +x_98 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +x_99 = l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3(x_81, x_98, x_3, x_4, x_5, x_6); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +if (lean_obj_tag(x_99) == 0) +{ +lean_dec_ref(x_99); +x_52 = x_59; +x_53 = lean_box(0); +goto block_57; +} +else +{ +uint8_t x_100; +lean_dec_ref(x_2); +x_100 = !lean_is_exclusive(x_99); +if (x_100 == 0) +{ +return x_99; +} +else +{ +lean_object* x_101; lean_object* x_102; +x_101 = lean_ctor_get(x_99, 0); +lean_inc(x_101); +lean_dec(x_99); +x_102 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_102, 0, x_101); +return x_102; +} +} +} +} +else +{ +lean_object* x_103; uint8_t x_104; +x_103 = lean_ctor_get(x_82, 0); +lean_inc(x_103); +lean_dec(x_82); +x_104 = lean_unbox(x_103); +lean_dec(x_103); +if (x_104 == 0) +{ +lean_dec(x_76); +lean_free_object(x_60); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec(x_1); +x_52 = x_59; +x_53 = lean_box(0); +goto block_57; +} +else +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; +x_105 = l_Lean_Meta_rwMatcher___closed__11; +x_106 = l_Nat_reprFast(x_1); +x_107 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_107, 0, x_106); +x_108 = l_Lean_MessageData_ofFormat(x_107); +x_109 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_109, 0, x_105); +lean_ctor_set(x_109, 1, x_108); +x_110 = l_Lean_Meta_rwMatcher___closed__13; +x_111 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_111, 0, x_109); +lean_ctor_set(x_111, 1, x_110); +x_112 = l_Nat_reprFast(x_79); +lean_ctor_set_tag(x_60, 3); +lean_ctor_set(x_60, 0, x_112); +x_113 = l_Lean_MessageData_ofFormat(x_60); +x_114 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_114, 0, x_111); +lean_ctor_set(x_114, 1, x_113); +x_115 = l_Lean_Meta_rwMatcher___closed__15; +x_116 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_116, 0, x_114); +lean_ctor_set(x_116, 1, x_115); +x_117 = l_Lean_MessageData_ofConstName(x_76, x_58); +x_118 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_118, 0, x_116); +lean_ctor_set(x_118, 1, x_117); +x_119 = l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3(x_81, x_118, x_3, x_4, x_5, x_6); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +if (lean_obj_tag(x_119) == 0) +{ +lean_dec_ref(x_119); +x_52 = x_59; +x_53 = lean_box(0); +goto block_57; +} +else +{ +lean_object* x_120; lean_object* x_121; lean_object* x_122; +lean_dec_ref(x_2); +x_120 = lean_ctor_get(x_119, 0); +lean_inc(x_120); +if (lean_is_exclusive(x_119)) { + lean_ctor_release(x_119, 0); + x_121 = x_119; +} else { + lean_dec_ref(x_119); + x_121 = lean_box(0); +} +if (lean_is_scalar(x_121)) { + x_122 = lean_alloc_ctor(1, 1, 0); +} else { + x_122 = x_121; +} +lean_ctor_set(x_122, 0, x_120); +return x_122; +} +} +} +} +else +{ +lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; +lean_dec(x_76); +lean_free_object(x_60); +x_123 = lean_box(0); +x_124 = lean_array_get(x_123, x_78, x_1); +lean_dec(x_1); +lean_dec(x_78); +x_125 = lean_box(x_58); +lean_inc_ref(x_2); +lean_inc(x_124); +x_126 = lean_alloc_closure((void*)(l_Lean_Meta_rwMatcher___lam__0___boxed), 9, 3); +lean_closure_set(x_126, 0, x_124); +lean_closure_set(x_126, 1, x_125); +lean_closure_set(x_126, 2, x_2); +x_127 = l_Lean_Meta_rwMatcher___closed__7; +x_128 = l_Lean_Expr_constLevels_x21(x_75); +lean_dec_ref(x_75); +lean_inc(x_124); +x_129 = l_Lean_mkConst(x_124, x_128); +x_130 = l_Lean_Meta_rwMatcher___closed__16; +x_131 = l_Lean_Expr_getAppNumArgs(x_2); +lean_inc(x_131); +x_132 = lean_mk_array(x_131, x_130); +x_133 = lean_unsigned_to_nat(1u); +x_134 = lean_nat_sub(x_131, x_133); +lean_dec(x_131); +lean_inc_ref(x_2); +x_135 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_2, x_132, x_134); +x_136 = l_Lean_mkAppN(x_129, x_135); +lean_dec_ref(x_135); +x_137 = lean_box(x_59); +x_138 = lean_box(x_58); +lean_inc_ref(x_2); +lean_inc(x_124); +x_139 = lean_alloc_closure((void*)(l_Lean_Meta_rwMatcher___lam__1___boxed), 11, 6); +lean_closure_set(x_139, 0, x_136); +lean_closure_set(x_139, 1, x_127); +lean_closure_set(x_139, 2, x_137); +lean_closure_set(x_139, 3, x_124); +lean_closure_set(x_139, 4, x_2); +lean_closure_set(x_139, 5, x_138); +x_140 = l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___closed__1; +lean_inc(x_6); +lean_inc_ref(x_5); +lean_inc(x_4); +lean_inc_ref(x_3); +x_141 = l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg(x_127, x_126, x_139, x_59, x_140, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_141) == 0) +{ +lean_dec(x_124); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +return x_141; +} +else +{ +lean_object* x_142; lean_object* x_143; lean_object* x_144; uint8_t x_145; +x_142 = lean_ctor_get(x_141, 0); +lean_inc(x_142); +x_143 = lean_box(x_59); +x_144 = lean_alloc_closure((void*)(l_Lean_Meta_rwMatcher___lam__2___boxed), 8, 2); +lean_closure_set(x_144, 0, x_2); +lean_closure_set(x_144, 1, x_143); +x_145 = l_Lean_Exception_isInterrupt(x_142); +if (x_145 == 0) +{ +uint8_t x_146; +lean_inc(x_142); +x_146 = l_Lean_Exception_isRuntime(x_142); +x_19 = x_127; +x_20 = lean_box(0); +x_21 = x_144; +x_22 = x_124; +x_23 = x_141; +x_24 = x_142; +x_25 = x_146; +goto block_45; +} +else +{ +x_19 = x_127; +x_20 = lean_box(0); +x_21 = x_144; +x_22 = x_124; +x_23 = x_141; +x_24 = x_142; +x_25 = x_145; +goto block_45; +} +} +} +} +else +{ +uint8_t x_147; +lean_dec(x_76); +lean_dec_ref(x_75); +lean_free_object(x_60); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +lean_dec(x_1); +x_147 = !lean_is_exclusive(x_77); +if (x_147 == 0) +{ +return x_77; +} +else +{ +lean_object* x_148; lean_object* x_149; +x_148 = lean_ctor_get(x_77, 0); +lean_inc(x_148); +lean_dec(x_77); +x_149 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_149, 0, x_148); +return x_149; +} +} +} +} +else +{ +lean_object* x_150; uint8_t x_151; +x_150 = lean_ctor_get(x_60, 0); +lean_inc(x_150); +lean_dec(x_60); +x_151 = lean_unbox(x_150); +lean_dec(x_150); +if (x_151 == 0) +{ +lean_object* x_152; lean_object* x_153; lean_object* x_154; uint8_t x_155; +lean_dec(x_1); +x_152 = l_Lean_Meta_rwMatcher___closed__7; +x_153 = l_Lean_isTracingEnabledFor___at___00Lean_Meta_rwMatcher_spec__2___redArg(x_152, x_5); +x_154 = lean_ctor_get(x_153, 0); +lean_inc(x_154); +lean_dec_ref(x_153); +x_155 = lean_unbox(x_154); +lean_dec(x_154); +if (x_155 == 0) +{ +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +x_46 = x_59; +x_47 = lean_box(0); +goto block_51; +} +else +{ +lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; +x_156 = l_Lean_Meta_rwMatcher___closed__9; +lean_inc_ref(x_2); +x_157 = l_Lean_indentExpr(x_2); +x_158 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_158, 0, x_156); +lean_ctor_set(x_158, 1, x_157); +x_159 = l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3(x_152, x_158, x_3, x_4, x_5, x_6); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +if (lean_obj_tag(x_159) == 0) +{ +lean_dec_ref(x_159); +x_46 = x_59; +x_47 = lean_box(0); +goto block_51; +} +else +{ +lean_object* x_160; lean_object* x_161; lean_object* x_162; +lean_dec_ref(x_2); +x_160 = lean_ctor_get(x_159, 0); +lean_inc(x_160); +if (lean_is_exclusive(x_159)) { + lean_ctor_release(x_159, 0); + x_161 = x_159; +} else { + lean_dec_ref(x_159); + x_161 = lean_box(0); +} +if (lean_is_scalar(x_161)) { + x_162 = lean_alloc_ctor(1, 1, 0); +} else { + x_162 = x_161; +} +lean_ctor_set(x_162, 0, x_160); +return x_162; +} +} +} +else +{ +lean_object* x_163; lean_object* x_164; lean_object* x_165; +x_163 = l_Lean_Expr_getAppFn(x_2); +x_164 = l_Lean_Expr_constName_x21(x_163); +lean_inc(x_6); +lean_inc_ref(x_5); +lean_inc(x_4); +lean_inc_ref(x_3); +lean_inc(x_164); +x_165 = lean_get_congr_match_equations_for(x_164, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_165) == 0) +{ +lean_object* x_166; lean_object* x_167; uint8_t x_168; +x_166 = lean_ctor_get(x_165, 0); +lean_inc(x_166); +lean_dec_ref(x_165); +x_167 = lean_array_get_size(x_166); +x_168 = lean_nat_dec_lt(x_1, x_167); +if (x_168 == 0) +{ +lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; uint8_t x_173; +lean_dec(x_166); +lean_dec_ref(x_163); +x_169 = l_Lean_Meta_rwMatcher___closed__7; +x_170 = l_Lean_isTracingEnabledFor___at___00Lean_Meta_rwMatcher_spec__2___redArg(x_169, x_5); +x_171 = lean_ctor_get(x_170, 0); +lean_inc(x_171); +if (lean_is_exclusive(x_170)) { + lean_ctor_release(x_170, 0); + x_172 = x_170; +} else { + lean_dec_ref(x_170); + x_172 = lean_box(0); +} +x_173 = lean_unbox(x_171); +lean_dec(x_171); +if (x_173 == 0) +{ +lean_dec(x_172); +lean_dec(x_164); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec(x_1); +x_52 = x_59; +x_53 = lean_box(0); +goto block_57; +} +else +{ +lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; +x_174 = l_Lean_Meta_rwMatcher___closed__11; +x_175 = l_Nat_reprFast(x_1); +if (lean_is_scalar(x_172)) { + x_176 = lean_alloc_ctor(3, 1, 0); +} else { + x_176 = x_172; + lean_ctor_set_tag(x_176, 3); +} +lean_ctor_set(x_176, 0, x_175); +x_177 = l_Lean_MessageData_ofFormat(x_176); +x_178 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_178, 0, x_174); +lean_ctor_set(x_178, 1, x_177); +x_179 = l_Lean_Meta_rwMatcher___closed__13; +x_180 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_180, 0, x_178); +lean_ctor_set(x_180, 1, x_179); +x_181 = l_Nat_reprFast(x_167); +x_182 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_182, 0, x_181); +x_183 = l_Lean_MessageData_ofFormat(x_182); +x_184 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_184, 0, x_180); +lean_ctor_set(x_184, 1, x_183); +x_185 = l_Lean_Meta_rwMatcher___closed__15; +x_186 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_186, 0, x_184); +lean_ctor_set(x_186, 1, x_185); +x_187 = l_Lean_MessageData_ofConstName(x_164, x_58); +x_188 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_188, 0, x_186); +lean_ctor_set(x_188, 1, x_187); +x_189 = l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3(x_169, x_188, x_3, x_4, x_5, x_6); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +if (lean_obj_tag(x_189) == 0) +{ +lean_dec_ref(x_189); +x_52 = x_59; +x_53 = lean_box(0); +goto block_57; +} +else +{ +lean_object* x_190; lean_object* x_191; lean_object* x_192; +lean_dec_ref(x_2); +x_190 = lean_ctor_get(x_189, 0); +lean_inc(x_190); +if (lean_is_exclusive(x_189)) { + lean_ctor_release(x_189, 0); + x_191 = x_189; +} else { + lean_dec_ref(x_189); + x_191 = lean_box(0); +} +if (lean_is_scalar(x_191)) { + x_192 = lean_alloc_ctor(1, 1, 0); +} else { + x_192 = x_191; +} +lean_ctor_set(x_192, 0, x_190); +return x_192; +} +} +} +else +{ +lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; +lean_dec(x_164); +x_193 = lean_box(0); +x_194 = lean_array_get(x_193, x_166, x_1); +lean_dec(x_1); +lean_dec(x_166); +x_195 = lean_box(x_58); +lean_inc_ref(x_2); +lean_inc(x_194); +x_196 = lean_alloc_closure((void*)(l_Lean_Meta_rwMatcher___lam__0___boxed), 9, 3); +lean_closure_set(x_196, 0, x_194); +lean_closure_set(x_196, 1, x_195); +lean_closure_set(x_196, 2, x_2); +x_197 = l_Lean_Meta_rwMatcher___closed__7; +x_198 = l_Lean_Expr_constLevels_x21(x_163); +lean_dec_ref(x_163); +lean_inc(x_194); +x_199 = l_Lean_mkConst(x_194, x_198); +x_200 = l_Lean_Meta_rwMatcher___closed__16; +x_201 = l_Lean_Expr_getAppNumArgs(x_2); +lean_inc(x_201); +x_202 = lean_mk_array(x_201, x_200); +x_203 = lean_unsigned_to_nat(1u); +x_204 = lean_nat_sub(x_201, x_203); +lean_dec(x_201); +lean_inc_ref(x_2); +x_205 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_2, x_202, x_204); +x_206 = l_Lean_mkAppN(x_199, x_205); +lean_dec_ref(x_205); +x_207 = lean_box(x_59); +x_208 = lean_box(x_58); +lean_inc_ref(x_2); +lean_inc(x_194); +x_209 = lean_alloc_closure((void*)(l_Lean_Meta_rwMatcher___lam__1___boxed), 11, 6); +lean_closure_set(x_209, 0, x_206); +lean_closure_set(x_209, 1, x_197); +lean_closure_set(x_209, 2, x_207); +lean_closure_set(x_209, 3, x_194); +lean_closure_set(x_209, 4, x_2); +lean_closure_set(x_209, 5, x_208); +x_210 = l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___closed__1; +lean_inc(x_6); +lean_inc_ref(x_5); +lean_inc(x_4); +lean_inc_ref(x_3); +x_211 = l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg(x_197, x_196, x_209, x_59, x_210, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_211) == 0) +{ +lean_dec(x_194); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +return x_211; +} +else +{ +lean_object* x_212; lean_object* x_213; lean_object* x_214; uint8_t x_215; +x_212 = lean_ctor_get(x_211, 0); +lean_inc(x_212); +x_213 = lean_box(x_59); +x_214 = lean_alloc_closure((void*)(l_Lean_Meta_rwMatcher___lam__2___boxed), 8, 2); +lean_closure_set(x_214, 0, x_2); +lean_closure_set(x_214, 1, x_213); +x_215 = l_Lean_Exception_isInterrupt(x_212); +if (x_215 == 0) +{ +uint8_t x_216; +lean_inc(x_212); +x_216 = l_Lean_Exception_isRuntime(x_212); +x_19 = x_197; +x_20 = lean_box(0); +x_21 = x_214; +x_22 = x_194; +x_23 = x_211; +x_24 = x_212; +x_25 = x_216; +goto block_45; +} +else +{ +x_19 = x_197; +x_20 = lean_box(0); +x_21 = x_214; +x_22 = x_194; +x_23 = x_211; +x_24 = x_212; +x_25 = x_215; +goto block_45; +} +} +} +} +else +{ +lean_object* x_217; lean_object* x_218; lean_object* x_219; +lean_dec(x_164); +lean_dec_ref(x_163); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +lean_dec(x_1); +x_217 = lean_ctor_get(x_165, 0); +lean_inc(x_217); +if (lean_is_exclusive(x_165)) { + lean_ctor_release(x_165, 0); + x_218 = x_165; +} else { + lean_dec_ref(x_165); + x_218 = lean_box(0); +} +if (lean_is_scalar(x_218)) { + x_219 = lean_alloc_ctor(1, 1, 0); +} else { + x_219 = x_218; +} +lean_ctor_set(x_219, 0, x_217); +return x_219; +} +} +} +} +else +{ +lean_object* x_220; +lean_dec(x_1); +x_220 = l___private_Init_While_0__Lean_Loop_forIn_loop___at___00Lean_Meta_rwMatcher_spec__11(x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_220) == 0) +{ +uint8_t x_221; +x_221 = !lean_is_exclusive(x_220); +if (x_221 == 0) +{ +lean_object* x_222; lean_object* x_223; lean_object* x_224; +x_222 = lean_ctor_get(x_220, 0); +x_223 = lean_box(0); +x_224 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_224, 0, x_222); +lean_ctor_set(x_224, 1, x_223); +lean_ctor_set_uint8(x_224, sizeof(void*)*2, x_59); +lean_ctor_set(x_220, 0, x_224); +return x_220; +} +else +{ +lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; +x_225 = lean_ctor_get(x_220, 0); +lean_inc(x_225); +lean_dec(x_220); +x_226 = lean_box(0); +x_227 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_227, 0, x_225); +lean_ctor_set(x_227, 1, x_226); +lean_ctor_set_uint8(x_227, sizeof(void*)*2, x_59); +x_228 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_228, 0, x_227); +return x_228; +} +} +else +{ +uint8_t x_229; +x_229 = !lean_is_exclusive(x_220); +if (x_229 == 0) +{ +return x_220; +} +else +{ +lean_object* x_230; lean_object* x_231; +x_230 = lean_ctor_get(x_220, 0); +lean_inc(x_230); +lean_dec(x_220); +x_231 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_231, 0, x_230); +return x_231; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0___boxed(lean_object* x_1, 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_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_isMatcherApp___at___00Lean_Meta_rwMatcher_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_Meta_isMatcherApp___at___00Lean_Meta_rwMatcher_spec__1(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = l_Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0(x_1, x_2, x_3); +lean_dec(x_3); +x_5 = lean_box(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at___00Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4_spec__18___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; lean_object* x_8; +x_7 = l_Lean_PersistentHashMap_containsAtAux___at___00Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4_spec__18(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_6); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +x_8 = lean_box(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; uint8_t x_6; lean_object* x_7; +x_5 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_6 = l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4(x_1, x_2, x_5, x_4); +lean_dec(x_4); +x_7 = lean_box(x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00Lean_Meta_rwMatcher_spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_instantiateMVars___at___00Lean_Meta_rwMatcher_spec__5(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec(x_3); +lean_dec_ref(x_2); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00Lean_Meta_rwMatcher_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_isTracingEnabledFor___at___00Lean_Meta_rwMatcher_spec__2(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec(x_3); +lean_dec_ref(x_2); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_Meta_rwMatcher_spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_throwError___at___00Lean_Meta_rwMatcher_spec__6(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_5); +x_13 = l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10(x_1, x_2, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10); +return x_13; +} +} +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_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) { +_start: +{ +lean_object* x_8; +x_8 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__15(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12(x_1, x_2, x_3, x_4); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec(x_2); +lean_dec_ref(x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0___redArg(x_1, x_2); +lean_dec(x_2); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00Lean_Meta_rwMatcher_spec__2___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_isTracingEnabledFor___at___00Lean_Meta_rwMatcher_spec__2___redArg(x_1, x_2); +lean_dec_ref(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__15___redArg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__15___redArg(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_isMatcherApp___at___00Lean_Meta_rwMatcher_spec__1___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_Meta_isMatcherApp___at___00Lean_Meta_rwMatcher_spec__1___redArg(x_1, x_2); +lean_dec(x_2); +lean_dec_ref(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0___redArg(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Option_get___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__13___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__13(x_1, x_2); +lean_dec_ref(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Option_get___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__16___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__16(x_1, x_2); +lean_dec_ref(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_Meta_rwMatcher_spec__6___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_throwError___at___00Lean_Meta_rwMatcher_spec__6___redArg(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec(x_3); +lean_dec_ref(x_2); +return x_7; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_rwMatcher_spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +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___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_rwMatcher_spec__4(x_4, x_5, x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at___00Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4_spec__18___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = l_Lean_PersistentHashMap_containsAtAux___at___00Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4_spec__18___redArg(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec_ref(x_1); +x_5 = lean_box(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__14_spec__16___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +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___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__14_spec__16(x_4, x_5, x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_addMessageContextFull___at___00Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3_spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_addMessageContextFull___at___00Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3_spec__4(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec(x_3); +lean_dec_ref(x_2); +return x_7; +} +} +LEAN_EXPORT lean_object* l___private_Init_While_0__Lean_Loop_forIn_loop___at___00Lean_Meta_rwMatcher_spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l___private_Init_While_0__Lean_Loop_forIn_loop___at___00Lean_Meta_rwMatcher_spec__11(x_1, x_2, x_3, x_4, x_5); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; uint8_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_5 = l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4___redArg(x_1, x_4, x_3); +lean_dec(x_3); +x_6 = lean_box(x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00Lean_Meta_rwMatcher_spec__5___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_instantiateMVars___at___00Lean_Meta_rwMatcher_spec__5___redArg(x_1, x_2); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12___redArg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12___redArg(x_1); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Meta_rwMatcher_spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +size_t x_10; size_t x_11; lean_object* x_12; +x_10 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_11 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_12 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Meta_rwMatcher_spec__9(x_1, x_10, x_11, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_8); +lean_dec_ref(x_7); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec_ref(x_1); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_6); +lean_dec_ref(x_5); +lean_dec(x_4); +lean_dec_ref(x_3); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__14(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_8); +lean_dec(x_6); +lean_dec_ref(x_5); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_rwMatcher___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Meta_rwMatcher(x_1, x_2, x_3, x_4, x_5, x_6); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_4); +x_12 = l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg(x_1, x_2, x_3, x_11, x_5, x_6, x_7, x_8, x_9); +return x_12; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; size_t x_13; size_t x_14; lean_object* x_15; +x_12 = lean_unbox(x_2); +x_13 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_14 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_15 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8(x_1, x_12, x_3, x_13, x_14, x_6, x_7, x_8, x_9, x_10); +lean_dec_ref(x_3); +return x_15; +} +} +lean_object* initialize_Lean_Meta_Tactic_Simp_Types(uint8_t builtin); +lean_object* initialize_Lean_Meta_Match_MatcherApp_Transform(uint8_t builtin); +lean_object* initialize_Lean_Meta_Tactic_Assumption(uint8_t builtin); +lean_object* initialize_Lean_Meta_Tactic_Refl(uint8_t builtin); +lean_object* initialize_Lean_Meta_Tactic_Simp_Rewrite(uint8_t builtin); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Meta_Match_Rewrite(uint8_t builtin) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Lean_Meta_Tactic_Simp_Types(builtin); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Match_MatcherApp_Transform(builtin); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Assumption(builtin); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Refl(builtin); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Simp_Rewrite(builtin); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Meta_rwIfWith___closed__0 = _init_l_Lean_Meta_rwIfWith___closed__0(); +lean_mark_persistent(l_Lean_Meta_rwIfWith___closed__0); +l_Lean_Meta_rwIfWith___closed__1 = _init_l_Lean_Meta_rwIfWith___closed__1(); +lean_mark_persistent(l_Lean_Meta_rwIfWith___closed__1); +l_Lean_Meta_rwIfWith___closed__2 = _init_l_Lean_Meta_rwIfWith___closed__2(); +lean_mark_persistent(l_Lean_Meta_rwIfWith___closed__2); +l_Lean_Meta_rwIfWith___closed__3 = _init_l_Lean_Meta_rwIfWith___closed__3(); +lean_mark_persistent(l_Lean_Meta_rwIfWith___closed__3); +l_Lean_Meta_rwIfWith___closed__4 = _init_l_Lean_Meta_rwIfWith___closed__4(); +lean_mark_persistent(l_Lean_Meta_rwIfWith___closed__4); +l_Lean_Meta_rwIfWith___closed__5 = _init_l_Lean_Meta_rwIfWith___closed__5(); +lean_mark_persistent(l_Lean_Meta_rwIfWith___closed__5); +l_Lean_Meta_rwIfWith___closed__6 = _init_l_Lean_Meta_rwIfWith___closed__6(); +lean_mark_persistent(l_Lean_Meta_rwIfWith___closed__6); +l_Lean_Meta_rwIfWith___closed__7 = _init_l_Lean_Meta_rwIfWith___closed__7(); +lean_mark_persistent(l_Lean_Meta_rwIfWith___closed__7); +l_Lean_Meta_rwIfWith___closed__8 = _init_l_Lean_Meta_rwIfWith___closed__8(); +lean_mark_persistent(l_Lean_Meta_rwIfWith___closed__8); +l_Lean_Meta_rwIfWith___closed__9 = _init_l_Lean_Meta_rwIfWith___closed__9(); +lean_mark_persistent(l_Lean_Meta_rwIfWith___closed__9); +l_Lean_Meta_rwIfWith___closed__10 = _init_l_Lean_Meta_rwIfWith___closed__10(); +lean_mark_persistent(l_Lean_Meta_rwIfWith___closed__10); +l_Lean_Meta_rwIfWith___closed__11 = _init_l_Lean_Meta_rwIfWith___closed__11(); +lean_mark_persistent(l_Lean_Meta_rwIfWith___closed__11); +l_Lean_Meta_rwIfWith___closed__12 = _init_l_Lean_Meta_rwIfWith___closed__12(); +lean_mark_persistent(l_Lean_Meta_rwIfWith___closed__12); +l_Lean_Meta_rwIfWith___closed__13 = _init_l_Lean_Meta_rwIfWith___closed__13(); +lean_mark_persistent(l_Lean_Meta_rwIfWith___closed__13); +l_Lean_Meta_rwIfWith___closed__14 = _init_l_Lean_Meta_rwIfWith___closed__14(); +lean_mark_persistent(l_Lean_Meta_rwIfWith___closed__14); +l_Lean_Meta_rwIfWith___closed__15 = _init_l_Lean_Meta_rwIfWith___closed__15(); +lean_mark_persistent(l_Lean_Meta_rwIfWith___closed__15); +l_Lean_Meta_rwIfWith___closed__16 = _init_l_Lean_Meta_rwIfWith___closed__16(); +lean_mark_persistent(l_Lean_Meta_rwIfWith___closed__16); +l_Lean_Meta_rwIfWith___closed__17 = _init_l_Lean_Meta_rwIfWith___closed__17(); +lean_mark_persistent(l_Lean_Meta_rwIfWith___closed__17); +l_Lean_Meta_rwIfWith___closed__18 = _init_l_Lean_Meta_rwIfWith___closed__18(); +lean_mark_persistent(l_Lean_Meta_rwIfWith___closed__18); +l_Lean_Meta_rwIfWith___closed__19 = _init_l_Lean_Meta_rwIfWith___closed__19(); +lean_mark_persistent(l_Lean_Meta_rwIfWith___closed__19); +l_Lean_Meta_rwIfWith___closed__20 = _init_l_Lean_Meta_rwIfWith___closed__20(); +lean_mark_persistent(l_Lean_Meta_rwIfWith___closed__20); +l_Lean_Meta_rwIfWith___closed__21 = _init_l_Lean_Meta_rwIfWith___closed__21(); +lean_mark_persistent(l_Lean_Meta_rwIfWith___closed__21); +l_Lean_Meta_rwIfWith___closed__22 = _init_l_Lean_Meta_rwIfWith___closed__22(); +lean_mark_persistent(l_Lean_Meta_rwIfWith___closed__22); +l_Lean_Meta_rwIfWith___closed__23 = _init_l_Lean_Meta_rwIfWith___closed__23(); +lean_mark_persistent(l_Lean_Meta_rwIfWith___closed__23); +l_Lean_Meta_rwIfWith___closed__24 = _init_l_Lean_Meta_rwIfWith___closed__24(); +lean_mark_persistent(l_Lean_Meta_rwIfWith___closed__24); +l_Lean_Meta_rwIfWith___closed__25 = _init_l_Lean_Meta_rwIfWith___closed__25(); +lean_mark_persistent(l_Lean_Meta_rwIfWith___closed__25); +l_Lean_Meta_rwIfWith___closed__26 = _init_l_Lean_Meta_rwIfWith___closed__26(); +lean_mark_persistent(l_Lean_Meta_rwIfWith___closed__26); +l_Lean_Meta_rwIfWith___closed__27 = _init_l_Lean_Meta_rwIfWith___closed__27(); +lean_mark_persistent(l_Lean_Meta_rwIfWith___closed__27); +l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4___redArg___closed__0 = _init_l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4___redArg___closed__0(); +l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4___redArg___closed__1 = _init_l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Meta_rwMatcher_spec__0_spec__0_spec__4___redArg___closed__1(); +l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___closed__0 = _init_l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___closed__0(); +l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__0 = _init_l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__0(); +lean_mark_persistent(l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__0); +l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__1 = _init_l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__1(); +lean_mark_persistent(l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__1); +l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__2 = _init_l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__2(); +l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__3 = _init_l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__3(); +lean_mark_persistent(l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__3); +l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__4 = _init_l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__4(); +lean_mark_persistent(l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__4); +l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__5 = _init_l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__5(); +l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12___redArg___closed__0 = _init_l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12___redArg___closed__0(); +lean_mark_persistent(l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12___redArg___closed__0); +l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12___redArg___closed__1 = _init_l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12___redArg___closed__1(); +lean_mark_persistent(l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12___redArg___closed__1); +l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12___redArg___closed__2 = _init_l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12___redArg___closed__2(); +lean_mark_persistent(l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10_spec__12___redArg___closed__2); +l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__6 = _init_l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__6(); +lean_mark_persistent(l_Lean_withTraceNode___at___00Lean_Meta_rwMatcher_spec__10___redArg___closed__6); +l_Lean_Meta_rwMatcher___lam__0___closed__0 = _init_l_Lean_Meta_rwMatcher___lam__0___closed__0(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___lam__0___closed__0); +l_Lean_Meta_rwMatcher___lam__0___closed__1 = _init_l_Lean_Meta_rwMatcher___lam__0___closed__1(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___lam__0___closed__1); +l_Lean_Meta_rwMatcher___lam__0___closed__2 = _init_l_Lean_Meta_rwMatcher___lam__0___closed__2(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___lam__0___closed__2); +l_Lean_Meta_rwMatcher___lam__0___closed__3 = _init_l_Lean_Meta_rwMatcher___lam__0___closed__3(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___lam__0___closed__3); +l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___closed__1 = _init_l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___closed__1(); +lean_mark_persistent(l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___closed__1); +l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___closed__2 = _init_l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___closed__2(); +lean_mark_persistent(l_Lean_addTrace___at___00Lean_Meta_rwMatcher_spec__3___closed__2); +l_Lean_Meta_rwMatcher___closed__0 = _init_l_Lean_Meta_rwMatcher___closed__0(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___closed__0); +l_Lean_Meta_rwMatcher___closed__1 = _init_l_Lean_Meta_rwMatcher___closed__1(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___closed__1); +l_Lean_Meta_rwMatcher___closed__2 = _init_l_Lean_Meta_rwMatcher___closed__2(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___closed__2); +l_Lean_Meta_rwMatcher___closed__3 = _init_l_Lean_Meta_rwMatcher___closed__3(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___closed__3); +l_Lean_Meta_rwMatcher___closed__4 = _init_l_Lean_Meta_rwMatcher___closed__4(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___closed__4); +l_Lean_Meta_rwMatcher___closed__5 = _init_l_Lean_Meta_rwMatcher___closed__5(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___closed__5); +l_Lean_Meta_rwMatcher___closed__6 = _init_l_Lean_Meta_rwMatcher___closed__6(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___closed__6); +l_Lean_Meta_rwMatcher___closed__7 = _init_l_Lean_Meta_rwMatcher___closed__7(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___closed__7); +l_Lean_Meta_rwMatcher___closed__8 = _init_l_Lean_Meta_rwMatcher___closed__8(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___closed__8); +l_Lean_Meta_rwMatcher___closed__9 = _init_l_Lean_Meta_rwMatcher___closed__9(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___closed__9); +l_Lean_Meta_rwMatcher___closed__10 = _init_l_Lean_Meta_rwMatcher___closed__10(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___closed__10); +l_Lean_Meta_rwMatcher___closed__11 = _init_l_Lean_Meta_rwMatcher___closed__11(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___closed__11); +l_Lean_Meta_rwMatcher___closed__12 = _init_l_Lean_Meta_rwMatcher___closed__12(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___closed__12); +l_Lean_Meta_rwMatcher___closed__13 = _init_l_Lean_Meta_rwMatcher___closed__13(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___closed__13); +l_Lean_Meta_rwMatcher___closed__14 = _init_l_Lean_Meta_rwMatcher___closed__14(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___closed__14); +l_Lean_Meta_rwMatcher___closed__15 = _init_l_Lean_Meta_rwMatcher___closed__15(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___closed__15); +l_Lean_Meta_rwMatcher___closed__16 = _init_l_Lean_Meta_rwMatcher___closed__16(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___closed__16); +l_Lean_Meta_rwMatcher___closed__17 = _init_l_Lean_Meta_rwMatcher___closed__17(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___closed__17); +l_Lean_Meta_rwMatcher___closed__18 = _init_l_Lean_Meta_rwMatcher___closed__18(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___closed__18); +l_Lean_Meta_rwMatcher___closed__19 = _init_l_Lean_Meta_rwMatcher___closed__19(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___closed__19); +l_Lean_Meta_rwMatcher___closed__20 = _init_l_Lean_Meta_rwMatcher___closed__20(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___closed__20); +l_Lean_Meta_rwMatcher___closed__21 = _init_l_Lean_Meta_rwMatcher___closed__21(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___closed__21); +l_Lean_Meta_rwMatcher___lam__1___closed__0 = _init_l_Lean_Meta_rwMatcher___lam__1___closed__0(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___lam__1___closed__0); +l_Lean_Meta_rwMatcher___lam__1___closed__1 = _init_l_Lean_Meta_rwMatcher___lam__1___closed__1(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___lam__1___closed__1); +l_Lean_Meta_rwMatcher___lam__1___closed__2 = _init_l_Lean_Meta_rwMatcher___lam__1___closed__2(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___lam__1___closed__2); +l_Lean_Meta_rwMatcher___lam__1___closed__3 = _init_l_Lean_Meta_rwMatcher___lam__1___closed__3(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___lam__1___closed__3); +l_Lean_Meta_rwMatcher___lam__1___closed__4 = _init_l_Lean_Meta_rwMatcher___lam__1___closed__4(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___lam__1___closed__4); +l_Lean_Meta_rwMatcher___lam__1___closed__5 = _init_l_Lean_Meta_rwMatcher___lam__1___closed__5(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___lam__1___closed__5); +l_Lean_Meta_rwMatcher___lam__1___closed__6 = _init_l_Lean_Meta_rwMatcher___lam__1___closed__6(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___lam__1___closed__6); +l_Lean_Meta_rwMatcher___lam__1___closed__7 = _init_l_Lean_Meta_rwMatcher___lam__1___closed__7(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___lam__1___closed__7); +l_Lean_Meta_rwMatcher___lam__1___closed__8 = _init_l_Lean_Meta_rwMatcher___lam__1___closed__8(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___lam__1___closed__8); +l_Lean_Meta_rwMatcher___lam__1___closed__9 = _init_l_Lean_Meta_rwMatcher___lam__1___closed__9(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___lam__1___closed__9); +l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__0 = _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__0(); +lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__0); +l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__1 = _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__1(); +lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__1); +l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__2 = _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__2(); +lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__2); +l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__3 = _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__3(); +lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__3); +l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__4 = _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__4(); +lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__4); +l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__5 = _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__5(); +lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_rwMatcher_spec__8___closed__5); +l_Lean_Meta_rwMatcher___lam__1___closed__10 = _init_l_Lean_Meta_rwMatcher___lam__1___closed__10(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___lam__1___closed__10); +l_Lean_Meta_rwMatcher___lam__1___closed__11 = _init_l_Lean_Meta_rwMatcher___lam__1___closed__11(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___lam__1___closed__11); +l_Lean_Meta_rwMatcher___lam__1___closed__12 = _init_l_Lean_Meta_rwMatcher___lam__1___closed__12(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___lam__1___closed__12); +l_Lean_Meta_rwMatcher___lam__1___closed__13 = _init_l_Lean_Meta_rwMatcher___lam__1___closed__13(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___lam__1___closed__13); +l_Lean_Meta_rwMatcher___lam__1___closed__14 = _init_l_Lean_Meta_rwMatcher___lam__1___closed__14(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___lam__1___closed__14); +l_Lean_Meta_rwMatcher___lam__1___closed__15 = _init_l_Lean_Meta_rwMatcher___lam__1___closed__15(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___lam__1___closed__15); +l_Lean_Meta_rwMatcher___lam__1___closed__16 = _init_l_Lean_Meta_rwMatcher___lam__1___closed__16(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___lam__1___closed__16); +l_Lean_Meta_rwMatcher___lam__1___closed__17 = _init_l_Lean_Meta_rwMatcher___lam__1___closed__17(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___lam__1___closed__17); +l_Lean_Meta_rwMatcher___lam__1___closed__18 = _init_l_Lean_Meta_rwMatcher___lam__1___closed__18(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___lam__1___closed__18); +l_Lean_Meta_rwMatcher___lam__1___closed__19 = _init_l_Lean_Meta_rwMatcher___lam__1___closed__19(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___lam__1___closed__19); +l_Lean_Meta_rwMatcher___lam__1___closed__20 = _init_l_Lean_Meta_rwMatcher___lam__1___closed__20(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___lam__1___closed__20); +l_Lean_Meta_rwMatcher___lam__1___closed__21 = _init_l_Lean_Meta_rwMatcher___lam__1___closed__21(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___lam__1___closed__21); +l_Lean_Meta_rwMatcher___lam__1___closed__22 = _init_l_Lean_Meta_rwMatcher___lam__1___closed__22(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___lam__1___closed__22); +l_Lean_Meta_rwMatcher___lam__1___closed__23 = _init_l_Lean_Meta_rwMatcher___lam__1___closed__23(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___lam__1___closed__23); +l_Lean_Meta_rwMatcher___lam__1___closed__24 = _init_l_Lean_Meta_rwMatcher___lam__1___closed__24(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___lam__1___closed__24); +l_Lean_Meta_rwMatcher___lam__1___closed__25 = _init_l_Lean_Meta_rwMatcher___lam__1___closed__25(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___lam__1___closed__25); +l_Lean_Meta_rwMatcher___lam__1___closed__26 = _init_l_Lean_Meta_rwMatcher___lam__1___closed__26(); +lean_mark_persistent(l_Lean_Meta_rwMatcher___lam__1___closed__26); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Meta/Tactic/FunInd.c b/stage0/stdlib/Lean/Meta/Tactic/FunInd.c index defcfa4219da..ab4e0cf44977 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/FunInd.c +++ b/stage0/stdlib/Lean/Meta/Tactic/FunInd.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.Tactic.FunInd -// Imports: public import Lean.Meta.Tactic.Simp.Types import Lean.Meta.Match.MatcherApp.Transform import Lean.Meta.Injective import Lean.Meta.ArgsPacker import Lean.Elab.PreDefinition.WF.Eqns import Lean.Elab.PreDefinition.Structural.Eqns import Lean.Elab.PreDefinition.Structural.FindRecArg import Lean.Meta.Tactic.ElimInfo import Lean.Meta.Tactic.FunIndInfo import Lean.Data.Array import Lean.Meta.Tactic.Simp.Rewrite import Lean.Meta.Tactic.Refl import Lean.Meta.Tactic.Replace +// Imports: public import Lean.Meta.Tactic.Simp.Types import Lean.Meta.Match.MatcherApp.Transform import Lean.Meta.Match.Rewrite import Lean.Meta.Injective import Lean.Meta.ArgsPacker import Lean.Elab.PreDefinition.WF.Eqns import Lean.Elab.PreDefinition.Structural.Eqns import Lean.Elab.PreDefinition.Structural.FindRecArg import Lean.Meta.Tactic.ElimInfo import Lean.Meta.Tactic.FunIndInfo import Lean.Data.Array import Lean.Meta.Tactic.Simp.Rewrite import Lean.Meta.Tactic.Replace #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -16,13 +16,13 @@ extern "C" { LEAN_EXPORT lean_object* l_Lean_Elab_FixedParamPerm_forallTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__20___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__19___redArg___lam__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_inProdLambdaLastArg___closed__0; -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___closed__16; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_PersistentArray_forIn___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanupAfter_allHeqToEq_spec__0_spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Structural_Positions_mapMwith___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__13___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); lean_object* l_Lean_Expr_fvarId_x3f(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__initFn___closed__16_00___x40_Lean_Meta_Tactic_FunInd_922168289____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__1___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkEqHEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases_spec__3___closed__0; @@ -30,12 +30,12 @@ static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_ static lean_object* l_panic___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__10___closed__0; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwMData___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_registerReservedNameAction(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20_spec__23___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___00Lean_Meta_withIncRecDepth___at___00__private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit___at___00Lean_Meta_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__1_spec__2_spec__9_spec__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13_spec__32___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Elab_Structural_Positions_groupAndSort___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__6_spec__7___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInduction___closed__2; -LEAN_EXPORT lean_object* l_Lean_Meta_isMatcherApp___at___00Lean_Tactic_FunInd_rwMatcher_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__1___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_inferArgumentTypesN(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -43,14 +43,9 @@ LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_Matche static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__29___redArg___lam__3___closed__13; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__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___private_Lean_Meta_Tactic_FunInd_0__initFn_00___x40_Lean_Meta_Tactic_FunInd_922168289____hygCtx___hyg_2____boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___closed__14; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18___redArg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6___redArg(lean_object*, size_t, size_t, lean_object*, lean_object*); static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg___closed__8; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__initFn___closed__19_00___x40_Lean_Meta_Tactic_FunInd_922168289____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l_panic___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Meta_reduceRecMatcher_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_isEqvAux___at___00Lean_Elab_Structural_Positions_groupAndSort___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__6_spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00__private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0_spec__6_spec__9(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__5(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -61,15 +56,13 @@ LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_match LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_refinedArguments_go_spec__0(lean_object*, lean_object*, lean_object*, 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_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__8___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__29___redArg___lam__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__25_spec__27_spec__41___lam__0___closed__2; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__20___closed__9; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M_tell___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__20; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Structural_Positions_groupAndSort___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__6_spec__9_spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___closed__8; -LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__2___boxed(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_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___redArg___lam__3___closed__0; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_projectMutualInduct_spec__1___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -77,10 +70,13 @@ static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_bu lean_object* l_List_lengthTR___redArg(lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_setNaryFunIndInfo_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_unpackMutualInduction_doRealize___closed__1; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2_spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg___closed__9; +LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13___redArg___closed__4; LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__10___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12___closed__1; lean_object* l_Lean_Core_instMonadCoreM___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__10___closed__3; @@ -114,10 +110,10 @@ static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknow LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__21___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00Lean_mkConstWithLevelParams___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__2_spec__3_spec__5_spec__11___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase___lam__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5___lam__3(lean_object*, lean_object*, lean_object*, uint8_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static size_t l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7___redArg___closed__1; LEAN_EXPORT lean_object* l___private_Init_Data_Array_QSort_Basic_0__Array_qsort_sort___at___00Lean_Elab_Structural_Positions_groupAndSort___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__6_spec__11___redArg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M_branch___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__2___closed__0; @@ -129,14 +125,13 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___00__private_Lean_Meta_Ta static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__4___closed__0; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__21___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__14___redArg___lam__0___boxed(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_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__8___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_unpackMutualInduction_doRealize___lam__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInduction___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___closed__5; static lean_object* l_Lean_withExporting___at___00Lean_withoutExporting___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__14_spec__39___redArg___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__21___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapFinIdxM_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__5___redArg___closed__5; -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__0___closed__0; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00__private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit_visitForall___at___00__private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit___at___00Lean_Meta_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__1_spec__2_spec__5_spec__6___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withErasedFVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__3___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -151,6 +146,7 @@ LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___00Lean_throwUnknownIdentifie static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__4___closed__10; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_unpackMutualInduction_doRealize___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__5___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7___redArg(lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00__private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0_spec__6_spec__10___redArg(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -176,7 +172,6 @@ lean_object* l_Lean_ConstantInfo_type(lean_object*); static lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__25_spec__27_spec__41___lam__0___closed__7; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__6___closed__14; lean_object* lean_whnf(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__10(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Environment_header(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_projectMutualInduct_spec__1(uint8_t, 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*); @@ -199,6 +194,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Transform_0__Lean_Meta_transformW LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_setNaryFunIndInfo(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_levelParams(lean_object*); +static lean_object* l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__17; LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_indentD(lean_object*); double lean_float_div(double, double); @@ -214,12 +210,10 @@ uint8_t l_Lean_Exception_isInterrupt(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M_ask___redArg(lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00__private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0_spec__3_spec__4(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__6___closed__7; -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__23; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00Lean_mkConstWithLevelParams___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__2_spec__3_spec__5_spec__11_spec__14_spec__16(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at___00Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3_spec__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwIfWith___closed__7; lean_object* l_Lean_mkAppB(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__3(lean_object*, size_t, size_t, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_refinedArguments___closed__0; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInduction___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -240,15 +234,14 @@ LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_Matche LEAN_EXPORT lean_object* l_Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__11(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withIncRecDepth___at___00__private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit___at___00Lean_Meta_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__1_spec__2_spec__9___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_rwMatcher(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__17___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6___redArg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__9(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural(uint8_t, 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_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t lean_uint64_lor(uint64_t, uint64_t); LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00Lean_mkConstWithLevelParams___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__2_spec__3_spec__5_spec__11_spec__14_spec__15_spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Meta_forallMetaTelescope(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13_spec__33___redArg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_inProdLambdaLastArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -271,6 +264,7 @@ static lean_object* l_panic___at___00Lean_Meta_MatcherApp_transform___at___00__p LEAN_EXPORT lean_object* l_Lean_getConstVal___at___00Lean_mkConstWithLevelParams___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__2_spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isApp(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___closed__2; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__4; LEAN_EXPORT uint8_t l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__25_spec__27_spec__41___lam__0(uint8_t, uint8_t, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__5(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -280,7 +274,6 @@ LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Uns lean_object* l_Lean_Environment_findConstVal_x3f(lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_withExporting___at___00Lean_withoutExporting___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__14_spec__39___redArg___closed__3; static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__29___redArg___lam__3___closed__8; -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__21; LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__24___redArg___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withExporting___at___00Lean_withoutExporting___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__14_spec__39___redArg___lam__0(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_sort___override(lean_object*); @@ -296,6 +289,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunI LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_refinedArguments_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_unpackMutualInduction_doRealize___lam__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_array_push(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__4___redArg(lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__29___redArg___lam__5___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__1___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -305,8 +299,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunI LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14___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_Tactic_FunInd_rwIfWith___closed__18; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__2___closed__0; size_t lean_usize_mul(size_t, size_t); static lean_object* l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__8; @@ -315,6 +307,7 @@ static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_de static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M2_run___redArg___closed__0; LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__28; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit_visitPost___at___00__private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit___at___00Lean_Meta_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__1_spec__2_spec__4(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -328,13 +321,13 @@ extern lean_object* l_Lean_unknownIdentifierMessageTag; LEAN_EXPORT lean_object* l___private_Std_Data_DHashMap_Internal_Defs_0__Std_DHashMap_Internal_Raw_u2080_expand_go___at___00Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00__private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0_spec__6_spec__10_spec__11___redArg(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_isFunCasesName___closed__1; uint8_t lean_usize_dec_eq(size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00Lean_Tactic_FunInd_rwMatcher_spec__3___redArg(lean_object*, lean_object*); static lean_object* l_Lean_Tactic_FunInd_isFunInductName___closed__0; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__5___boxed(lean_object**); lean_object* l_Lean_KVMap_find(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_contains___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__14(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Structural_Positions_groupAndSort___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__6___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__9___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M2_branch___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13_spec__35___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -347,47 +340,47 @@ LEAN_EXPORT lean_object* l_Array_filterMapM___at___00__private_Lean_Meta_Tactic_ LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_withLetDecls_go_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ST_Prim_Ref_get___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__4_spec__4_spec__7_spec__18_spec__28_spec__31___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_Tactic_FunInd_rwMatcher___lam__0(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2___closed__0; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanupAfter_allHeqToEq_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__29___redArg___lam__3___closed__7; lean_object* l_Lean_Expr_mdata___override(lean_object*, lean_object*); lean_object* l___private_Lean_Log_0__Lean_MessageData_appendDescriptionWidgetIfNamed(lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__1_spec__7(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_panic___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__27___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_mapTR_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__4(lean_object*, lean_object*); static lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__5___closed__0; +static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__26; lean_object* l_Lean_Elab_Structural_getRecArgInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_rwMatcher___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__19___redArg___lam__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__20___closed__0; LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__19_spec__24___boxed(lean_object**); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__10(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_setNaryFunIndInfo_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanupAfter_cleanupAfter_x3f_spec__1(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_PersistentArray_forInAux___at___00Lean_PersistentArray_forIn___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanupAfter_cleanupAfter_x3f_spec__1_spec__1_spec__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_filterMapM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__7___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t l_Lean_ExprStructEq_hash(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___closed__4; static lean_object* l_Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12___closed__5; LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_isMatcherApp___at___00Lean_Tactic_FunInd_rwMatcher_spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00__private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0_spec__3___redArg(lean_object*, lean_object*); lean_object* l_Lean_Meta_mkExpectedTypeHint(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_proj___override(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__16; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___closed__1; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__initFn___closed__8_00___x40_Lean_Meta_Tactic_FunInd_922168289____hygCtx___hyg_2_; +static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___closed__0; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__17___closed__4; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_initFn___lam__0___closed__0_00___x40_Lean_Meta_Tactic_FunInd_4256658486____hygCtx___hyg_2_; +static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__29; LEAN_EXPORT lean_object* l___private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit___at___00Lean_Meta_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__1_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_array_fset(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_isMatcherApp___at___00Lean_Tactic_FunInd_rwMatcher_spec__1___redArg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__16; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__0___closed__1; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__4___closed__5; lean_object* l_Lean_Syntax_getTailPos_x3f(lean_object*, uint8_t); @@ -402,12 +395,12 @@ uint8_t l_Lean_instBEqMessageSeverity_beq(uint8_t, uint8_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M_exec___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__24___redArg___lam__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at___00Lean_Core_withIncRecDepth___at___00__private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0_spec__5_spec__7___redArg___closed__6; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20_spec__23___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M_localMapM___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_isFunInductName___boxed(lean_object*, lean_object*); static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__16___lam__2___closed__3; LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__24___redArg___lam__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_refinedArguments___closed__1; -static lean_object* l_Lean_Tactic_FunInd_rwIfWith___closed__22; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__initFn___closed__9_00___x40_Lean_Meta_Tactic_FunInd_922168289____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l_Lean_Meta_withErasedFVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__3(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*); @@ -436,9 +429,9 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunI LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__4_spec__6___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___00Lean_Core_withIncRecDepth___at___00__private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0_spec__5_spec__7(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__9___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_MVarId_refl(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__9; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_setNaryFunIndInfo___closed__4; +static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__24; +static lean_object* l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__20; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_initFn___closed__0_00___x40_Lean_Meta_Tactic_FunInd_4256658486____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__8(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Nat_blt___boxed(lean_object*, lean_object*); @@ -453,13 +446,13 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withErasedFVars___at___00__private_Lean_Met LEAN_EXPORT lean_object* l_Lean_withoutExporting___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__8___closed__1; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__9; -static lean_object* l_Lean_Tactic_FunInd_rwIfWith___closed__17; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__11___closed__1; lean_object* lean_array_fget_borrowed(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__4___closed__3; lean_object* l_Lean_Meta_PProdN_projM(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__10___closed__0; +LEAN_EXPORT lean_object* l_List_mapTR_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__7(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__21(uint8_t, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__4_spec__4_spec__7_spec__18___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases_spec__0(size_t, size_t, lean_object*); @@ -468,6 +461,7 @@ uint8_t l_Lean_Expr_isAppOf(lean_object*, lean_object*); static lean_object* l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__5; lean_object* l_Lean_Meta_PProdN_packLambdas___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__11___closed__4; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__14___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_withLetDecls_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13_spec__33___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__13; @@ -480,7 +474,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at__ LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_refinedArguments_go_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_unpackMutualInduction_doRealize___closed__3; -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___closed__7; lean_object* lean_io_get_num_heartbeats(); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__17(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -505,17 +498,16 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalD LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInduction___lam__0(uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_contains___at___00Lean_Meta_withErasedFVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__3_spec__3___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at___00Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3_spec__12___redArg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase___lam__2___closed__0; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__4(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_instInhabitedPersistentArrayNode_default(lean_object*); lean_object* l_Lean_MessageData_note(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M_branch___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Structural_Positions_groupAndSort___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__6_spec__9_spec__12___redArg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7_spec__14___redArg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkApp4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__10(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__0___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_withLetDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInduction___lam__2___closed__1; static lean_object* l_Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1___closed__2; @@ -537,25 +529,24 @@ LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_m LEAN_EXPORT lean_object* l___private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit_visitLet___at___00__private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit___at___00Lean_Meta_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__1_spec__2_spec__7___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__5___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Tactic_FunInd_rwMatcher_spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__16___lam__0(lean_object*, lean_object*); static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__7(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_InductiveVal_numTypeFormers(lean_object*); -lean_object* l_Lean_Exception_toMessageData(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_PersistentArray_forInAux___at___00Lean_PersistentArray_forIn___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanupAfter_allHeqToEq_spec__0_spec__0_spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00Lean_mkConstWithLevelParams___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__2_spec__3_spec__5_spec__11_spec__14_spec__15_spec__16___redArg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5___closed__0; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M_run___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20_spec__23(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase___lam__0___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_FixedParamPerm_forallTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__20___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withErasedFVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12___closed__11; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstVal___at___00Lean_mkConstWithLevelParams___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases_spec__2_spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -567,10 +558,8 @@ LEAN_EXPORT lean_object* l_panic___at___00Lean_getConstInfoRec___at___00__privat LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_PersistentArray_forInAux___at___00Lean_PersistentArray_forIn___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanupAfter_allHeqToEq_spec__0_spec__0_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_withLetDecls___redArg(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___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__0___closed__2; static lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__5___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwIfWith___closed__21; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M_tell___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Structural_Positions_groupAndSort___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__6_spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__29___redArg___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -579,20 +568,20 @@ LEAN_EXPORT lean_object* l_Lean_Meta_transform___at___00__private_Lean_Meta_Tact LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__16(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__29___redArg___lam__3___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__12(lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_lambdaBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__9_spec__15___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M2_branch(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7_spec__15(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_utf8_byte_size(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___boxed(lean_object**); -static lean_object* l_Lean_Tactic_FunInd_rwIfWith___closed__12; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__21(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_back_x21___redArg(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_mkLambdaFVarsMasked___closed__0; static lean_object* l_Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12___lam__2___closed__2; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__6___closed__8; +static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__33; lean_object* l_Lean_Level_ofNat(lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__4___redArg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_appArg_x21(lean_object*); @@ -604,15 +593,14 @@ lean_object* l_ReaderT_instMonad___redArg(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00Lean_mkConstWithLevelParams___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__2_spec__3_spec__5_spec__11_spec__14___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12___boxed__const__1; static lean_object* l_panic___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__26___closed__0; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15___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_getConstInfoInduct___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__2___closed__0; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_PersistentArray_forIn___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanupAfter_allHeqToEq_spec__0_spec__1_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*); extern lean_object* l_Lean_Elab_WF_eqnInfoExt; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3___redArg(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__4(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M2_run___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_index(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_inLastArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -622,12 +610,9 @@ static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_de LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_projectMutualInduct_spec__1___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__4_spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Tactic_FunInd_isFunInductName___closed__1; -static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M_run___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Structural_Positions_groupAndSort___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__6_spec__9___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M_branch___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__8(size_t, size_t, lean_object*); static lean_object* l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_unpackMutualInduction_doRealize_spec__0___closed__9; LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*); @@ -636,13 +621,13 @@ uint8_t l_Lean_Expr_hasMVar(lean_object*); LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__9___redArg___lam__0(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Structural_IndGroupInst_brecOn(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__16___lam__2___closed__5; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__8___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__7___boxed(lean_object**); uint8_t l_Lean_Expr_isSemiOutParam(lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__24___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases_spec__1___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13_spec__33___redArg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Structural_Positions_mapMwith___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__13___redArg___closed__5; -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__13___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___redArg___lam__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M_run___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -653,10 +638,12 @@ uint8_t l_Lean_Expr_isLambda(lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__29_spec__32___redArg(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__9; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__initFn___closed__21_00___x40_Lean_Meta_Tactic_FunInd_922168289____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__4___closed__3; LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__16_spec__20_spec__27(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__30; LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__4_spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M_ask(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M_tell___redArg___boxed(lean_object*, lean_object*, lean_object*); @@ -664,29 +651,27 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunI static lean_object* l_Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12___lam__2___closed__3; lean_object* l_Nat_reprFast(lean_object*); lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_withUserNamesImpl(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_Basic_0__Array_mapMUnsafe_map___at___00Lean_Tactic_FunInd_rwMatcher_spec__2___boxed(lean_object*, lean_object*, lean_object*); size_t lean_ptr_addr(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___00Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7_spec__14_spec__17___redArg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isTypeCorrect(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__0(lean_object*); +static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___closed__0; LEAN_EXPORT lean_object* l_Lean_Elab_FixedParamPerm_forallTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__20___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_initFn___lam__0_00___x40_Lean_Meta_Tactic_FunInd_4256658486____hygCtx___hyg_2_(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__1; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__15(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_FixedParamPerm_instantiateLambda(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__9___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_initFn___lam__0___closed__4_00___x40_Lean_Meta_Tactic_FunInd_4256658486____hygCtx___hyg_2_; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__2___closed__4; LEAN_EXPORT lean_object* l_Array_isEqvAux___at___00Lean_Elab_Structural_Positions_groupAndSort___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__6_spec__10___redArg___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___closed__11; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__5; -LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_rwMatcher___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__2___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingImp(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_Structural_eqnInfoExt; @@ -702,24 +687,21 @@ LEAN_EXPORT lean_object* l_Lean_throwUnknownIdentifierAt___at___00Lean_throwUnkn size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_panic___at___00Lean_Elab_Structural_Positions_mapMwith___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__13_spec__22(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Tactic_FunInd_isFunInductName___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__1___closed__5; LEAN_EXPORT lean_object* l_Array_contains___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__6___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__47___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit_visitPost___at___00__private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit___at___00Lean_Meta_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__1_spec__2_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___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__29___redArg___lam__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__11___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Tactic_FunInd_isFunInductName(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__1___closed__1; -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___closed__9; LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__24___redArg___lam__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_inProdLambdaLastArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwIfWith___closed__9; LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__19___redArg___lam__2(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__15___redArg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__1___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6___redArg___closed__0; LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_refinedArguments_go_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__29___redArg___lam__3___closed__3; @@ -730,8 +712,8 @@ LEAN_EXPORT lean_object* l_Lean_throwUnknownIdentifierAt___at___00Lean_throwUnkn LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_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*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__21; +static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__21___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_PersistentArray_forIn___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanupAfter_cleanupAfter_x3f_spec__1_spec__2(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00Lean_mkConstWithLevelParams___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__2_spec__3_spec__5_spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -748,13 +730,11 @@ static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_bu static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__6___closed__19; LEAN_EXPORT lean_object* l_Lean_withoutExporting___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__14(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__8___closed__2; -static lean_object* l_Lean_Tactic_FunInd_rwIfWith___closed__24; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withErasedFVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__18___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_getFunInductName(lean_object*, uint8_t); +static size_t l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7___redArg___closed__0; LEAN_EXPORT lean_object* l_Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases_spec__3___closed__1; -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___closed__4; lean_object* lean_st_ref_take(lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at___00Lean_Core_withIncRecDepth___at___00__private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0_spec__5_spec__7___redArg___closed__3; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -779,63 +759,53 @@ LEAN_EXPORT lean_object* l_panic___at___00__private_Lean_Meta_Tactic_FunInd_0__L static lean_object* l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13___redArg___closed__5; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_setNaryFunIndInfo___closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5___redArg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_inlineExpr(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13_spec__33(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__30___redArg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Elab_Structural_Positions_groupAndSort___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__6_spec__12(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__15_spec__27___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit_visitForall___at___00__private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit___at___00Lean_Meta_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__1_spec__2_spec__5___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_expr_eqv(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__11; lean_object* l_Lean_Meta_Match_Extension_getMatcherInfo_x3f(lean_object*, lean_object*); static lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__25_spec__27_spec__41___lam__0___closed__1; LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__19___boxed(lean_object**); static lean_object* l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_unpackMutualInduction_doRealize_spec__0___closed__0; -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__15; LEAN_EXPORT lean_object* l_Lean_getConstInfoRec___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__7; -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__13; LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__14___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00Lean_mkConstWithLevelParams___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__2_spec__3_spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00__private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit_visitForall___at___00__private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit___at___00Lean_Meta_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__1_spec__2_spec__5_spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__10___closed__2; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2_spec__9(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___redArg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwIfWith___closed__25; LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__initFn___closed__6_00___x40_Lean_Meta_Tactic_FunInd_922168289____hygCtx___hyg_2_; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deduplicateIHs___closed__0; static lean_object* l_Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12___closed__12; LEAN_EXPORT lean_object* l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M_localMapM___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___boxed(lean_object**); uint64_t lean_uint64_shift_right(uint64_t, uint64_t); static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg___closed__12; lean_object* l_Lean_Meta_Match_MatcherInfo_getMotivePos(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_PersistentArray_forInAux___at___00Lean_PersistentArray_forIn___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanupAfter_allHeqToEq_spec__0_spec__0_spec__3_spec__4(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___00__private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit_visitLet___at___00__private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit___at___00Lean_Meta_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__1_spec__2_spec__7_spec__9___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwIfWith___closed__5; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__3___closed__1; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__20___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M_exec___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__6___closed__0; LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__12(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__initFn_00___x40_Lean_Meta_Tactic_FunInd_922168289____hygCtx___hyg_2_(); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__21___boxed(lean_object**); lean_object* l_instMonadEST(lean_object*, lean_object*); lean_object* lean_nat_div(lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__3; lean_object* l_Lean_MVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_empty(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__4___closed__8; -LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6_spec__13(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__25_spec__27_spec__41___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Expr_isEq(lean_object*); +static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__21(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_unpackMutualInduction_doRealize_spec__0___lam__0___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__10___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -851,8 +821,9 @@ lean_object* l_Lean_Elab_Structural_Positions_numIndices(lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__15___redArg___lam__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_QSort_Basic_0__Array_qsort_sort___at___00Lean_Elab_Structural_Positions_groupAndSort___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__6_spec__11___redArg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___closed__13; +static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__25; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__initFn___closed__15_00___x40_Lean_Meta_Tactic_FunInd_922168289____hygCtx___hyg_2_; +static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___closed__2; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__4___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*); @@ -861,10 +832,10 @@ static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_bu LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M_ask___redArg___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_initFn___lam__0___closed__1_00___x40_Lean_Meta_Tactic_FunInd_4256658486____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l_Lean_withExporting___at___00Lean_withoutExporting___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__14_spec__39___redArg___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deduplicateIHs_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_kabstract(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__15___redArg___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7_spec__14(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withExporting___at___00Lean_withoutExporting___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__14_spec__39___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkHEqRefl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -875,17 +846,14 @@ static lean_object* l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___ LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__8___redArg(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__5___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__1(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__4_spec__4_spec__7___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_contains___at___00Lean_Meta_withErasedFVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__3_spec__3(lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwIfWith___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__11(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__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_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__10___redArg(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_getPrefix(lean_object*); @@ -895,7 +863,6 @@ lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withMVarContextImp(lean_ob lean_object* l_Lean_throwError___redArg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__16___lam__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__7(lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0___redArg(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__8___closed__0; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInduction___lam__2___closed__4; static lean_object* l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_unpackMutualInduction_doRealize_spec__0___closed__5; @@ -906,7 +873,6 @@ static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_bu LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__3___boxed(lean_object**); static lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__25_spec__27_spec__41___lam__0___closed__5; lean_object* l_Lean_Expr_projIdx_x21(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__8___redArg(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_projectMutualInduct___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__12___lam__0(uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withExporting___at___00Lean_withoutExporting___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__14_spec__39___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -915,7 +881,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__5___closed__0; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__4_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* l_Lean_MessageData_ofFormat(lean_object*); -static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__2; static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg___closed__20; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_refinedArguments_spec__1___redArg(uint8_t, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_PersistentArray_forIn___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanupAfter_cleanupAfter_x3f_spec__1_spec__2_spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -951,7 +916,6 @@ static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_de static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_withLetDecls___redArg___closed__1; static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__29___redArg___lam__3___closed__2; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_refinedArguments___closed__2; -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__22; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkAbsurd(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapFinIdxM_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__5___redArg___closed__4; @@ -962,9 +926,7 @@ static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_bu static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_initFn___lam__0___closed__3_00___x40_Lean_Meta_Tactic_FunInd_4256658486____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13_spec__36___redArg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg___closed__5; -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__6; double lean_float_of_nat(lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__10; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__8; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases_spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -976,6 +938,7 @@ lean_object* lean_st_ref_get(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13_spec__32(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13_spec__36___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__19; LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00Lean_mkConstWithLevelParams___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__2_spec__3_spec__5_spec__11_spec__14_spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -985,10 +948,8 @@ lean_object* lean_array_pop(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkAppM(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___closed__12; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Meta_withErasedFVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__3_spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__24___redArg___lam__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwIfWith___closed__26; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases_spec__3(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_getParamRevDeps(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -998,7 +959,7 @@ LEAN_EXPORT lean_object* l_panic___at___00__private_Lean_Meta_Tactic_FunInd_0__L static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__5___closed__1; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__initFn___closed__12_00___x40_Lean_Meta_Tactic_FunInd_922168289____hygCtx___hyg_2_; LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_isFunCasesName(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0___redArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__4(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___00Lean_PersistentArray_forIn___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanupAfter_cleanupAfter_x3f_spec__1_spec__1(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__6___closed__10; lean_object* lean_st_mk_ref(lean_object*); @@ -1006,7 +967,6 @@ static lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at__ static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M_run___redArg___closed__0; lean_object* l_Lean_MVarId_tryClearMany_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_zipWithMAux___at___00Lean_Elab_Structural_Positions_mapMwith___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__13_spec__23___redArg(lean_object*, lean_object*, 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_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__4(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__6___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInduction(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Option_get___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13_spec__37___boxed(lean_object*, lean_object*); @@ -1019,10 +979,8 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunI LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__16___lam__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__26___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwIfWith___closed__13; LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__18___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInduction___lam__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18_spec__21___redArg___boxed(lean_object*, lean_object*, 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_MessageData_hasTag(lean_object*, lean_object*); static lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13_spec__33___redArg___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__4_spec__6___redArg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -1062,11 +1020,9 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunI static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___redArg___lam__3___closed__3; LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__29___boxed(lean_object**); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Structural_Positions_mapMwith___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__13_spec__21(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapFinIdxM_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__5___redArg___closed__6; static lean_object* l_Lean_Elab_Structural_Positions_groupAndSort___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__6___redArg___closed__1; -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18_spec__21(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_revert(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_io_mono_nanos_now(); lean_object* l_Lean_Elab_FixedParamPerm_instantiateForall(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1075,23 +1031,23 @@ LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Uns LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__20(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Structural_Positions_groupAndSort___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__6___redArg___closed__0; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___closed__0; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__8(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_unpackMutualInduction_doRealize___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__14___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___closed__1; lean_object* l_Lean_Expr_constName_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00Lean_mkConstWithLevelParams___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__2_spec__3_spec__5_spec__11_spec__14___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__10___closed__5; extern lean_object* l_Lean_instInhabitedExpr; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars___closed__5; -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8___redArg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit___at___00Lean_Meta_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__1_spec__2_spec__8(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_rwIfWith___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_unpackMutualInduction(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_PersistentArray_forInAux___at___00Lean_PersistentArray_forIn___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanupAfter_allHeqToEq_spec__0_spec__0_spec__3(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13___redArg___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_withRewrittenMotive(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__4_spec__4_spec__7_spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___00Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6_spec__12_spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00__private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0_spec__6_spec__9___redArg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit_visitForall___at___00__private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit___at___00Lean_Meta_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__1_spec__2_spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_clearValue___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1127,22 +1083,20 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunI LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00Lean_mkConstWithLevelParams___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__2_spec__3_spec__5_spec__11_spec__14_spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__15___boxed(lean_object**); +static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__34; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInduction___lam__2___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__0(lean_object*); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13_spec__36___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__4_spec__4_spec__7_spec__18___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfoDefn___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__0___closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__3(lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__29___redArg___lam__3___closed__14; LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_letE___override(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); lean_object* l_Lean_Meta_mapErrorImp___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__1(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12___lam__2___closed__4; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__initFn___closed__5_00___x40_Lean_Meta_Tactic_FunInd_922168289____hygCtx___hyg_2_; lean_object* l_Lean_Meta_mkLetFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_refinedArguments_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_isMatcherApp___at___00Lean_Tactic_FunInd_rwMatcher_spec__1___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__21(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); @@ -1155,7 +1109,6 @@ lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); lean_object* l_Lean_FVarId_isLetVar___redArg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentArray_append___redArg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6_spec__13___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Core_instMonadCoreM___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__initFn___closed__17_00___x40_Lean_Meta_Tactic_FunInd_922168289____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Structural_Positions_groupAndSort___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__6_spec__9___redArg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); @@ -1169,16 +1122,13 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunI LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__2___redArg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_warningAsError; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_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_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_StateT_instMonad___redArg___lam__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLCtx___at___00Lean_Meta_withErasedFVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__3_spec__4___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_foldRevM_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__13___closed__1; static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__16___lam__2___closed__1; -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__4; -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__4___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__10___boxed(lean_object**); -lean_object* lean_get_congr_match_equations_for(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_AsyncConstantInfo_toConstantInfo(lean_object*); LEAN_EXPORT lean_object* l_Array_contains___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__14___boxed(lean_object*, lean_object*); @@ -1191,12 +1141,11 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_foldlM___at___00Lean_LocalContex static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__6___closed__5; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__16___closed__1; extern lean_object* l_Lean_trace_profiler_threshold; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20_spec__23___redArg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deduplicateIHs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_While_0__Lean_Loop_forIn_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_setNaryFunIndInfo_spec__2___redArg___closed__2; static lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__5___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__16(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__8(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__11___redArg(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__18; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1208,18 +1157,16 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00__private_Lean_Meta_ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_initFn___lam__0___closed__7_00___x40_Lean_Meta_Tactic_FunInd_4256658486____hygCtx___hyg_2_; -static size_t l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3___redArg___closed__0; uint8_t l_Lean_LocalDecl_isLet(lean_object*, uint8_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit_visitForall___at___00__private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit___at___00Lean_Meta_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__1_spec__2_spec__5___lam__0(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_contains___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deduplicateIHs_spec__0(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__7___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_extract___redArg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_PersistentArray_forIn___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanupAfter_cleanupAfter_x3f_spec__1_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_withErasedFVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__3___redArg___closed__0; -lean_object* l_Lean_MVarId_assumption(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__28___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Array_contains___at___00Lean_Meta_withErasedFVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__3_spec__3_spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -1232,21 +1179,21 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunI lean_object* l_StateT_bind(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg___closed__10; lean_object* l_Array_range(lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__12; static lean_object* l_Array_mapFinIdxM_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__5___redArg___closed__1; LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__19(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstVal___at___00Lean_mkConstWithLevelParams___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__2_spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Subarray_toArray___redArg(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations___lam__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__23; LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__14___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__2___closed__1; lean_object* l_Lean_LocalDecl_fvarId(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_Meta_withLCtx___at___00Lean_Meta_withErasedFVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__18_spec__33(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__9___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__initFn___closed__0_00___x40_Lean_Meta_Tactic_FunInd_922168289____hygCtx___hyg_2_; -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Tactic_FunInd_rwMatcher_spec__2(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__1___closed__0; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__8(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___closed__0; @@ -1254,9 +1201,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunI lean_object* l_Lean_Meta_Match_forallAltVarsTelescope___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__5___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_appFn_x21(lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7___redArg___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_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg___closed__13; -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__19; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deduplicateIHs_spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit___at___00Lean_Meta_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__1_spec__2___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1265,7 +1210,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunI LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__9___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_WellFounded_opaqueFix_u2083___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_refinedArguments_spec__2___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Tactic_FunInd_rwMatcher_spec__6(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit_visitLambda___at___00__private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit___at___00Lean_Meta_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__1_spec__2_spec__6(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___00Lean_PersistentArray_forIn___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanupAfter_allHeqToEq_spec__0_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1281,9 +1225,7 @@ LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Uns LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanupAfter_allHeqToEq_spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00__private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0_spec__6_spec__11___redArg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_assert(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18_spec__21___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_PersistentArray_forInAux___at___00Lean_PersistentArray_forIn___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanupAfter_allHeqToEq_spec__0_spec__0_spec__2(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__7(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_initFn___lam__0___closed__6_00___x40_Lean_Meta_Tactic_FunInd_4256658486____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_unpackMutualInduction_doRealize___lam__4(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_withLetDecls_go___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1292,6 +1234,7 @@ static lean_object* l_Lean_Meta_MatcherApp_transform___at___00__private_Lean_Met LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanupAfter_allHeqToEq___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_borrowed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkFalseElim(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Meta_withErasedFVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__3_spec__5(lean_object*, lean_object*, size_t, size_t, lean_object*); @@ -1300,7 +1243,6 @@ LEAN_EXPORT lean_object* l_panic___at___00Lean_Meta_MatcherApp_transform___at___ static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_setNaryFunIndInfo___closed__5; static lean_object* l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___closed__6; LEAN_EXPORT lean_object* l_Lean_getConstVal___at___00Lean_mkConstWithLevelParams___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases_spec__2_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___closed__1; LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__29___redArg___lam__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1310,8 +1252,8 @@ lean_object* lean_usize_to_nat(size_t); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_refinedArguments_go_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase___lam__2___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_isFunCasesName___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7_spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static double l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13___redArg___closed__3; -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at___00Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3_spec__12___redArg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__19___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Array_contains___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deduplicateIHs_spec__0_spec__0(lean_object*, lean_object*, size_t, size_t); @@ -1329,12 +1271,12 @@ uint8_t l_Lean_Environment_contains(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__8(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__4_spec__4_spec__7_spec__18_spec__28___redArg___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_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__29___redArg___lam__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20___redArg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13___redArg___closed__1; lean_object* l_Lean_Meta_ArgsPacker_numFuncs(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInduction___closed__0; -LEAN_EXPORT lean_object* l___private_Init_While_0__Lean_Loop_forIn_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofExpr(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__6(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1356,20 +1298,17 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunI static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInduction___lam__2___closed__2; LEAN_EXPORT uint8_t l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_unpackMutualInduction_doRealize_spec__0___lam__0(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLCtx___at___00Lean_Meta_withErasedFVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__3_spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___closed__3; LEAN_EXPORT lean_object* l___private_Init_Data_Array_QSort_Basic_0__Array_qsort_sort___at___00Lean_Elab_Structural_Positions_groupAndSort___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__6_spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_unpackMutualInduction_doRealize_spec__0___closed__3; -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18_spec__21___redArg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_insert___at___00__private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0_spec__6(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__4___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__3; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__8___closed__7; +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__4___redArg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__17___closed__0; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__1___closed__3; lean_object* l_Lean_Meta_getFunCasesName(lean_object*, uint8_t); lean_object* l_Lean_Expr_eta(lean_object*); LEAN_EXPORT uint8_t l_Array_contains___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__6(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00Lean_Tactic_FunInd_rwMatcher_spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00Lean_mkConstWithLevelParams___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__2_spec__3_spec__5_spec__11_spec__14_spec__16___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__4_spec__10(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_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_refinedArguments_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1377,10 +1316,9 @@ LEAN_EXPORT lean_object* l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_Fun LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_insert___at___00__private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0_spec__6___redArg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_finIdxOf_x3f___at___00Array_idxOf_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__11_spec__18(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__0___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_rwIfWith___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwIfWith___closed__19; lean_object* l_Lean_Expr_getAppNumArgs(lean_object*); lean_object* l_Lean_LocalDecl_userName(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__14(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__6___closed__9; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__6___closed__15; lean_object* l_Lean_Meta_Context_config(lean_object*); @@ -1399,10 +1337,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunI static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__20___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Structural_Positions_mapMwith___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__13___redArg___closed__2; -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg___closed__0; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3___redArg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__30___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__19___redArg___lam__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__28___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1418,12 +1353,12 @@ LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___00Lean_throwUnknownIdentifie static lean_object* l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__4; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__21___closed__0; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_PersistentArray_forInAux___at___00Lean_PersistentArray_forIn___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanupAfter_cleanupAfter_x3f_spec__1_spec__1_spec__3_spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__0; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__4___closed__5; lean_object* l_Lean_Meta_MatcherApp_inferMatchType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_zipMasked___redArg(lean_object*, lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__29___redArg___lam__3___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations___lam__0___closed__0; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1432,10 +1367,12 @@ LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00__private_Lean_Meta_Ta LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__5___redArg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deduplicateIHs___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanupAfter_allHeqToEq_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___closed__3; LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__4_spec__4_spec__7_spec__18_spec__28_spec__31_spec__33___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__18___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_heqToEq(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit_visitForall___at___00__private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit___at___00Lean_Meta_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__1_spec__2_spec__5(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7_spec__15___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00__private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit_visitForall___at___00__private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit___at___00Lean_Meta_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__1_spec__2_spec__5_spec__6___redArg(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_Transform_0__Lean_Meta_transformWithCache_visit___at___00Lean_Meta_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__1_spec__2___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars___lam__0___closed__0; @@ -1451,7 +1388,6 @@ static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_de static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanupAfter_cleanupAfter_x3f___lam__0___closed__0; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00__private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit_visitForall___at___00__private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit___at___00Lean_Meta_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__1_spec__2_spec__5_spec__6___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_FixedParamPerm_buildArgs___redArg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_constLevels_x21(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__15___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_instBEqFVarId_beq(lean_object*, lean_object*); @@ -1461,17 +1397,12 @@ lean_object* l_Lean_Expr_bindingDomain_x21(lean_object*); lean_object* l_Lean_Expr_letValue_x21(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__28___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Subarray_empty(lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__14; -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__13(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_withLetDecls___redArg___closed__0; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__19(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__27___closed__4; -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__24___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwIfWith___closed__23; -static lean_object* l_Lean_Tactic_FunInd_rwIfWith___closed__0; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_unpackMutualInduction_doRealize___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__19___redArg___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__9(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1481,7 +1412,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00__private_Lea LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__4_spec__8___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_Structural_instInhabitedEqnInfo_default; LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at___00Lean_getConstVal___at___00Lean_mkConstWithLevelParams___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__2_spec__3_spec__5___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___closed__0; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__4___closed__0; static lean_object* l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__2; static lean_object* l_Lean_Elab_Structural_Positions_mapMwith___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__13___redArg___closed__3; @@ -1493,10 +1424,10 @@ static lean_object* l_Lean_throwMaxRecDepthAt___at___00Lean_Core_withIncRecDepth LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__7(lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_Meta_instMonadMetaM___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_withLetDecls_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6_spec__13___redArg(size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__9___redArg___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_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars___boxed__const__1; @@ -1506,7 +1437,6 @@ LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at___00__private_Lean_M LEAN_EXPORT lean_object* l_panic___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_instMonadMetaM___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInduction___lam__2___closed__0; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at___00Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3_spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__15___redArg___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addMessageContextFull___at___00Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__2_spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1515,6 +1445,7 @@ static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_de uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Array_contains___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__6_spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkApp3(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__32; LEAN_EXPORT lean_object* l___private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__4_spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withExporting___at___00Lean_withoutExporting___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__14_spec__39___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1538,9 +1469,8 @@ static lean_object* l_Lean_throwMaxRecDepthAt___at___00Lean_Core_withIncRecDepth LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__5(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isProj(lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwIfWith___closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__17(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_assertIHs___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__1___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1556,34 +1486,32 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunI LEAN_EXPORT lean_object* l_Lean_getConstInfoInduct___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__12; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit___at___00Lean_Meta_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__1_spec__2_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*); -static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_withLetDecls_go_spec__0___redArg(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__29___redArg___closed__0; LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__19___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__19___redArg(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Meta_isMatcherAppCore(lean_object*, lean_object*); static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg___closed__14; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_unpackMutualInduction_doRealize(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__4___closed__7; LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__15(lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00__private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0_spec__6_spec__10(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__2___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15___redArg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withExporting___at___00Lean_withoutExporting___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__14_spec__39___redArg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwIfWith___closed__11; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__4_spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__29(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentArray_toArray___redArg(lean_object*); lean_object* l_Lean_Meta_mkEqRefl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__1___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_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0___redArg(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_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__8___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__1(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_indentExpr(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_lambdaBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__9_spec__25___redArg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__4___closed__4; LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__19_spec__24(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___closed__2; uint8_t l_Lean_Expr_isHEq(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit___at___00Lean_Meta_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__1_spec__2_spec__3(lean_object*, lean_object*, uint8_t, uint8_t, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_While_0__Lean_Loop_forIn_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_mkLambdaFVarsMasked_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1597,12 +1525,10 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunI static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__16___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__30___redArg(lean_object*, lean_object*); lean_object* l_Lean_EnvironmentHeader_moduleNames(lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7___redArg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__9___redArg(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__4___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_Meta_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__1___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5___lam__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwIfWith___closed__15; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___boxed(lean_object**); static lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Structural_Positions_groupAndSort___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__6_spec__9___redArg___closed__0; lean_object* l_Lean_Meta_unfoldDefinition_x3f(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1615,7 +1541,6 @@ LEAN_EXPORT lean_object* l_Array_idxOfAux___at___00Array_finIdxOf_x3f___at___00A LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_refinedArguments_spec__1___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__4_spec__4_spec__7_spec__18_spec__28_spec__32___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Environment_isSafeDefinition(lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwIfWith___closed__3; static lean_object* l_Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12___closed__2; static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg___closed__15; uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); @@ -1633,12 +1558,10 @@ lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__8___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_SavedState_restore___redArg(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__35; lean_object* l_Lean_MapDeclarationExtension_find_x3f___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_setNaryFunIndInfo___closed__9; -uint8_t l_Lean_Meta_Simp_isEqnThmHypothesis(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_rwIfWith(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__20___closed__1; -static lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__1___closed__0; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_setNaryFunIndInfo___closed__6; lean_object* l_Lean_Meta_instInhabitedMetaM___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_initFn___lam__0___closed__8_00___x40_Lean_Meta_Tactic_FunInd_4256658486____hygCtx___hyg_2_; @@ -1649,7 +1572,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunI lean_object* l_Lean_MVarId_clear___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_MatcherApp_toExpr(lean_object*); -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3(lean_object*, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withErasedFVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_type(lean_object*); static lean_object* l_Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12___closed__7; @@ -1657,18 +1579,15 @@ static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknow LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__4_spec__4_spec__7___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t lean_uint64_xor(uint64_t, uint64_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M_run(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2(lean_object*, size_t, size_t, lean_object*); lean_object* l_Array_toSubarray___redArg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__19___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Std_Data_DHashMap_Internal_Defs_0__Std_DHashMap_Internal_Raw_u2080_expand_go___at___00Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00__private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0_spec__6_spec__10_spec__11(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_initFn___lam__0___closed__2_00___x40_Lean_Meta_Tactic_FunInd_4256658486____hygCtx___hyg_2_; lean_object* lean_panic_fn(lean_object*, lean_object*); lean_object* l_Lean_Meta_saveState___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_setNaryFunIndInfo_spec__2___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit_visitLet___at___00__private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit___at___00Lean_Meta_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__1_spec__2_spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__11(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__3___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanupAfter_go___boxed(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_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__16___lam__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Transform_0__Lean_Core_transform_visit_visitPost___at___00__private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1695,20 +1614,17 @@ lean_object* lean_nat_sub(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__8___closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Structural_Positions_mapMwith___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__13___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_setNaryFunIndInfo___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___boxed(lean_object**); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase___lam__1___closed__1; -static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5___closed__2; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13_spec__36(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__15___redArg___lam__5(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0___closed__0; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__22___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1___closed__0; lean_object* l_Lean_Expr_getAppFn(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9___redArg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_append___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__25_spec__27_spec__41(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18_spec__21___redArg___lam__0___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_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__4_spec__4_spec__7_spec__18_spec__28_spec__32___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__3; lean_object* l_Lean_Meta_Simp_mkCongr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1717,7 +1633,6 @@ lean_object* l_ReaderT_instApplicativeOfMonad___redArg___lam__4(lean_object*, le lean_object* l_Lean_Meta_PProdN_mk(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__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___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwMData(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg___closed__18; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__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*); @@ -1725,52 +1640,52 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunI LEAN_EXPORT lean_object* l_Array_zipWithMAux___at___00Lean_Elab_Structural_Positions_mapMwith___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__13_spec__23___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t lean_uint64_shift_left(uint64_t, uint64_t); +static lean_object* l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7___redArg___closed__2; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__0(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInduction___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__30(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_isPrivateName(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__6___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__18; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_projectMutualInduct_spec__1___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkArrow(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__21___closed__1; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__1; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Structural_Positions_mapMwith___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__13_spec__21___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___00Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7_spec__14_spec__17(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__22___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_findIdx_x3f_loop___redArg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addMessageContextFull___at___00Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__2_spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20_spec__23___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7_spec__15___redArg(size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_withLetDecls_go___redArg___lam__0___boxed(lean_object*, lean_object*, 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_Std_DHashMap_Internal_AssocList_contains___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00__private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0_spec__6_spec__9___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__9___redArg___lam__1(uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Option_get___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13_spec__37(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14___redArg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___closed__3; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__10; static lean_object* l_Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12___closed__9; LEAN_EXPORT lean_object* l_panic___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_withLetDecls_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__19___redArg(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_string_memcmp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1___closed__1; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__20___closed__11; lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); lean_object* l_Lean_Meta_check___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6_spec__12___redArg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_letToHave(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__1(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__1; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__initFn___closed__18_00___x40_Lean_Meta_Tactic_FunInd_922168289____hygCtx___hyg_2_; lean_object* l_Lean_Environment_getModuleIdxFor_x3f(lean_object*, lean_object*); static lean_object* l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__13; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__11(uint8_t, 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_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__1___closed__4; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__13(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6_spec__12(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Subarray_size___redArg(lean_object*); LEAN_EXPORT lean_object* l_Lean_observing_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanupAfter_cleanupAfter_x3f_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00Lean_Tactic_FunInd_rwMatcher_spec__3___redArg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__4_spec__7(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__8___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__8(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1793,8 +1708,6 @@ lean_object* l_Lean_Meta_getUnfoldEqnFor_x3f(lean_object*, uint8_t, lean_object* LEAN_EXPORT lean_object* l_Lean_Elab_Structural_Positions_groupAndSort___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__6___redArg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_PersistentArray_forInAux___at___00Lean_PersistentArray_forIn___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanupAfter_cleanupAfter_x3f_spec__1_spec__1_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_refinedArguments_go_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__1_spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwIfWith___closed__16; LEAN_EXPORT lean_object* l_Lean_Elab_FixedParamPerm_forallTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__20___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___redArg___lam__3___closed__2; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__4___closed__9; @@ -1807,7 +1720,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunI LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__5___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInduction___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__10___closed__3; -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__16; static lean_object* l_Lean_Elab_Structural_Positions_mapMwith___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__13___redArg___closed__4; LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__15___redArg___lam__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__20___boxed(lean_object**); @@ -1830,7 +1742,6 @@ LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta LEAN_EXPORT lean_object* l_Lean_Elab_FixedParamPerm_forallTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__20___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___redArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__6; -LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__6___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_withIncRecDepth___at___00__private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit___at___00Lean_Meta_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__1_spec__2_spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations___lam__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -1838,7 +1749,6 @@ size_t lean_usize_add(size_t, size_t); LEAN_EXPORT uint8_t l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__19(lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_assertIHs_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__22(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___closed__10; lean_object* l_mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Structural_Positions_groupAndSort___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__6___redArg___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1854,8 +1764,6 @@ LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tac LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_FVarId_getDecl___redArg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__24___redArg(lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Meta_mkEqOfHEq(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___closed__15; lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAux(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withIncRecDepth___at___00__private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit___at___00Lean_Meta_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__1_spec__2_spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__3; @@ -1867,7 +1775,6 @@ LEAN_EXPORT lean_object* l_Lean_Core_transform___at___00__private_Lean_Meta_Tact LEAN_EXPORT uint8_t l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Array_contains___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__14_spec__25(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__18(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); -static lean_object* l_Lean_Tactic_FunInd_rwIfWith___closed__14; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__4___redArg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanupAfter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_fvar___override(lean_object*); @@ -1876,14 +1783,12 @@ size_t lean_array_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__20___boxed(lean_object**); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_refinedArguments_go_spec__1___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__20___closed__6; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M_eval___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M_localMapM_spec__0(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_trace_profiler; -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18_spec__21___redArg___lam__0(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__16_spec__20___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__initFn___closed__2_00___x40_Lean_Meta_Tactic_FunInd_922168289____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l_panic___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__26(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1893,10 +1798,7 @@ LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_m LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases_spec__1___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__11___closed__0; static lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_setNaryFunIndInfo_spec__2___redArg___closed__0; -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__7; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M_exec___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static size_t l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3___redArg___closed__1; -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__4_spec__10___redArg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkFVar(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Array_contains___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deduplicateIHs_spec__0_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -1908,7 +1810,6 @@ static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_fo LEAN_EXPORT lean_object* l_panic___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*); lean_object* l_Lean_FVarId_getType___redArg(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___00Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6_spec__12_spec__15___redArg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_refinedArguments___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__initFn___closed__20_00___x40_Lean_Meta_Tactic_FunInd_922168289____hygCtx___hyg_2_; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__initFn___closed__7_00___x40_Lean_Meta_Tactic_FunInd_922168289____hygCtx___hyg_2_; @@ -1926,7 +1827,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunI LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__4_spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__4_spec__4_spec__7_spec__18_spec__28_spec__31_spec__33___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_bindingBody_x21(lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwIfWith___closed__8; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_getLocalInstances___redArg(lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__29___redArg___lam__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*); @@ -1951,7 +1851,6 @@ lean_object* l_Lean_Meta_Match_MatcherInfo_getNumDiscrEqs(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___closed__0; lean_object* l_instDecidableEqNat___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__7; -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_inProdLambdaLastArg___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__16___lam__0___boxed(lean_object*, lean_object*); @@ -1963,31 +1862,25 @@ LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_ArgsPacker_unpack(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___redArg___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_Tactic_FunInd_rwMatcher___lam__1___closed__17; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00__private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0_spec__6_spec__11(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__12___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural___boxed(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_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5(lean_object*, uint8_t, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__29___redArg___lam__3___closed__12; LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__1(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__15___redArg___lam__4(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__4_spec__4_spec__7_spec__18_spec__28_spec__31_spec__33(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at___00Lean_Core_withIncRecDepth___at___00__private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0_spec__5_spec__7___redArg___closed__0; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Structural_Positions_groupAndSort___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__6_spec__9(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12___closed__8; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M_localM(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__10___redArg___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__6___boxed(lean_object**); -static lean_object* l_Lean_Tactic_FunInd_rwIfWith___closed__20; +static lean_object* l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__18; lean_object* l_Lean_Expr_headBeta(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanupAfter_cleanupAfter_x3f___lam__0___boxed(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_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M_exec(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwIfWith___closed__4; static lean_object* l_Lean_getConstInfoRec___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__3___closed__0; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__47___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_WF_instInhabitedEqnInfo_default; @@ -1998,6 +1891,7 @@ lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M_localMapM_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_withRewrittenMotive___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__10___closed__4; +static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__1___closed__2; static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg___closed__19; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__13___closed__2; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__initFn___closed__4_00___x40_Lean_Meta_Tactic_FunInd_922168289____hygCtx___hyg_2_; @@ -2016,7 +1910,6 @@ lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__4_spec__4_spec__7_spec__18_spec__28___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__8___redArg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___00__private_Std_Data_DHashMap_Internal_Defs_0__Std_DHashMap_Internal_Raw_u2080_expand_go___at___00Std_DHashMap_Internal_Raw_u2080_expand___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00__private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0_spec__6_spec__10_spec__11_spec__12___redArg(lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__19(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2033,11 +1926,9 @@ LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_M static lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__29___redArg___lam__3___closed__6; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__2___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__20(uint8_t, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_lambdaBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__9_spec__15___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_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg___closed__3; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_setNaryFunIndInfo___closed__0; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfoInduct___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__2___closed__1; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__11___closed__3; @@ -2057,7 +1948,6 @@ LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___00Lean_Core_withIncRec LEAN_EXPORT lean_object* l_Lean_observing_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanupAfter_cleanupAfter_x3f_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_exceptEmoji___redArg(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Meta_withErasedFVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__3_spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__2; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__11___closed__2; static lean_object* l_Lean_getConstInfoRec___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__3___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwMData___redArg___boxed(lean_object*, lean_object*); @@ -2082,7 +1972,6 @@ LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___00Lean_Core_withIncRec static lean_object* l_Array_mapFinIdxM_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__5___redArg___closed__2; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_PersistentArray_forInAux___at___00Lean_PersistentArray_forIn___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanupAfter_allHeqToEq_spec__0_spec__0_spec__3_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 uint8_t l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Array_contains___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__6_spec__9(lean_object*, lean_object*, size_t, size_t); -LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6_spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_mdataExpr_x21(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_mkLambdaFVarsMasked(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Environment_setExporting(lean_object*, uint8_t); @@ -2092,7 +1981,6 @@ static lean_object* l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_Fun LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M2_run___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_PersistentArray_forInAux___at___00Lean_PersistentArray_forIn___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanupAfter_cleanupAfter_x3f_spec__1_spec__1_spec__3_spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isFVar(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__20___closed__0; extern lean_object* l_Lean_Elab_Structural_instInhabitedRecArgInfo_default; uint8_t l_Lean_Expr_isForall(lean_object*); @@ -2112,13 +2000,14 @@ lean_object* l_Lean_InductiveVal_numCtors(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___00Lean_Meta_withIncRecDepth___at___00__private_Lean_Meta_Transform_0__Lean_Meta_transformWithCache_visit___at___00Lean_Meta_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__1_spec__2_spec__9_spec__12___redArg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLCtx___at___00Lean_Meta_withErasedFVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__3_spec__4___redArg(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_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__9___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__0; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars___lam__0___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__24___redArg___lam__5(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_consumeMData(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg___closed__6; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_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_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__20___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__5___closed__0; static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__6___closed__13; @@ -2126,13 +2015,11 @@ LEAN_EXPORT lean_object* l___private_Init_While_0__Lean_Loop_forIn_loop___at___0 LEAN_EXPORT lean_object* l_Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__4_spec__4_spec__7_spec__18(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__27___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_withLetDecls(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_zip___redArg(lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__0___closed__1; uint8_t l_Lean_isCasesOnRecursor(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanupAfter_cleanupAfter_x3f(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2141,11 +2028,10 @@ LEAN_EXPORT lean_object* l_Lean_Meta_lambdaBoundedTelescope___at___00__private_L lean_object* lean_expr_instantiate1(lean_object*, lean_object*); static lean_object* l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_unpackMutualInduction_doRealize_spec__0___closed__4; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__7___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwIfWith___closed__6; lean_object* l_List_get_x21Internal___redArg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13_spec__35_spec__39___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__24___redArg___lam__4(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__27; LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__21___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__25_spec__27_spec__41___lam__0___closed__4; LEAN_EXPORT lean_object* l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2159,8 +2045,6 @@ LEAN_EXPORT lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknow LEAN_EXPORT lean_object* l_Lean_withoutExporting___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__14___redArg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkLevelParam(lean_object*); lean_object* l_Lean_isInductiveCore_x3f(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__7___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_unpackMutualInduction_doRealize___lam__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__11___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__29___redArg___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2171,23 +2055,19 @@ static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_de LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__23(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__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_forallBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__8___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_MVarId_hrefl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__11(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_FixedParamPerm_forallTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__20(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___closed__1; LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__24___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__10___closed__1; static lean_object* l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__14; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addDecl(lean_object*, uint8_t, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00Lean_Tactic_FunInd_rwMatcher_spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageLog_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__23___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_StateT_instMonad___redArg___lam__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwIfWith___closed__10; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20_spec__23___redArg___lam__0(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__17___closed__3; size_t lean_usize_land(size_t, size_t); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__29___redArg___lam__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*); @@ -2196,22 +2076,23 @@ LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at___00Lean_getConstInfo_ lean_object* l_StateT_pure(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_setFunIndInfo(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__28(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_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_elimOptParam(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__28___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_withIncRecDepth___at___00__private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0_spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfoRec___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__3___closed__2; LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_projectMutualInduct_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__0___closed__3; LEAN_EXPORT lean_object* l_Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__25(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__15___boxed(lean_object**); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__4_spec__8___redArg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__4___redArg___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__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_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0___redArg___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__31; LEAN_EXPORT lean_object* l_Lean_Meta_withLCtx___at___00Lean_Meta_withErasedFVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__18_spec__33___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__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_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__24___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); double lean_float_sub(double, double); LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__9___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -44219,9727 +44100,2352 @@ x_8 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_inProdLambdaLast return x_8; } } -static lean_object* _init_l_Lean_Tactic_FunInd_rwIfWith___closed__0() { +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__0___redArg(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("cond", 4, 4); -return x_1; +lean_object* x_4; lean_object* x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; +x_4 = lean_ctor_get(x_2, 2); +x_5 = lean_ctor_get(x_2, 13); +x_6 = l_Lean_checkTraceOption(x_5, x_4, x_1); +x_7 = lean_box(x_6); +x_8 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_8, 0, x_7); +return x_8; } } -static lean_object* _init_l_Lean_Tactic_FunInd_rwIfWith___closed__1() { +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Tactic_FunInd_rwIfWith___closed__0; -x_2 = l_Lean_Name_mkStr1(x_1); -return x_2; +lean_object* x_7; +x_7 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__0___redArg(x_1, x_4); +return x_7; } } -static lean_object* _init_l_Lean_Tactic_FunInd_rwIfWith___closed__2() { +LEAN_EXPORT lean_object* l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("dite", 4, 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwIfWith___closed__3() { -_start: +lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_8 = lean_ctor_get(x_5, 5); +x_9 = l_Lean_addMessageContextFull___at___00Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__2_spec__7(x_2, x_3, x_4, x_5, x_6); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Tactic_FunInd_rwIfWith___closed__2; -x_2 = l_Lean_Name_mkStr1(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwIfWith___closed__4() { -_start: +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_st_ref_take(x_6); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("ite", 3, 3); -return x_1; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwIfWith___closed__5() { -_start: +lean_object* x_14; uint8_t x_15; +x_14 = lean_ctor_get(x_12, 4); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Tactic_FunInd_rwIfWith___closed__4; -x_2 = l_Lean_Name_mkStr1(x_1); -return x_2; -} +lean_object* x_16; double x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_16 = lean_ctor_get(x_14, 0); +x_17 = l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13___redArg___closed__0; +x_18 = 0; +x_19 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__21___closed__3; +x_20 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_20, 0, x_1); +lean_ctor_set(x_20, 1, x_19); +lean_ctor_set_float(x_20, sizeof(void*)*2, x_17); +lean_ctor_set_float(x_20, sizeof(void*)*2 + 8, x_17); +lean_ctor_set_uint8(x_20, sizeof(void*)*2 + 16, x_18); +x_21 = l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__2___closed__0; +x_22 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_11); +lean_ctor_set(x_22, 2, x_21); +lean_inc(x_8); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_8); +lean_ctor_set(x_23, 1, x_22); +x_24 = l_Lean_PersistentArray_push___redArg(x_16, x_23); +lean_ctor_set(x_14, 0, x_24); +x_25 = lean_st_ref_set(x_6, x_12); +x_26 = lean_box(0); +lean_ctor_set(x_9, 0, x_26); +return x_9; } -static lean_object* _init_l_Lean_Tactic_FunInd_rwIfWith___closed__6() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("if_neg", 6, 6); -return x_1; +uint64_t x_27; lean_object* x_28; double x_29; uint8_t x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_27 = lean_ctor_get_uint64(x_14, sizeof(void*)*1); +x_28 = lean_ctor_get(x_14, 0); +lean_inc(x_28); +lean_dec(x_14); +x_29 = l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13___redArg___closed__0; +x_30 = 0; +x_31 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__21___closed__3; +x_32 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_32, 0, x_1); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set_float(x_32, sizeof(void*)*2, x_29); +lean_ctor_set_float(x_32, sizeof(void*)*2 + 8, x_29); +lean_ctor_set_uint8(x_32, sizeof(void*)*2 + 16, x_30); +x_33 = l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__2___closed__0; +x_34 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_11); +lean_ctor_set(x_34, 2, x_33); +lean_inc(x_8); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_8); +lean_ctor_set(x_35, 1, x_34); +x_36 = l_Lean_PersistentArray_push___redArg(x_28, x_35); +x_37 = lean_alloc_ctor(0, 1, 8); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set_uint64(x_37, sizeof(void*)*1, x_27); +lean_ctor_set(x_12, 4, x_37); +x_38 = lean_st_ref_set(x_6, x_12); +x_39 = lean_box(0); +lean_ctor_set(x_9, 0, x_39); +return x_9; } } -static lean_object* _init_l_Lean_Tactic_FunInd_rwIfWith___closed__7() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Tactic_FunInd_rwIfWith___closed__6; -x_2 = l_Lean_Name_mkStr1(x_1); -return x_2; +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint64_t x_49; lean_object* x_50; lean_object* x_51; double x_52; uint8_t x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_40 = lean_ctor_get(x_12, 4); +x_41 = lean_ctor_get(x_12, 0); +x_42 = lean_ctor_get(x_12, 1); +x_43 = lean_ctor_get(x_12, 2); +x_44 = lean_ctor_get(x_12, 3); +x_45 = lean_ctor_get(x_12, 5); +x_46 = lean_ctor_get(x_12, 6); +x_47 = lean_ctor_get(x_12, 7); +x_48 = lean_ctor_get(x_12, 8); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_40); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_12); +x_49 = lean_ctor_get_uint64(x_40, sizeof(void*)*1); +x_50 = lean_ctor_get(x_40, 0); +lean_inc_ref(x_50); +if (lean_is_exclusive(x_40)) { + lean_ctor_release(x_40, 0); + x_51 = x_40; +} else { + lean_dec_ref(x_40); + x_51 = lean_box(0); } +x_52 = l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13___redArg___closed__0; +x_53 = 0; +x_54 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__21___closed__3; +x_55 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_55, 0, x_1); +lean_ctor_set(x_55, 1, x_54); +lean_ctor_set_float(x_55, sizeof(void*)*2, x_52); +lean_ctor_set_float(x_55, sizeof(void*)*2 + 8, x_52); +lean_ctor_set_uint8(x_55, sizeof(void*)*2 + 16, x_53); +x_56 = l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__2___closed__0; +x_57 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_11); +lean_ctor_set(x_57, 2, x_56); +lean_inc(x_8); +x_58 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_58, 0, x_8); +lean_ctor_set(x_58, 1, x_57); +x_59 = l_Lean_PersistentArray_push___redArg(x_50, x_58); +if (lean_is_scalar(x_51)) { + x_60 = lean_alloc_ctor(0, 1, 8); +} else { + x_60 = x_51; } -static lean_object* _init_l_Lean_Tactic_FunInd_rwIfWith___closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(6u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set_uint64(x_60, sizeof(void*)*1, x_49); +x_61 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_61, 0, x_41); +lean_ctor_set(x_61, 1, x_42); +lean_ctor_set(x_61, 2, x_43); +lean_ctor_set(x_61, 3, x_44); +lean_ctor_set(x_61, 4, x_60); +lean_ctor_set(x_61, 5, x_45); +lean_ctor_set(x_61, 6, x_46); +lean_ctor_set(x_61, 7, x_47); +lean_ctor_set(x_61, 8, x_48); +x_62 = lean_st_ref_set(x_6, x_61); +x_63 = lean_box(0); +lean_ctor_set(x_9, 0, x_63); +return x_9; } } -static lean_object* _init_l_Lean_Tactic_FunInd_rwIfWith___closed__9() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("if_pos", 6, 6); -return x_1; -} +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; uint64_t x_76; lean_object* x_77; lean_object* x_78; double x_79; uint8_t x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_64 = lean_ctor_get(x_9, 0); +lean_inc(x_64); +lean_dec(x_9); +x_65 = lean_st_ref_take(x_6); +x_66 = lean_ctor_get(x_65, 4); +lean_inc_ref(x_66); +x_67 = lean_ctor_get(x_65, 0); +lean_inc_ref(x_67); +x_68 = lean_ctor_get(x_65, 1); +lean_inc(x_68); +x_69 = lean_ctor_get(x_65, 2); +lean_inc_ref(x_69); +x_70 = lean_ctor_get(x_65, 3); +lean_inc_ref(x_70); +x_71 = lean_ctor_get(x_65, 5); +lean_inc_ref(x_71); +x_72 = lean_ctor_get(x_65, 6); +lean_inc_ref(x_72); +x_73 = lean_ctor_get(x_65, 7); +lean_inc_ref(x_73); +x_74 = lean_ctor_get(x_65, 8); +lean_inc_ref(x_74); +if (lean_is_exclusive(x_65)) { + lean_ctor_release(x_65, 0); + lean_ctor_release(x_65, 1); + lean_ctor_release(x_65, 2); + lean_ctor_release(x_65, 3); + lean_ctor_release(x_65, 4); + lean_ctor_release(x_65, 5); + lean_ctor_release(x_65, 6); + lean_ctor_release(x_65, 7); + lean_ctor_release(x_65, 8); + x_75 = x_65; +} else { + lean_dec_ref(x_65); + x_75 = lean_box(0); } -static lean_object* _init_l_Lean_Tactic_FunInd_rwIfWith___closed__10() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Tactic_FunInd_rwIfWith___closed__9; -x_2 = l_Lean_Name_mkStr1(x_1); -return x_2; +x_76 = lean_ctor_get_uint64(x_66, sizeof(void*)*1); +x_77 = lean_ctor_get(x_66, 0); +lean_inc_ref(x_77); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + x_78 = x_66; +} else { + lean_dec_ref(x_66); + x_78 = lean_box(0); } +x_79 = l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13___redArg___closed__0; +x_80 = 0; +x_81 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__21___closed__3; +x_82 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_82, 0, x_1); +lean_ctor_set(x_82, 1, x_81); +lean_ctor_set_float(x_82, sizeof(void*)*2, x_79); +lean_ctor_set_float(x_82, sizeof(void*)*2 + 8, x_79); +lean_ctor_set_uint8(x_82, sizeof(void*)*2 + 16, x_80); +x_83 = l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__2___closed__0; +x_84 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_84, 0, x_82); +lean_ctor_set(x_84, 1, x_64); +lean_ctor_set(x_84, 2, x_83); +lean_inc(x_8); +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_8); +lean_ctor_set(x_85, 1, x_84); +x_86 = l_Lean_PersistentArray_push___redArg(x_77, x_85); +if (lean_is_scalar(x_78)) { + x_87 = lean_alloc_ctor(0, 1, 8); +} else { + x_87 = x_78; } -static lean_object* _init_l_Lean_Tactic_FunInd_rwIfWith___closed__11() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("dif_neg", 7, 7); -return x_1; +lean_ctor_set(x_87, 0, x_86); +lean_ctor_set_uint64(x_87, sizeof(void*)*1, x_76); +if (lean_is_scalar(x_75)) { + x_88 = lean_alloc_ctor(0, 9, 0); +} else { + x_88 = x_75; } +lean_ctor_set(x_88, 0, x_67); +lean_ctor_set(x_88, 1, x_68); +lean_ctor_set(x_88, 2, x_69); +lean_ctor_set(x_88, 3, x_70); +lean_ctor_set(x_88, 4, x_87); +lean_ctor_set(x_88, 5, x_71); +lean_ctor_set(x_88, 6, x_72); +lean_ctor_set(x_88, 7, x_73); +lean_ctor_set(x_88, 8, x_74); +x_89 = lean_st_ref_set(x_6, x_88); +x_90 = lean_box(0); +x_91 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_91, 0, x_90); +return x_91; } -static lean_object* _init_l_Lean_Tactic_FunInd_rwIfWith___closed__12() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Tactic_FunInd_rwIfWith___closed__11; -x_2 = l_Lean_Name_mkStr1(x_1); -return x_2; } } -static lean_object* _init_l_Lean_Tactic_FunInd_rwIfWith___closed__13() { +static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___closed__0() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("dif_pos", 7, 7); +x_1 = lean_mk_string_unchecked("rwLetWith failed:", 17, 17); return x_1; } } -static lean_object* _init_l_Lean_Tactic_FunInd_rwIfWith___closed__14() { +static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___closed__1() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Tactic_FunInd_rwIfWith___closed__13; -x_2 = l_Lean_Name_mkStr1(x_1); +x_1 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___closed__0; +x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Tactic_FunInd_rwIfWith___closed__15() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Bool", 4, 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwIfWith___closed__16() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("true", 4, 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwIfWith___closed__17() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Tactic_FunInd_rwIfWith___closed__16; -x_2 = l_Lean_Tactic_FunInd_rwIfWith___closed__15; -x_3 = l_Lean_Name_mkStr2(x_2, x_1); -return x_3; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwIfWith___closed__18() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Tactic_FunInd_rwIfWith___closed__17; -x_3 = l_Lean_mkConst(x_2, x_1); -return x_3; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwIfWith___closed__19() { +static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___closed__2() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("false", 5, 5); +x_1 = lean_mk_string_unchecked("not a let expression or `", 25, 25); return x_1; } } -static lean_object* _init_l_Lean_Tactic_FunInd_rwIfWith___closed__20() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Tactic_FunInd_rwIfWith___closed__19; -x_2 = l_Lean_Tactic_FunInd_rwIfWith___closed__15; -x_3 = l_Lean_Name_mkStr2(x_2, x_1); -return x_3; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwIfWith___closed__21() { +static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___closed__3() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Tactic_FunInd_rwIfWith___closed__20; -x_3 = l_Lean_mkConst(x_2, x_1); -return x_3; +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -static lean_object* _init_l_Lean_Tactic_FunInd_rwIfWith___closed__22() { +static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("cond_neg", 8, 8); +x_1 = lean_mk_string_unchecked("` is not definitionally equal to `", 34, 34); return x_1; } } -static lean_object* _init_l_Lean_Tactic_FunInd_rwIfWith___closed__23() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Tactic_FunInd_rwIfWith___closed__22; -x_2 = l_Lean_Tactic_FunInd_rwIfWith___closed__15; -x_3 = l_Lean_Name_mkStr2(x_2, x_1); -return x_3; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwIfWith___closed__24() { +static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(5u); -x_2 = lean_mk_empty_array_with_capacity(x_1); +x_1 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Tactic_FunInd_rwIfWith___closed__25() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("cond_pos", 8, 8); -return x_1; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwIfWith___closed__26() { -_start: +lean_object* x_8; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_43; +x_43 = l_Lean_Expr_isLet(x_2); +if (x_43 == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Tactic_FunInd_rwIfWith___closed__25; -x_2 = l_Lean_Tactic_FunInd_rwIfWith___closed__15; -x_3 = l_Lean_Name_mkStr2(x_2, x_1); -return x_3; -} +x_14 = x_3; +x_15 = x_4; +x_16 = x_5; +x_17 = x_6; +x_18 = lean_box(0); +goto block_42; } -LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_rwIfWith(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +else { -lean_object* x_8; -lean_inc_ref(x_2); -x_8 = l_Lean_Meta_instantiateMVarsIfMVarApp___redArg(x_2, x_4); -if (lean_obj_tag(x_8) == 0) +lean_object* x_44; lean_object* x_45; +x_44 = l_Lean_Expr_letValue_x21(x_2); +lean_inc(x_6); +lean_inc_ref(x_5); +lean_inc(x_4); +lean_inc_ref(x_3); +lean_inc_ref(x_1); +x_45 = l_Lean_Meta_isExprDefEq(x_44, x_1, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_45) == 0) { -lean_object* x_9; lean_object* x_10; lean_object* x_16; uint8_t x_17; -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -if (lean_is_exclusive(x_8)) { - lean_ctor_release(x_8, 0); - x_10 = x_8; -} else { - lean_dec_ref(x_8); - x_10 = lean_box(0); +uint8_t x_46; +x_46 = !lean_is_exclusive(x_45); +if (x_46 == 0) +{ +lean_object* x_47; uint8_t x_48; +x_47 = lean_ctor_get(x_45, 0); +x_48 = lean_unbox(x_47); +if (x_48 == 0) +{ +lean_free_object(x_45); +lean_dec(x_47); +x_14 = x_3; +x_15 = x_4; +x_16 = x_5; +x_17 = x_6; +x_18 = lean_box(0); +goto block_42; } -x_16 = l_Lean_Expr_cleanupAnnotations(x_9); -x_17 = l_Lean_Expr_isApp(x_16); -if (x_17 == 0) +else { -lean_dec_ref(x_16); +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; lean_dec(x_6); lean_dec_ref(x_5); lean_dec(x_4); lean_dec_ref(x_3); +x_49 = l_Lean_Expr_letBody_x21(x_2); +lean_dec_ref(x_2); +x_50 = lean_expr_instantiate1(x_49, x_1); lean_dec_ref(x_1); -goto block_15; +lean_dec_ref(x_49); +x_51 = lean_box(0); +x_52 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_51); +x_53 = lean_unbox(x_47); +lean_dec(x_47); +lean_ctor_set_uint8(x_52, sizeof(void*)*2, x_53); +lean_ctor_set(x_45, 0, x_52); +return x_45; +} } else { -lean_object* x_18; uint8_t x_19; -lean_inc_ref(x_16); -x_18 = l_Lean_Expr_appFnCleanup___redArg(x_16); -x_19 = l_Lean_Expr_isApp(x_18); -if (x_19 == 0) +lean_object* x_54; uint8_t x_55; +x_54 = lean_ctor_get(x_45, 0); +lean_inc(x_54); +lean_dec(x_45); +x_55 = lean_unbox(x_54); +if (x_55 == 0) { -lean_dec_ref(x_18); -lean_dec_ref(x_16); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -lean_dec_ref(x_1); -goto block_15; +lean_dec(x_54); +x_14 = x_3; +x_15 = x_4; +x_16 = x_5; +x_17 = x_6; +x_18 = lean_box(0); +goto block_42; } else { -lean_object* x_20; uint8_t x_21; -lean_inc_ref(x_18); -x_20 = l_Lean_Expr_appFnCleanup___redArg(x_18); -x_21 = l_Lean_Expr_isApp(x_20); -if (x_21 == 0) -{ -lean_dec_ref(x_20); -lean_dec_ref(x_18); -lean_dec_ref(x_16); +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; lean_object* x_61; lean_dec(x_6); lean_dec_ref(x_5); lean_dec(x_4); lean_dec_ref(x_3); +x_56 = l_Lean_Expr_letBody_x21(x_2); +lean_dec_ref(x_2); +x_57 = lean_expr_instantiate1(x_56, x_1); lean_dec_ref(x_1); -goto block_15; +lean_dec_ref(x_56); +x_58 = lean_box(0); +x_59 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_58); +x_60 = lean_unbox(x_54); +lean_dec(x_54); +lean_ctor_set_uint8(x_59, sizeof(void*)*2, x_60); +x_61 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_61, 0, x_59); +return x_61; +} +} } else { -lean_object* x_22; uint8_t x_23; -lean_inc_ref(x_20); -x_22 = l_Lean_Expr_appFnCleanup___redArg(x_20); -x_23 = l_Lean_Expr_isApp(x_22); -if (x_23 == 0) -{ -lean_dec_ref(x_22); -lean_dec_ref(x_20); -lean_dec_ref(x_18); -lean_dec_ref(x_16); +uint8_t x_62; lean_dec(x_6); lean_dec_ref(x_5); lean_dec(x_4); lean_dec_ref(x_3); +lean_dec_ref(x_2); lean_dec_ref(x_1); -goto block_15; +x_62 = !lean_is_exclusive(x_45); +if (x_62 == 0) +{ +return x_45; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; -x_24 = lean_ctor_get(x_16, 1); -lean_inc_ref(x_24); -lean_dec_ref(x_16); -x_25 = lean_ctor_get(x_18, 1); -lean_inc_ref(x_25); -lean_dec_ref(x_18); -x_26 = lean_ctor_get(x_20, 1); -lean_inc_ref(x_26); -lean_dec_ref(x_20); -x_27 = lean_ctor_get(x_22, 1); -lean_inc_ref(x_27); -x_28 = l_Lean_Expr_appFnCleanup___redArg(x_22); -x_29 = l_Lean_Tactic_FunInd_rwIfWith___closed__1; -x_30 = l_Lean_Expr_isConstOf(x_28, x_29); -if (x_30 == 0) -{ -uint8_t x_31; -x_31 = l_Lean_Expr_isApp(x_28); -if (x_31 == 0) -{ -lean_dec_ref(x_28); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -lean_dec_ref(x_1); -goto block_15; +lean_object* x_63; lean_object* x_64; +x_63 = lean_ctor_get(x_45, 0); +lean_inc(x_63); +lean_dec(x_45); +x_64 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_64, 0, x_63); +return x_64; } -else +} +} +block_13: { -lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; -x_32 = lean_ctor_get(x_28, 1); -lean_inc_ref(x_32); -x_33 = l_Lean_Expr_appFnCleanup___redArg(x_28); -x_34 = l_Lean_Tactic_FunInd_rwIfWith___closed__3; -x_35 = l_Lean_Expr_isConstOf(x_33, x_34); -if (x_35 == 0) +lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_box(0); +x_10 = 1; +x_11 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_11, 0, x_2); +lean_ctor_set(x_11, 1, x_9); +lean_ctor_set_uint8(x_11, sizeof(void*)*2, x_10); +x_12 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_12, 0, x_11); +return x_12; +} +block_42: { -lean_object* x_36; uint8_t x_37; -x_36 = l_Lean_Tactic_FunInd_rwIfWith___closed__5; -x_37 = l_Lean_Expr_isConstOf(x_33, x_36); -if (x_37 == 0) +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_19 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__21___closed__2; +x_20 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__0___redArg(x_19, x_16); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +lean_dec_ref(x_20); +x_22 = lean_unbox(x_21); +lean_dec(x_21); +if (x_22 == 0) { -lean_dec_ref(x_33); -lean_dec_ref(x_32); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); +lean_dec(x_17); +lean_dec_ref(x_16); +lean_dec(x_15); +lean_dec_ref(x_14); lean_dec_ref(x_1); -goto block_15; +x_8 = lean_box(0); +goto block_13; } else { -lean_object* x_38; -lean_dec(x_10); -lean_inc(x_6); -lean_inc_ref(x_5); -lean_inc(x_4); -lean_inc_ref(x_3); -lean_inc_ref(x_1); -x_38 = lean_infer_type(x_1, x_3, x_4, x_5, x_6); +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_23 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___closed__1; +x_24 = lean_unsigned_to_nat(30u); +lean_inc_ref(x_2); +x_25 = l_Lean_inlineExpr(x_2, x_24); +x_26 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_26, 0, x_23); +lean_ctor_set(x_26, 1, x_25); +x_27 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___closed__3; +x_28 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +x_29 = l_Lean_MessageData_ofExpr(x_1); +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +x_31 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___closed__5; +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +x_33 = l_Lean_Expr_letValue_x21(x_2); +x_34 = l_Lean_MessageData_ofExpr(x_33); +x_35 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_35, 0, x_32); +lean_ctor_set(x_35, 1, x_34); +x_36 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___redArg___lam__3___closed__3; +x_37 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +x_38 = l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__1(x_19, x_37, x_14, x_15, x_16, x_17); +lean_dec(x_17); +lean_dec_ref(x_16); +lean_dec(x_15); +lean_dec_ref(x_14); if (lean_obj_tag(x_38) == 0) { -lean_object* x_39; lean_object* x_40; -x_39 = lean_ctor_get(x_38, 0); -lean_inc(x_39); lean_dec_ref(x_38); -lean_inc(x_6); -lean_inc_ref(x_5); -lean_inc(x_4); -lean_inc_ref(x_3); -lean_inc_ref(x_27); -x_40 = l_Lean_Meta_isExprDefEq(x_27, x_39, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_40) == 0) -{ -uint8_t x_41; -x_41 = !lean_is_exclusive(x_40); -if (x_41 == 0) -{ -lean_object* x_42; lean_object* x_43; uint8_t x_44; -x_42 = lean_ctor_get(x_40, 0); -x_43 = l_Lean_Expr_constLevels_x21(x_33); -lean_dec_ref(x_33); -x_44 = lean_unbox(x_42); -lean_dec(x_42); -if (x_44 == 0) -{ -lean_object* x_45; -lean_free_object(x_40); -lean_inc(x_6); -lean_inc_ref(x_5); -lean_inc(x_4); -lean_inc_ref(x_3); -lean_inc_ref(x_1); -x_45 = lean_infer_type(x_1, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_45) == 0) -{ -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_45, 0); -lean_inc(x_46); -lean_dec_ref(x_45); -lean_inc_ref(x_27); -x_47 = l_Lean_mkNot(x_27); -x_48 = l_Lean_Meta_isExprDefEq(x_47, x_46, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_48) == 0) -{ -uint8_t x_49; -x_49 = !lean_is_exclusive(x_48); -if (x_49 == 0) -{ -lean_object* x_50; uint8_t x_51; -x_50 = lean_ctor_get(x_48, 0); -x_51 = lean_unbox(x_50); -lean_dec(x_50); -if (x_51 == 0) -{ -lean_object* x_52; lean_object* x_53; -lean_dec(x_43); -lean_dec_ref(x_32); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); -lean_dec_ref(x_1); -x_52 = lean_box(0); -x_53 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_53, 0, x_2); -lean_ctor_set(x_53, 1, x_52); -lean_ctor_set_uint8(x_53, sizeof(void*)*2, x_37); -lean_ctor_set(x_48, 0, x_53); -return x_48; +x_8 = lean_box(0); +goto block_13; } else { -lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +uint8_t x_39; lean_dec_ref(x_2); -x_54 = l_Lean_Tactic_FunInd_rwIfWith___closed__7; -x_55 = l_Lean_mkConst(x_54, x_43); -x_56 = l_Lean_Tactic_FunInd_rwIfWith___closed__8; -x_57 = lean_array_push(x_56, x_27); -x_58 = lean_array_push(x_57, x_26); -x_59 = lean_array_push(x_58, x_1); -x_60 = lean_array_push(x_59, x_32); -x_61 = lean_array_push(x_60, x_25); -lean_inc_ref(x_24); -x_62 = lean_array_push(x_61, x_24); -x_63 = l_Lean_mkAppN(x_55, x_62); -lean_dec_ref(x_62); -x_64 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_64, 0, x_63); -x_65 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_65, 0, x_24); -lean_ctor_set(x_65, 1, x_64); -lean_ctor_set_uint8(x_65, sizeof(void*)*2, x_37); -lean_ctor_set(x_48, 0, x_65); -return x_48; -} -} -else -{ -lean_object* x_66; uint8_t x_67; -x_66 = lean_ctor_get(x_48, 0); -lean_inc(x_66); -lean_dec(x_48); -x_67 = lean_unbox(x_66); -lean_dec(x_66); -if (x_67 == 0) +x_39 = !lean_is_exclusive(x_38); +if (x_39 == 0) { -lean_object* x_68; lean_object* x_69; lean_object* x_70; -lean_dec(x_43); -lean_dec_ref(x_32); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); -lean_dec_ref(x_1); -x_68 = lean_box(0); -x_69 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_69, 0, x_2); -lean_ctor_set(x_69, 1, x_68); -lean_ctor_set_uint8(x_69, sizeof(void*)*2, x_37); -x_70 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_70, 0, x_69); -return x_70; +return x_38; } else { -lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; -lean_dec_ref(x_2); -x_71 = l_Lean_Tactic_FunInd_rwIfWith___closed__7; -x_72 = l_Lean_mkConst(x_71, x_43); -x_73 = l_Lean_Tactic_FunInd_rwIfWith___closed__8; -x_74 = lean_array_push(x_73, x_27); -x_75 = lean_array_push(x_74, x_26); -x_76 = lean_array_push(x_75, x_1); -x_77 = lean_array_push(x_76, x_32); -x_78 = lean_array_push(x_77, x_25); -lean_inc_ref(x_24); -x_79 = lean_array_push(x_78, x_24); -x_80 = l_Lean_mkAppN(x_72, x_79); -lean_dec_ref(x_79); -x_81 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_81, 0, x_80); -x_82 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_82, 0, x_24); -lean_ctor_set(x_82, 1, x_81); -lean_ctor_set_uint8(x_82, sizeof(void*)*2, x_37); -x_83 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_83, 0, x_82); -return x_83; -} +lean_object* x_40; lean_object* x_41; +x_40 = lean_ctor_get(x_38, 0); +lean_inc(x_40); +lean_dec(x_38); +x_41 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_41, 0, x_40); +return x_41; } } -else -{ -uint8_t x_84; -lean_dec(x_43); -lean_dec_ref(x_32); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_84 = !lean_is_exclusive(x_48); -if (x_84 == 0) -{ -return x_48; } -else -{ -lean_object* x_85; lean_object* x_86; -x_85 = lean_ctor_get(x_48, 0); -lean_inc(x_85); -lean_dec(x_48); -x_86 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_86, 0, x_85); -return x_86; } } } -else +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__0___boxed(lean_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_87; -lean_dec(x_43); -lean_dec_ref(x_32); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); +lean_object* x_7; +x_7 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__0(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec(x_3); lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_87 = !lean_is_exclusive(x_45); -if (x_87 == 0) -{ -return x_45; +return x_7; } -else -{ -lean_object* x_88; lean_object* x_89; -x_88 = lean_ctor_get(x_45, 0); -lean_inc(x_88); -lean_dec(x_45); -x_89 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_89, 0, x_88); -return x_89; } +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__0___redArg(x_1, x_2); +lean_dec_ref(x_2); +return x_4; } } -else +LEAN_EXPORT lean_object* l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; +lean_object* x_8; +x_8 = l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__1(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_6); lean_dec_ref(x_5); lean_dec(x_4); lean_dec_ref(x_3); -lean_dec_ref(x_2); -x_90 = l_Lean_Tactic_FunInd_rwIfWith___closed__10; -x_91 = l_Lean_mkConst(x_90, x_43); -x_92 = l_Lean_Tactic_FunInd_rwIfWith___closed__8; -x_93 = lean_array_push(x_92, x_27); -x_94 = lean_array_push(x_93, x_26); -x_95 = lean_array_push(x_94, x_1); -x_96 = lean_array_push(x_95, x_32); -lean_inc_ref(x_25); -x_97 = lean_array_push(x_96, x_25); -x_98 = lean_array_push(x_97, x_24); -x_99 = l_Lean_mkAppN(x_91, x_98); -lean_dec_ref(x_98); -x_100 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_100, 0, x_99); -x_101 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_101, 0, x_25); -lean_ctor_set(x_101, 1, x_100); -lean_ctor_set_uint8(x_101, sizeof(void*)*2, x_37); -lean_ctor_set(x_40, 0, x_101); -return x_40; +return x_8; } } -else -{ -lean_object* x_102; lean_object* x_103; uint8_t x_104; -x_102 = lean_ctor_get(x_40, 0); -lean_inc(x_102); -lean_dec(x_40); -x_103 = l_Lean_Expr_constLevels_x21(x_33); -lean_dec_ref(x_33); -x_104 = lean_unbox(x_102); -lean_dec(x_102); -if (x_104 == 0) -{ -lean_object* x_105; -lean_inc(x_6); -lean_inc_ref(x_5); -lean_inc(x_4); -lean_inc_ref(x_3); -lean_inc_ref(x_1); -x_105 = lean_infer_type(x_1, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_105) == 0) -{ -lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_106 = lean_ctor_get(x_105, 0); -lean_inc(x_106); -lean_dec_ref(x_105); -lean_inc_ref(x_27); -x_107 = l_Lean_mkNot(x_27); -x_108 = l_Lean_Meta_isExprDefEq(x_107, x_106, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_108) == 0) +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___boxed(lean_object* x_1, lean_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_109; lean_object* x_110; uint8_t x_111; -x_109 = lean_ctor_get(x_108, 0); -lean_inc(x_109); -if (lean_is_exclusive(x_108)) { - lean_ctor_release(x_108, 0); - x_110 = x_108; -} else { - lean_dec_ref(x_108); - x_110 = lean_box(0); +lean_object* x_8; +x_8 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith(x_1, x_2, x_3, x_4, x_5, x_6); +return x_8; } -x_111 = lean_unbox(x_109); -lean_dec(x_109); -if (x_111 == 0) +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwMData___redArg(lean_object* x_1) { +_start: { -lean_object* x_112; lean_object* x_113; lean_object* x_114; -lean_dec(x_103); -lean_dec_ref(x_32); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); -lean_dec_ref(x_1); -x_112 = lean_box(0); -x_113 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_113, 0, x_2); -lean_ctor_set(x_113, 1, x_112); -lean_ctor_set_uint8(x_113, sizeof(void*)*2, x_37); -if (lean_is_scalar(x_110)) { - x_114 = lean_alloc_ctor(0, 1, 0); -} else { - x_114 = x_110; +lean_object* x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; +x_3 = l_Lean_Expr_consumeMData(x_1); +x_4 = lean_box(0); +x_5 = 1; +x_6 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_6, 0, x_3); +lean_ctor_set(x_6, 1, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*2, x_5); +x_7 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_7, 0, x_6); +return x_7; } -lean_ctor_set(x_114, 0, x_113); -return x_114; } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwMData(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; -lean_dec_ref(x_2); -x_115 = l_Lean_Tactic_FunInd_rwIfWith___closed__7; -x_116 = l_Lean_mkConst(x_115, x_103); -x_117 = l_Lean_Tactic_FunInd_rwIfWith___closed__8; -x_118 = lean_array_push(x_117, x_27); -x_119 = lean_array_push(x_118, x_26); -x_120 = lean_array_push(x_119, x_1); -x_121 = lean_array_push(x_120, x_32); -x_122 = lean_array_push(x_121, x_25); -lean_inc_ref(x_24); -x_123 = lean_array_push(x_122, x_24); -x_124 = l_Lean_mkAppN(x_116, x_123); -lean_dec_ref(x_123); -x_125 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_125, 0, x_124); -x_126 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_126, 0, x_24); -lean_ctor_set(x_126, 1, x_125); -lean_ctor_set_uint8(x_126, sizeof(void*)*2, x_37); -if (lean_is_scalar(x_110)) { - x_127 = lean_alloc_ctor(0, 1, 0); -} else { - x_127 = x_110; -} -lean_ctor_set(x_127, 0, x_126); -return x_127; +lean_object* x_7; +x_7 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwMData___redArg(x_1); +return x_7; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwMData___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -lean_object* x_128; lean_object* x_129; lean_object* x_130; -lean_dec(x_103); -lean_dec_ref(x_32); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); +lean_object* x_7; +x_7 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwMData(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_128 = lean_ctor_get(x_108, 0); -lean_inc(x_128); -if (lean_is_exclusive(x_108)) { - lean_ctor_release(x_108, 0); - x_129 = x_108; -} else { - lean_dec_ref(x_108); - x_129 = lean_box(0); -} -if (lean_is_scalar(x_129)) { - x_130 = lean_alloc_ctor(1, 1, 0); -} else { - x_130 = x_129; -} -lean_ctor_set(x_130, 0, x_128); -return x_130; +return x_7; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwMData___redArg___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_131; lean_object* x_132; lean_object* x_133; -lean_dec(x_103); -lean_dec_ref(x_32); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -lean_dec_ref(x_2); +lean_object* x_3; +x_3 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwMData___redArg(x_1); lean_dec_ref(x_1); -x_131 = lean_ctor_get(x_105, 0); -lean_inc(x_131); -if (lean_is_exclusive(x_105)) { - lean_ctor_release(x_105, 0); - x_132 = x_105; -} else { - lean_dec_ref(x_105); - x_132 = lean_box(0); -} -if (lean_is_scalar(x_132)) { - x_133 = lean_alloc_ctor(1, 1, 0); -} else { - x_133 = x_132; -} -lean_ctor_set(x_133, 0, x_131); -return x_133; +return x_3; } } -else +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { +_start: { -lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -lean_dec_ref(x_2); -x_134 = l_Lean_Tactic_FunInd_rwIfWith___closed__10; -x_135 = l_Lean_mkConst(x_134, x_103); -x_136 = l_Lean_Tactic_FunInd_rwIfWith___closed__8; -x_137 = lean_array_push(x_136, x_27); -x_138 = lean_array_push(x_137, x_26); -x_139 = lean_array_push(x_138, x_1); -x_140 = lean_array_push(x_139, x_32); -lean_inc_ref(x_25); -x_141 = lean_array_push(x_140, x_25); -x_142 = lean_array_push(x_141, x_24); -x_143 = l_Lean_mkAppN(x_135, x_142); -lean_dec_ref(x_142); -x_144 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_144, 0, x_143); -x_145 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_145, 0, x_25); -lean_ctor_set(x_145, 1, x_144); -lean_ctor_set_uint8(x_145, sizeof(void*)*2, x_37); -x_146 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_146, 0, x_145); -return x_146; -} -} +uint8_t x_8; +x_8 = lean_usize_dec_lt(x_6, x_5); +if (x_8 == 0) +{ +lean_inc_ref(x_7); +return x_7; } else { -uint8_t x_147; -lean_dec_ref(x_33); -lean_dec_ref(x_32); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_147 = !lean_is_exclusive(x_40); -if (x_147 == 0) +lean_object* x_9; uint8_t x_10; +x_9 = lean_array_uget(x_4, x_6); +x_10 = l_Lean_Expr_isConstOf(x_1, x_9); +if (x_10 == 0) { -return x_40; +size_t x_11; size_t x_12; +lean_dec(x_9); +x_11 = 1; +x_12 = lean_usize_add(x_6, x_11); +{ +size_t _tmp_5 = x_12; +lean_object* _tmp_6 = x_2; +x_6 = _tmp_5; +x_7 = _tmp_6; +} +goto _start; } else { -lean_object* x_148; lean_object* x_149; -x_148 = lean_ctor_get(x_40, 0); -lean_inc(x_148); -lean_dec(x_40); -x_149 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_149, 0, x_148); -return x_149; +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_9); +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_14); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_3); +return x_16; } } } -else +} +static lean_object* _init_l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__0() { +_start: { -uint8_t x_150; -lean_dec_ref(x_33); -lean_dec_ref(x_32); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_150 = !lean_is_exclusive(x_38); -if (x_150 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Eq", 2, 2); +return x_1; +} +} +static lean_object* _init_l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__1() { +_start: { -return x_38; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__0; +x_2 = l_Lean_Name_mkStr1(x_1); +return x_2; } -else +} +static lean_object* _init_l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__2() { +_start: { -lean_object* x_151; lean_object* x_152; -x_151 = lean_ctor_get(x_38, 0); -lean_inc(x_151); -lean_dec(x_38); -x_152 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_152, 0, x_151); -return x_152; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Not an equality: `", 18, 18); +return x_1; } } +static lean_object* _init_l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -else +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -lean_object* x_153; -lean_dec(x_10); -lean_inc(x_6); -lean_inc_ref(x_5); -lean_inc(x_4); -lean_inc_ref(x_3); -lean_inc_ref(x_1); -x_153 = lean_infer_type(x_1, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_153) == 0) +lean_object* x_11; +if (lean_obj_tag(x_3) == 5) { -lean_object* x_154; lean_object* x_155; -x_154 = lean_ctor_get(x_153, 0); -lean_inc(x_154); -lean_dec_ref(x_153); -lean_inc(x_6); -lean_inc_ref(x_5); -lean_inc(x_4); -lean_inc_ref(x_3); -lean_inc_ref(x_27); -x_155 = l_Lean_Meta_isExprDefEq(x_27, x_154, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_155) == 0) +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_17 = lean_ctor_get(x_3, 0); +lean_inc_ref(x_17); +x_18 = lean_ctor_get(x_3, 1); +lean_inc_ref(x_18); +lean_dec_ref(x_3); +x_19 = lean_array_set(x_4, x_5, x_18); +x_20 = lean_unsigned_to_nat(1u); +x_21 = lean_nat_sub(x_5, x_20); +lean_dec(x_5); +x_3 = x_17; +x_4 = x_19; +x_5 = x_21; +goto _start; +} +else { -uint8_t x_156; -x_156 = !lean_is_exclusive(x_155); -if (x_156 == 0) +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; +lean_dec(x_5); +x_23 = lean_box(0); +x_24 = lean_box(0); +x_25 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanupAfter_cleanupAfter_x3f___lam__0___closed__0; +x_26 = lean_array_size(x_2); +x_27 = 0; +x_28 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__0(x_3, x_25, x_24, x_2, x_26, x_27, x_25); +x_29 = !lean_is_exclusive(x_28); +if (x_29 == 0) { -lean_object* x_157; lean_object* x_158; uint8_t x_159; -x_157 = lean_ctor_get(x_155, 0); -x_158 = l_Lean_Expr_constLevels_x21(x_33); -lean_dec_ref(x_33); -x_159 = lean_unbox(x_157); -lean_dec(x_157); -if (x_159 == 0) +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_28, 0); +x_31 = lean_ctor_get(x_28, 1); +lean_dec(x_31); +if (lean_obj_tag(x_30) == 0) { -lean_object* x_160; -lean_free_object(x_155); -lean_inc(x_6); -lean_inc_ref(x_5); -lean_inc(x_4); -lean_inc_ref(x_3); -lean_inc_ref(x_1); -x_160 = lean_infer_type(x_1, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_160) == 0) +lean_free_object(x_28); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec_ref(x_4); +lean_dec_ref(x_3); +x_11 = lean_box(0); +goto block_16; +} +else { -lean_object* x_161; lean_object* x_162; lean_object* x_163; -x_161 = lean_ctor_get(x_160, 0); -lean_inc(x_161); -lean_dec_ref(x_160); -lean_inc_ref(x_27); -x_162 = l_Lean_mkNot(x_27); -x_163 = l_Lean_Meta_isExprDefEq(x_162, x_161, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_163) == 0) +lean_object* x_32; +x_32 = lean_ctor_get(x_30, 0); +lean_inc(x_32); +lean_dec_ref(x_30); +if (lean_obj_tag(x_32) == 1) { -uint8_t x_164; -x_164 = !lean_is_exclusive(x_163); -if (x_164 == 0) +lean_object* x_33; uint8_t x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +lean_dec_ref(x_32); +x_34 = 1; +lean_inc(x_9); +lean_inc_ref(x_8); +lean_inc(x_7); +lean_inc_ref(x_6); +x_35 = l_Lean_Meta_getUnfoldEqnFor_x3f(x_33, x_34, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_35) == 0) { -lean_object* x_165; uint8_t x_166; -x_165 = lean_ctor_get(x_163, 0); -x_166 = lean_unbox(x_165); -lean_dec(x_165); -if (x_166 == 0) +uint8_t x_36; +x_36 = !lean_is_exclusive(x_35); +if (x_36 == 0) { -lean_object* x_167; lean_object* x_168; -lean_dec(x_158); -lean_dec_ref(x_32); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); +lean_object* x_37; +x_37 = lean_ctor_get(x_35, 0); +if (lean_obj_tag(x_37) == 1) +{ +uint8_t x_38; +lean_free_object(x_35); lean_dec_ref(x_1); -x_167 = lean_box(0); -x_168 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_168, 0, x_2); -lean_ctor_set(x_168, 1, x_167); -lean_ctor_set_uint8(x_168, sizeof(void*)*2, x_35); -lean_ctor_set(x_163, 0, x_168); -return x_163; -} -else +x_38 = !lean_is_exclusive(x_37); +if (x_38 == 0) { -lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; -lean_dec_ref(x_2); -x_169 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__6___closed__0; -lean_inc_ref(x_1); -x_170 = lean_array_push(x_169, x_1); -lean_inc_ref(x_24); -x_171 = l_Lean_Expr_beta(x_24, x_170); -x_172 = l_Lean_Tactic_FunInd_rwIfWith___closed__12; -x_173 = l_Lean_mkConst(x_172, x_158); -x_174 = l_Lean_Tactic_FunInd_rwIfWith___closed__8; -x_175 = lean_array_push(x_174, x_27); -x_176 = lean_array_push(x_175, x_26); -x_177 = lean_array_push(x_176, x_1); -x_178 = lean_array_push(x_177, x_32); -x_179 = lean_array_push(x_178, x_25); -x_180 = lean_array_push(x_179, x_24); -x_181 = l_Lean_mkAppN(x_173, x_180); -lean_dec_ref(x_180); -x_182 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_182, 0, x_181); -x_183 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_183, 0, x_171); -lean_ctor_set(x_183, 1, x_182); -lean_ctor_set_uint8(x_183, sizeof(void*)*2, x_35); -lean_ctor_set(x_163, 0, x_183); -return x_163; -} -} -else +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_39 = lean_ctor_get(x_37, 0); +x_40 = l_Lean_Expr_constLevels_x21(x_3); +lean_dec_ref(x_3); +x_41 = l_Lean_mkConst(x_39, x_40); +x_42 = l_Lean_mkAppN(x_41, x_4); +lean_dec_ref(x_4); +lean_inc(x_9); +lean_inc_ref(x_8); +lean_inc(x_7); +lean_inc_ref(x_6); +lean_inc_ref(x_42); +x_43 = lean_infer_type(x_42, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_43) == 0) { -lean_object* x_184; uint8_t x_185; -x_184 = lean_ctor_get(x_163, 0); -lean_inc(x_184); -lean_dec(x_163); -x_185 = lean_unbox(x_184); -lean_dec(x_184); -if (x_185 == 0) +uint8_t x_44; +x_44 = !lean_is_exclusive(x_43); +if (x_44 == 0) { -lean_object* x_186; lean_object* x_187; lean_object* x_188; -lean_dec(x_158); -lean_dec_ref(x_32); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); -lean_dec_ref(x_1); -x_186 = lean_box(0); -x_187 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_187, 0, x_2); -lean_ctor_set(x_187, 1, x_186); -lean_ctor_set_uint8(x_187, sizeof(void*)*2, x_35); -x_188 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_188, 0, x_187); -return x_188; +lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; +x_45 = lean_ctor_get(x_43, 0); +x_46 = l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__1; +x_47 = lean_unsigned_to_nat(3u); +x_48 = l_Lean_Expr_isAppOfArity(x_45, x_46, x_47); +if (x_48 == 0) +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +lean_free_object(x_43); +lean_dec(x_45); +lean_free_object(x_37); +x_49 = l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__3; +x_50 = l_Lean_MessageData_ofExpr(x_42); +lean_ctor_set_tag(x_28, 7); +lean_ctor_set(x_28, 1, x_50); +lean_ctor_set(x_28, 0, x_49); +x_51 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___redArg___lam__3___closed__3; +x_52 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_52, 0, x_28); +lean_ctor_set(x_52, 1, x_51); +x_53 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___redArg(x_52, x_6, x_7, x_8, x_9); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); +return x_53; } else { -lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; -lean_dec_ref(x_2); -x_189 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__6___closed__0; -lean_inc_ref(x_1); -x_190 = lean_array_push(x_189, x_1); -lean_inc_ref(x_24); -x_191 = l_Lean_Expr_beta(x_24, x_190); -x_192 = l_Lean_Tactic_FunInd_rwIfWith___closed__12; -x_193 = l_Lean_mkConst(x_192, x_158); -x_194 = l_Lean_Tactic_FunInd_rwIfWith___closed__8; -x_195 = lean_array_push(x_194, x_27); -x_196 = lean_array_push(x_195, x_26); -x_197 = lean_array_push(x_196, x_1); -x_198 = lean_array_push(x_197, x_32); -x_199 = lean_array_push(x_198, x_25); -x_200 = lean_array_push(x_199, x_24); -x_201 = l_Lean_mkAppN(x_193, x_200); -lean_dec_ref(x_200); -x_202 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_202, 0, x_201); -x_203 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_203, 0, x_191); -lean_ctor_set(x_203, 1, x_202); -lean_ctor_set_uint8(x_203, sizeof(void*)*2, x_35); -x_204 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_204, 0, x_203); -return x_204; -} +lean_object* x_54; lean_object* x_55; +lean_free_object(x_28); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); +x_54 = l_Lean_Expr_appArg_x21(x_45); +lean_dec(x_45); +lean_ctor_set(x_37, 0, x_42); +x_55 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_37); +lean_ctor_set_uint8(x_55, sizeof(void*)*2, x_34); +lean_ctor_set(x_43, 0, x_55); +return x_43; } } else { -uint8_t x_205; -lean_dec(x_158); -lean_dec_ref(x_32); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_205 = !lean_is_exclusive(x_163); -if (x_205 == 0) +lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; +x_56 = lean_ctor_get(x_43, 0); +lean_inc(x_56); +lean_dec(x_43); +x_57 = l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__1; +x_58 = lean_unsigned_to_nat(3u); +x_59 = l_Lean_Expr_isAppOfArity(x_56, x_57, x_58); +if (x_59 == 0) { -return x_163; +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +lean_dec(x_56); +lean_free_object(x_37); +x_60 = l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__3; +x_61 = l_Lean_MessageData_ofExpr(x_42); +lean_ctor_set_tag(x_28, 7); +lean_ctor_set(x_28, 1, x_61); +lean_ctor_set(x_28, 0, x_60); +x_62 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___redArg___lam__3___closed__3; +x_63 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_63, 0, x_28); +lean_ctor_set(x_63, 1, x_62); +x_64 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___redArg(x_63, x_6, x_7, x_8, x_9); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); +return x_64; } else { -lean_object* x_206; lean_object* x_207; -x_206 = lean_ctor_get(x_163, 0); -lean_inc(x_206); -lean_dec(x_163); -x_207 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_207, 0, x_206); -return x_207; +lean_object* x_65; lean_object* x_66; lean_object* x_67; +lean_free_object(x_28); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); +x_65 = l_Lean_Expr_appArg_x21(x_56); +lean_dec(x_56); +lean_ctor_set(x_37, 0, x_42); +x_66 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_66, 0, x_65); +lean_ctor_set(x_66, 1, x_37); +lean_ctor_set_uint8(x_66, sizeof(void*)*2, x_34); +x_67 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_67, 0, x_66); +return x_67; } } } else { -uint8_t x_208; -lean_dec(x_158); -lean_dec_ref(x_32); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_208 = !lean_is_exclusive(x_160); -if (x_208 == 0) +uint8_t x_68; +lean_dec_ref(x_42); +lean_free_object(x_37); +lean_free_object(x_28); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); +x_68 = !lean_is_exclusive(x_43); +if (x_68 == 0) { -return x_160; +return x_43; } else { -lean_object* x_209; lean_object* x_210; -x_209 = lean_ctor_get(x_160, 0); -lean_inc(x_209); -lean_dec(x_160); -x_210 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_210, 0, x_209); -return x_210; +lean_object* x_69; lean_object* x_70; +x_69 = lean_ctor_get(x_43, 0); +lean_inc(x_69); +lean_dec(x_43); +x_70 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_70, 0, x_69); +return x_70; } } } else { -lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_71 = lean_ctor_get(x_37, 0); +lean_inc(x_71); +lean_dec(x_37); +x_72 = l_Lean_Expr_constLevels_x21(x_3); lean_dec_ref(x_3); -lean_dec_ref(x_2); -x_211 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__6___closed__0; -lean_inc_ref(x_1); -x_212 = lean_array_push(x_211, x_1); -lean_inc_ref(x_25); -x_213 = l_Lean_Expr_beta(x_25, x_212); -x_214 = l_Lean_Tactic_FunInd_rwIfWith___closed__14; -x_215 = l_Lean_mkConst(x_214, x_158); -x_216 = l_Lean_Tactic_FunInd_rwIfWith___closed__8; -x_217 = lean_array_push(x_216, x_27); -x_218 = lean_array_push(x_217, x_26); -x_219 = lean_array_push(x_218, x_1); -x_220 = lean_array_push(x_219, x_32); -x_221 = lean_array_push(x_220, x_25); -x_222 = lean_array_push(x_221, x_24); -x_223 = l_Lean_mkAppN(x_215, x_222); -lean_dec_ref(x_222); -x_224 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_224, 0, x_223); -x_225 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_225, 0, x_213); -lean_ctor_set(x_225, 1, x_224); -lean_ctor_set_uint8(x_225, sizeof(void*)*2, x_35); -lean_ctor_set(x_155, 0, x_225); -return x_155; -} -} -else -{ -lean_object* x_226; lean_object* x_227; uint8_t x_228; -x_226 = lean_ctor_get(x_155, 0); -lean_inc(x_226); -lean_dec(x_155); -x_227 = l_Lean_Expr_constLevels_x21(x_33); -lean_dec_ref(x_33); -x_228 = lean_unbox(x_226); -lean_dec(x_226); -if (x_228 == 0) -{ -lean_object* x_229; -lean_inc(x_6); -lean_inc_ref(x_5); -lean_inc(x_4); -lean_inc_ref(x_3); -lean_inc_ref(x_1); -x_229 = lean_infer_type(x_1, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_229) == 0) -{ -lean_object* x_230; lean_object* x_231; lean_object* x_232; -x_230 = lean_ctor_get(x_229, 0); -lean_inc(x_230); -lean_dec_ref(x_229); -lean_inc_ref(x_27); -x_231 = l_Lean_mkNot(x_27); -x_232 = l_Lean_Meta_isExprDefEq(x_231, x_230, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_232) == 0) +x_73 = l_Lean_mkConst(x_71, x_72); +x_74 = l_Lean_mkAppN(x_73, x_4); +lean_dec_ref(x_4); +lean_inc(x_9); +lean_inc_ref(x_8); +lean_inc(x_7); +lean_inc_ref(x_6); +lean_inc_ref(x_74); +x_75 = lean_infer_type(x_74, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_75) == 0) { -lean_object* x_233; lean_object* x_234; uint8_t x_235; -x_233 = lean_ctor_get(x_232, 0); -lean_inc(x_233); -if (lean_is_exclusive(x_232)) { - lean_ctor_release(x_232, 0); - x_234 = x_232; +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; uint8_t x_80; +x_76 = lean_ctor_get(x_75, 0); +lean_inc(x_76); +if (lean_is_exclusive(x_75)) { + lean_ctor_release(x_75, 0); + x_77 = x_75; } else { - lean_dec_ref(x_232); - x_234 = lean_box(0); + lean_dec_ref(x_75); + x_77 = lean_box(0); } -x_235 = lean_unbox(x_233); -lean_dec(x_233); -if (x_235 == 0) +x_78 = l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__1; +x_79 = lean_unsigned_to_nat(3u); +x_80 = l_Lean_Expr_isAppOfArity(x_76, x_78, x_79); +if (x_80 == 0) { -lean_object* x_236; lean_object* x_237; lean_object* x_238; -lean_dec(x_227); -lean_dec_ref(x_32); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); -lean_dec_ref(x_1); -x_236 = lean_box(0); -x_237 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_237, 0, x_2); -lean_ctor_set(x_237, 1, x_236); -lean_ctor_set_uint8(x_237, sizeof(void*)*2, x_35); -if (lean_is_scalar(x_234)) { - x_238 = lean_alloc_ctor(0, 1, 0); -} else { - x_238 = x_234; -} -lean_ctor_set(x_238, 0, x_237); -return x_238; +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +lean_dec(x_77); +lean_dec(x_76); +x_81 = l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__3; +x_82 = l_Lean_MessageData_ofExpr(x_74); +lean_ctor_set_tag(x_28, 7); +lean_ctor_set(x_28, 1, x_82); +lean_ctor_set(x_28, 0, x_81); +x_83 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___redArg___lam__3___closed__3; +x_84 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_84, 0, x_28); +lean_ctor_set(x_84, 1, x_83); +x_85 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___redArg(x_84, x_6, x_7, x_8, x_9); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); +return x_85; } else { -lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; -lean_dec_ref(x_2); -x_239 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__6___closed__0; -lean_inc_ref(x_1); -x_240 = lean_array_push(x_239, x_1); -lean_inc_ref(x_24); -x_241 = l_Lean_Expr_beta(x_24, x_240); -x_242 = l_Lean_Tactic_FunInd_rwIfWith___closed__12; -x_243 = l_Lean_mkConst(x_242, x_227); -x_244 = l_Lean_Tactic_FunInd_rwIfWith___closed__8; -x_245 = lean_array_push(x_244, x_27); -x_246 = lean_array_push(x_245, x_26); -x_247 = lean_array_push(x_246, x_1); -x_248 = lean_array_push(x_247, x_32); -x_249 = lean_array_push(x_248, x_25); -x_250 = lean_array_push(x_249, x_24); -x_251 = l_Lean_mkAppN(x_243, x_250); -lean_dec_ref(x_250); -x_252 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_252, 0, x_251); -x_253 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_253, 0, x_241); -lean_ctor_set(x_253, 1, x_252); -lean_ctor_set_uint8(x_253, sizeof(void*)*2, x_35); -if (lean_is_scalar(x_234)) { - x_254 = lean_alloc_ctor(0, 1, 0); +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +lean_free_object(x_28); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); +x_86 = l_Lean_Expr_appArg_x21(x_76); +lean_dec(x_76); +x_87 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_87, 0, x_74); +x_88 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_88, 0, x_86); +lean_ctor_set(x_88, 1, x_87); +lean_ctor_set_uint8(x_88, sizeof(void*)*2, x_34); +if (lean_is_scalar(x_77)) { + x_89 = lean_alloc_ctor(0, 1, 0); } else { - x_254 = x_234; + x_89 = x_77; } -lean_ctor_set(x_254, 0, x_253); -return x_254; +lean_ctor_set(x_89, 0, x_88); +return x_89; } } else { -lean_object* x_255; lean_object* x_256; lean_object* x_257; -lean_dec(x_227); -lean_dec_ref(x_32); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_255 = lean_ctor_get(x_232, 0); -lean_inc(x_255); -if (lean_is_exclusive(x_232)) { - lean_ctor_release(x_232, 0); - x_256 = x_232; +lean_object* x_90; lean_object* x_91; lean_object* x_92; +lean_dec_ref(x_74); +lean_free_object(x_28); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); +x_90 = lean_ctor_get(x_75, 0); +lean_inc(x_90); +if (lean_is_exclusive(x_75)) { + lean_ctor_release(x_75, 0); + x_91 = x_75; } else { - lean_dec_ref(x_232); - x_256 = lean_box(0); + lean_dec_ref(x_75); + x_91 = lean_box(0); } -if (lean_is_scalar(x_256)) { - x_257 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_91)) { + x_92 = lean_alloc_ctor(1, 1, 0); } else { - x_257 = x_256; + x_92 = x_91; +} +lean_ctor_set(x_92, 0, x_90); +return x_92; } -lean_ctor_set(x_257, 0, x_255); -return x_257; } } else { -lean_object* x_258; lean_object* x_259; lean_object* x_260; -lean_dec(x_227); -lean_dec_ref(x_32); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_258 = lean_ctor_get(x_229, 0); -lean_inc(x_258); -if (lean_is_exclusive(x_229)) { - lean_ctor_release(x_229, 0); - x_259 = x_229; -} else { - lean_dec_ref(x_229); - x_259 = lean_box(0); -} -if (lean_is_scalar(x_259)) { - x_260 = lean_alloc_ctor(1, 1, 0); -} else { - x_260 = x_259; -} -lean_ctor_set(x_260, 0, x_258); -return x_260; -} -} -else -{ -lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); +lean_object* x_93; +lean_dec(x_37); +lean_free_object(x_28); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec_ref(x_4); lean_dec_ref(x_3); -lean_dec_ref(x_2); -x_261 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__6___closed__0; -lean_inc_ref(x_1); -x_262 = lean_array_push(x_261, x_1); -lean_inc_ref(x_25); -x_263 = l_Lean_Expr_beta(x_25, x_262); -x_264 = l_Lean_Tactic_FunInd_rwIfWith___closed__14; -x_265 = l_Lean_mkConst(x_264, x_227); -x_266 = l_Lean_Tactic_FunInd_rwIfWith___closed__8; -x_267 = lean_array_push(x_266, x_27); -x_268 = lean_array_push(x_267, x_26); -x_269 = lean_array_push(x_268, x_1); -x_270 = lean_array_push(x_269, x_32); -x_271 = lean_array_push(x_270, x_25); -x_272 = lean_array_push(x_271, x_24); -x_273 = l_Lean_mkAppN(x_265, x_272); -lean_dec_ref(x_272); -x_274 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_274, 0, x_273); -x_275 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_275, 0, x_263); -lean_ctor_set(x_275, 1, x_274); -lean_ctor_set_uint8(x_275, sizeof(void*)*2, x_35); -x_276 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_276, 0, x_275); -return x_276; -} +x_93 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_93, 0, x_1); +lean_ctor_set(x_93, 1, x_23); +lean_ctor_set_uint8(x_93, sizeof(void*)*2, x_34); +lean_ctor_set(x_35, 0, x_93); +return x_35; } } else { -uint8_t x_277; -lean_dec_ref(x_33); -lean_dec_ref(x_32); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_277 = !lean_is_exclusive(x_155); -if (x_277 == 0) -{ -return x_155; -} -else +lean_object* x_94; +x_94 = lean_ctor_get(x_35, 0); +lean_inc(x_94); +lean_dec(x_35); +if (lean_obj_tag(x_94) == 1) { -lean_object* x_278; lean_object* x_279; -x_278 = lean_ctor_get(x_155, 0); -lean_inc(x_278); -lean_dec(x_155); -x_279 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_279, 0, x_278); -return x_279; -} -} +lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; +lean_dec_ref(x_1); +x_95 = lean_ctor_get(x_94, 0); +lean_inc(x_95); +if (lean_is_exclusive(x_94)) { + lean_ctor_release(x_94, 0); + x_96 = x_94; +} else { + lean_dec_ref(x_94); + x_96 = lean_box(0); } -else -{ -uint8_t x_280; -lean_dec_ref(x_33); -lean_dec_ref(x_32); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); +x_97 = l_Lean_Expr_constLevels_x21(x_3); lean_dec_ref(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_280 = !lean_is_exclusive(x_153); -if (x_280 == 0) +x_98 = l_Lean_mkConst(x_95, x_97); +x_99 = l_Lean_mkAppN(x_98, x_4); +lean_dec_ref(x_4); +lean_inc(x_9); +lean_inc_ref(x_8); +lean_inc(x_7); +lean_inc_ref(x_6); +lean_inc_ref(x_99); +x_100 = lean_infer_type(x_99, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_100) == 0) { -return x_153; +lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; uint8_t x_105; +x_101 = lean_ctor_get(x_100, 0); +lean_inc(x_101); +if (lean_is_exclusive(x_100)) { + lean_ctor_release(x_100, 0); + x_102 = x_100; +} else { + lean_dec_ref(x_100); + x_102 = lean_box(0); } -else +x_103 = l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__1; +x_104 = lean_unsigned_to_nat(3u); +x_105 = l_Lean_Expr_isAppOfArity(x_101, x_103, x_104); +if (x_105 == 0) { -lean_object* x_281; lean_object* x_282; -x_281 = lean_ctor_get(x_153, 0); -lean_inc(x_281); -lean_dec(x_153); -x_282 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_282, 0, x_281); -return x_282; -} -} -} -} +lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; +lean_dec(x_102); +lean_dec(x_101); +lean_dec(x_96); +x_106 = l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__3; +x_107 = l_Lean_MessageData_ofExpr(x_99); +lean_ctor_set_tag(x_28, 7); +lean_ctor_set(x_28, 1, x_107); +lean_ctor_set(x_28, 0, x_106); +x_108 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___redArg___lam__3___closed__3; +x_109 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_109, 0, x_28); +lean_ctor_set(x_109, 1, x_108); +x_110 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___redArg(x_109, x_6, x_7, x_8, x_9); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); +return x_110; } else { -lean_object* x_283; -lean_dec(x_10); -lean_inc(x_6); -lean_inc_ref(x_5); -lean_inc(x_4); -lean_inc_ref(x_3); -lean_inc_ref(x_1); -x_283 = lean_infer_type(x_1, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_283) == 0) -{ -lean_object* x_284; lean_object* x_285; lean_object* x_286; -x_284 = lean_ctor_get(x_283, 0); -lean_inc(x_284); -lean_dec_ref(x_283); -x_285 = l_Lean_Tactic_FunInd_rwIfWith___closed__18; -lean_inc(x_6); -lean_inc_ref(x_5); -lean_inc(x_4); -lean_inc_ref(x_3); -lean_inc_ref(x_26); -x_286 = l_Lean_Meta_mkEq(x_26, x_285, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_286) == 0) -{ -lean_object* x_287; lean_object* x_288; -x_287 = lean_ctor_get(x_286, 0); -lean_inc(x_287); -lean_dec_ref(x_286); -lean_inc(x_6); -lean_inc_ref(x_5); -lean_inc(x_4); -lean_inc_ref(x_3); -x_288 = l_Lean_Meta_isExprDefEq(x_284, x_287, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_288) == 0) -{ -uint8_t x_289; -x_289 = !lean_is_exclusive(x_288); -if (x_289 == 0) -{ -lean_object* x_290; lean_object* x_291; uint8_t x_292; -x_290 = lean_ctor_get(x_288, 0); -x_291 = l_Lean_Expr_constLevels_x21(x_28); -lean_dec_ref(x_28); -x_292 = lean_unbox(x_290); -lean_dec(x_290); -if (x_292 == 0) -{ -lean_object* x_293; -lean_free_object(x_288); -lean_inc(x_6); -lean_inc_ref(x_5); -lean_inc(x_4); -lean_inc_ref(x_3); -lean_inc_ref(x_1); -x_293 = lean_infer_type(x_1, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_293) == 0) -{ -lean_object* x_294; lean_object* x_295; lean_object* x_296; -x_294 = lean_ctor_get(x_293, 0); -lean_inc(x_294); -lean_dec_ref(x_293); -x_295 = l_Lean_Tactic_FunInd_rwIfWith___closed__21; -lean_inc(x_6); -lean_inc_ref(x_5); -lean_inc(x_4); -lean_inc_ref(x_3); -lean_inc_ref(x_26); -x_296 = l_Lean_Meta_mkEq(x_26, x_295, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_296) == 0) -{ -lean_object* x_297; lean_object* x_298; -x_297 = lean_ctor_get(x_296, 0); -lean_inc(x_297); -lean_dec_ref(x_296); -x_298 = l_Lean_Meta_isExprDefEq(x_294, x_297, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_298) == 0) -{ -uint8_t x_299; -x_299 = !lean_is_exclusive(x_298); -if (x_299 == 0) -{ -lean_object* x_300; uint8_t x_301; -x_300 = lean_ctor_get(x_298, 0); -x_301 = lean_unbox(x_300); -lean_dec(x_300); -if (x_301 == 0) -{ -lean_object* x_302; lean_object* x_303; -lean_dec(x_291); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); -lean_dec_ref(x_1); -x_302 = lean_box(0); -x_303 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_303, 0, x_2); -lean_ctor_set(x_303, 1, x_302); -lean_ctor_set_uint8(x_303, sizeof(void*)*2, x_30); -lean_ctor_set(x_298, 0, x_303); -return x_298; +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; +lean_free_object(x_28); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); +x_111 = l_Lean_Expr_appArg_x21(x_101); +lean_dec(x_101); +if (lean_is_scalar(x_96)) { + x_112 = lean_alloc_ctor(1, 1, 0); +} else { + x_112 = x_96; } -else -{ -lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; -lean_dec_ref(x_2); -x_304 = l_Lean_Tactic_FunInd_rwIfWith___closed__23; -x_305 = l_Lean_mkConst(x_304, x_291); -x_306 = l_Lean_Tactic_FunInd_rwIfWith___closed__24; -x_307 = lean_array_push(x_306, x_27); -x_308 = lean_array_push(x_307, x_26); -x_309 = lean_array_push(x_308, x_25); -lean_inc_ref(x_24); -x_310 = lean_array_push(x_309, x_24); -x_311 = lean_array_push(x_310, x_1); -x_312 = l_Lean_mkAppN(x_305, x_311); -lean_dec_ref(x_311); -x_313 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_313, 0, x_312); -x_314 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_314, 0, x_24); -lean_ctor_set(x_314, 1, x_313); -lean_ctor_set_uint8(x_314, sizeof(void*)*2, x_30); -lean_ctor_set(x_298, 0, x_314); -return x_298; +lean_ctor_set(x_112, 0, x_99); +x_113 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_113, 0, x_111); +lean_ctor_set(x_113, 1, x_112); +lean_ctor_set_uint8(x_113, sizeof(void*)*2, x_34); +if (lean_is_scalar(x_102)) { + x_114 = lean_alloc_ctor(0, 1, 0); +} else { + x_114 = x_102; } +lean_ctor_set(x_114, 0, x_113); +return x_114; } -else -{ -lean_object* x_315; uint8_t x_316; -x_315 = lean_ctor_get(x_298, 0); -lean_inc(x_315); -lean_dec(x_298); -x_316 = lean_unbox(x_315); -lean_dec(x_315); -if (x_316 == 0) -{ -lean_object* x_317; lean_object* x_318; lean_object* x_319; -lean_dec(x_291); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); -lean_dec_ref(x_1); -x_317 = lean_box(0); -x_318 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_318, 0, x_2); -lean_ctor_set(x_318, 1, x_317); -lean_ctor_set_uint8(x_318, sizeof(void*)*2, x_30); -x_319 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_319, 0, x_318); -return x_319; } else { -lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; -lean_dec_ref(x_2); -x_320 = l_Lean_Tactic_FunInd_rwIfWith___closed__23; -x_321 = l_Lean_mkConst(x_320, x_291); -x_322 = l_Lean_Tactic_FunInd_rwIfWith___closed__24; -x_323 = lean_array_push(x_322, x_27); -x_324 = lean_array_push(x_323, x_26); -x_325 = lean_array_push(x_324, x_25); -lean_inc_ref(x_24); -x_326 = lean_array_push(x_325, x_24); -x_327 = lean_array_push(x_326, x_1); -x_328 = l_Lean_mkAppN(x_321, x_327); -lean_dec_ref(x_327); -x_329 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_329, 0, x_328); -x_330 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_330, 0, x_24); -lean_ctor_set(x_330, 1, x_329); -lean_ctor_set_uint8(x_330, sizeof(void*)*2, x_30); -x_331 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_331, 0, x_330); -return x_331; +lean_object* x_115; lean_object* x_116; lean_object* x_117; +lean_dec_ref(x_99); +lean_dec(x_96); +lean_free_object(x_28); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); +x_115 = lean_ctor_get(x_100, 0); +lean_inc(x_115); +if (lean_is_exclusive(x_100)) { + lean_ctor_release(x_100, 0); + x_116 = x_100; +} else { + lean_dec_ref(x_100); + x_116 = lean_box(0); } +if (lean_is_scalar(x_116)) { + x_117 = lean_alloc_ctor(1, 1, 0); +} else { + x_117 = x_116; } +lean_ctor_set(x_117, 0, x_115); +return x_117; } -else -{ -uint8_t x_332; -lean_dec(x_291); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_332 = !lean_is_exclusive(x_298); -if (x_332 == 0) -{ -return x_298; } else { -lean_object* x_333; lean_object* x_334; -x_333 = lean_ctor_get(x_298, 0); -lean_inc(x_333); -lean_dec(x_298); -x_334 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_334, 0, x_333); -return x_334; +lean_object* x_118; lean_object* x_119; +lean_dec(x_94); +lean_free_object(x_28); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec_ref(x_4); +lean_dec_ref(x_3); +x_118 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_118, 0, x_1); +lean_ctor_set(x_118, 1, x_23); +lean_ctor_set_uint8(x_118, sizeof(void*)*2, x_34); +x_119 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_119, 0, x_118); +return x_119; } } } else { -uint8_t x_335; -lean_dec(x_294); -lean_dec(x_291); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); +uint8_t x_120; +lean_free_object(x_28); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec_ref(x_4); lean_dec_ref(x_3); -lean_dec_ref(x_2); lean_dec_ref(x_1); -x_335 = !lean_is_exclusive(x_296); -if (x_335 == 0) +x_120 = !lean_is_exclusive(x_35); +if (x_120 == 0) { -return x_296; +return x_35; } else { -lean_object* x_336; lean_object* x_337; -x_336 = lean_ctor_get(x_296, 0); -lean_inc(x_336); -lean_dec(x_296); -x_337 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_337, 0, x_336); -return x_337; +lean_object* x_121; lean_object* x_122; +x_121 = lean_ctor_get(x_35, 0); +lean_inc(x_121); +lean_dec(x_35); +x_122 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_122, 0, x_121); +return x_122; } } } else { -uint8_t x_338; -lean_dec(x_291); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); +lean_dec(x_32); +lean_free_object(x_28); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec_ref(x_4); lean_dec_ref(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_338 = !lean_is_exclusive(x_293); -if (x_338 == 0) -{ -return x_293; -} -else -{ -lean_object* x_339; lean_object* x_340; -x_339 = lean_ctor_get(x_293, 0); -lean_inc(x_339); -lean_dec(x_293); -x_340 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_340, 0, x_339); -return x_340; +x_11 = lean_box(0); +goto block_16; } } } else { -lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); +lean_object* x_123; +x_123 = lean_ctor_get(x_28, 0); +lean_inc(x_123); +lean_dec(x_28); +if (lean_obj_tag(x_123) == 0) +{ +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec_ref(x_4); lean_dec_ref(x_3); -lean_dec_ref(x_2); -x_341 = l_Lean_Tactic_FunInd_rwIfWith___closed__26; -x_342 = l_Lean_mkConst(x_341, x_291); -x_343 = l_Lean_Tactic_FunInd_rwIfWith___closed__24; -x_344 = lean_array_push(x_343, x_27); -x_345 = lean_array_push(x_344, x_26); -lean_inc_ref(x_25); -x_346 = lean_array_push(x_345, x_25); -x_347 = lean_array_push(x_346, x_24); -x_348 = lean_array_push(x_347, x_1); -x_349 = l_Lean_mkAppN(x_342, x_348); -lean_dec_ref(x_348); -x_350 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_350, 0, x_349); -x_351 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_351, 0, x_25); -lean_ctor_set(x_351, 1, x_350); -lean_ctor_set_uint8(x_351, sizeof(void*)*2, x_30); -lean_ctor_set(x_288, 0, x_351); -return x_288; -} +x_11 = lean_box(0); +goto block_16; } else { -lean_object* x_352; lean_object* x_353; uint8_t x_354; -x_352 = lean_ctor_get(x_288, 0); -lean_inc(x_352); -lean_dec(x_288); -x_353 = l_Lean_Expr_constLevels_x21(x_28); -lean_dec_ref(x_28); -x_354 = lean_unbox(x_352); -lean_dec(x_352); -if (x_354 == 0) -{ -lean_object* x_355; -lean_inc(x_6); -lean_inc_ref(x_5); -lean_inc(x_4); -lean_inc_ref(x_3); -lean_inc_ref(x_1); -x_355 = lean_infer_type(x_1, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_355) == 0) -{ -lean_object* x_356; lean_object* x_357; lean_object* x_358; -x_356 = lean_ctor_get(x_355, 0); -lean_inc(x_356); -lean_dec_ref(x_355); -x_357 = l_Lean_Tactic_FunInd_rwIfWith___closed__21; -lean_inc(x_6); -lean_inc_ref(x_5); -lean_inc(x_4); -lean_inc_ref(x_3); -lean_inc_ref(x_26); -x_358 = l_Lean_Meta_mkEq(x_26, x_357, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_358) == 0) +lean_object* x_124; +x_124 = lean_ctor_get(x_123, 0); +lean_inc(x_124); +lean_dec_ref(x_123); +if (lean_obj_tag(x_124) == 1) { -lean_object* x_359; lean_object* x_360; -x_359 = lean_ctor_get(x_358, 0); -lean_inc(x_359); -lean_dec_ref(x_358); -x_360 = l_Lean_Meta_isExprDefEq(x_356, x_359, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_360) == 0) +lean_object* x_125; uint8_t x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_124, 0); +lean_inc(x_125); +lean_dec_ref(x_124); +x_126 = 1; +lean_inc(x_9); +lean_inc_ref(x_8); +lean_inc(x_7); +lean_inc_ref(x_6); +x_127 = l_Lean_Meta_getUnfoldEqnFor_x3f(x_125, x_126, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_127) == 0) { -lean_object* x_361; lean_object* x_362; uint8_t x_363; -x_361 = lean_ctor_get(x_360, 0); -lean_inc(x_361); -if (lean_is_exclusive(x_360)) { - lean_ctor_release(x_360, 0); - x_362 = x_360; +lean_object* x_128; lean_object* x_129; +x_128 = lean_ctor_get(x_127, 0); +lean_inc(x_128); +if (lean_is_exclusive(x_127)) { + lean_ctor_release(x_127, 0); + x_129 = x_127; } else { - lean_dec_ref(x_360); - x_362 = lean_box(0); + lean_dec_ref(x_127); + x_129 = lean_box(0); } -x_363 = lean_unbox(x_361); -lean_dec(x_361); -if (x_363 == 0) +if (lean_obj_tag(x_128) == 1) { -lean_object* x_364; lean_object* x_365; lean_object* x_366; -lean_dec(x_353); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); +lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; +lean_dec(x_129); lean_dec_ref(x_1); -x_364 = lean_box(0); -x_365 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_365, 0, x_2); -lean_ctor_set(x_365, 1, x_364); -lean_ctor_set_uint8(x_365, sizeof(void*)*2, x_30); -if (lean_is_scalar(x_362)) { - x_366 = lean_alloc_ctor(0, 1, 0); +x_130 = lean_ctor_get(x_128, 0); +lean_inc(x_130); +if (lean_is_exclusive(x_128)) { + lean_ctor_release(x_128, 0); + x_131 = x_128; +} else { + lean_dec_ref(x_128); + x_131 = lean_box(0); +} +x_132 = l_Lean_Expr_constLevels_x21(x_3); +lean_dec_ref(x_3); +x_133 = l_Lean_mkConst(x_130, x_132); +x_134 = l_Lean_mkAppN(x_133, x_4); +lean_dec_ref(x_4); +lean_inc(x_9); +lean_inc_ref(x_8); +lean_inc(x_7); +lean_inc_ref(x_6); +lean_inc_ref(x_134); +x_135 = lean_infer_type(x_134, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_135) == 0) +{ +lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; uint8_t x_140; +x_136 = lean_ctor_get(x_135, 0); +lean_inc(x_136); +if (lean_is_exclusive(x_135)) { + lean_ctor_release(x_135, 0); + x_137 = x_135; } else { - x_366 = x_362; + lean_dec_ref(x_135); + x_137 = lean_box(0); } -lean_ctor_set(x_366, 0, x_365); -return x_366; +x_138 = l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__1; +x_139 = lean_unsigned_to_nat(3u); +x_140 = l_Lean_Expr_isAppOfArity(x_136, x_138, x_139); +if (x_140 == 0) +{ +lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; +lean_dec(x_137); +lean_dec(x_136); +lean_dec(x_131); +x_141 = l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__3; +x_142 = l_Lean_MessageData_ofExpr(x_134); +x_143 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_143, 0, x_141); +lean_ctor_set(x_143, 1, x_142); +x_144 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___redArg___lam__3___closed__3; +x_145 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_145, 0, x_143); +lean_ctor_set(x_145, 1, x_144); +x_146 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___redArg(x_145, x_6, x_7, x_8, x_9); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); +return x_146; } else { -lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; -lean_dec_ref(x_2); -x_367 = l_Lean_Tactic_FunInd_rwIfWith___closed__23; -x_368 = l_Lean_mkConst(x_367, x_353); -x_369 = l_Lean_Tactic_FunInd_rwIfWith___closed__24; -x_370 = lean_array_push(x_369, x_27); -x_371 = lean_array_push(x_370, x_26); -x_372 = lean_array_push(x_371, x_25); -lean_inc_ref(x_24); -x_373 = lean_array_push(x_372, x_24); -x_374 = lean_array_push(x_373, x_1); -x_375 = l_Lean_mkAppN(x_368, x_374); -lean_dec_ref(x_374); -x_376 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_376, 0, x_375); -x_377 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_377, 0, x_24); -lean_ctor_set(x_377, 1, x_376); -lean_ctor_set_uint8(x_377, sizeof(void*)*2, x_30); -if (lean_is_scalar(x_362)) { - x_378 = lean_alloc_ctor(0, 1, 0); +lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); +x_147 = l_Lean_Expr_appArg_x21(x_136); +lean_dec(x_136); +if (lean_is_scalar(x_131)) { + x_148 = lean_alloc_ctor(1, 1, 0); } else { - x_378 = x_362; + x_148 = x_131; } -lean_ctor_set(x_378, 0, x_377); -return x_378; +lean_ctor_set(x_148, 0, x_134); +x_149 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_149, 0, x_147); +lean_ctor_set(x_149, 1, x_148); +lean_ctor_set_uint8(x_149, sizeof(void*)*2, x_126); +if (lean_is_scalar(x_137)) { + x_150 = lean_alloc_ctor(0, 1, 0); +} else { + x_150 = x_137; +} +lean_ctor_set(x_150, 0, x_149); +return x_150; } } else { -lean_object* x_379; lean_object* x_380; lean_object* x_381; -lean_dec(x_353); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_379 = lean_ctor_get(x_360, 0); -lean_inc(x_379); -if (lean_is_exclusive(x_360)) { - lean_ctor_release(x_360, 0); - x_380 = x_360; +lean_object* x_151; lean_object* x_152; lean_object* x_153; +lean_dec_ref(x_134); +lean_dec(x_131); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); +x_151 = lean_ctor_get(x_135, 0); +lean_inc(x_151); +if (lean_is_exclusive(x_135)) { + lean_ctor_release(x_135, 0); + x_152 = x_135; } else { - lean_dec_ref(x_360); - x_380 = lean_box(0); + lean_dec_ref(x_135); + x_152 = lean_box(0); } -if (lean_is_scalar(x_380)) { - x_381 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_152)) { + x_153 = lean_alloc_ctor(1, 1, 0); } else { - x_381 = x_380; + x_153 = x_152; } -lean_ctor_set(x_381, 0, x_379); -return x_381; +lean_ctor_set(x_153, 0, x_151); +return x_153; } } else { -lean_object* x_382; lean_object* x_383; lean_object* x_384; -lean_dec(x_356); -lean_dec(x_353); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); +lean_object* x_154; lean_object* x_155; +lean_dec(x_128); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec_ref(x_4); lean_dec_ref(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_382 = lean_ctor_get(x_358, 0); -lean_inc(x_382); -if (lean_is_exclusive(x_358)) { - lean_ctor_release(x_358, 0); - x_383 = x_358; -} else { - lean_dec_ref(x_358); - x_383 = lean_box(0); -} -if (lean_is_scalar(x_383)) { - x_384 = lean_alloc_ctor(1, 1, 0); +x_154 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_154, 0, x_1); +lean_ctor_set(x_154, 1, x_23); +lean_ctor_set_uint8(x_154, sizeof(void*)*2, x_126); +if (lean_is_scalar(x_129)) { + x_155 = lean_alloc_ctor(0, 1, 0); } else { - x_384 = x_383; + x_155 = x_129; } -lean_ctor_set(x_384, 0, x_382); -return x_384; +lean_ctor_set(x_155, 0, x_154); +return x_155; } } else { -lean_object* x_385; lean_object* x_386; lean_object* x_387; -lean_dec(x_353); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); +lean_object* x_156; lean_object* x_157; lean_object* x_158; +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec_ref(x_4); lean_dec_ref(x_3); -lean_dec_ref(x_2); lean_dec_ref(x_1); -x_385 = lean_ctor_get(x_355, 0); -lean_inc(x_385); -if (lean_is_exclusive(x_355)) { - lean_ctor_release(x_355, 0); - x_386 = x_355; +x_156 = lean_ctor_get(x_127, 0); +lean_inc(x_156); +if (lean_is_exclusive(x_127)) { + lean_ctor_release(x_127, 0); + x_157 = x_127; } else { - lean_dec_ref(x_355); - x_386 = lean_box(0); + lean_dec_ref(x_127); + x_157 = lean_box(0); } -if (lean_is_scalar(x_386)) { - x_387 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_157)) { + x_158 = lean_alloc_ctor(1, 1, 0); } else { - x_387 = x_386; + x_158 = x_157; } -lean_ctor_set(x_387, 0, x_385); -return x_387; +lean_ctor_set(x_158, 0, x_156); +return x_158; } } else { -lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); +lean_dec(x_124); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); +lean_dec_ref(x_4); lean_dec_ref(x_3); -lean_dec_ref(x_2); -x_388 = l_Lean_Tactic_FunInd_rwIfWith___closed__26; -x_389 = l_Lean_mkConst(x_388, x_353); -x_390 = l_Lean_Tactic_FunInd_rwIfWith___closed__24; -x_391 = lean_array_push(x_390, x_27); -x_392 = lean_array_push(x_391, x_26); -lean_inc_ref(x_25); -x_393 = lean_array_push(x_392, x_25); -x_394 = lean_array_push(x_393, x_24); -x_395 = lean_array_push(x_394, x_1); -x_396 = l_Lean_mkAppN(x_389, x_395); -lean_dec_ref(x_395); -x_397 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_397, 0, x_396); -x_398 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_398, 0, x_25); -lean_ctor_set(x_398, 1, x_397); -lean_ctor_set_uint8(x_398, sizeof(void*)*2, x_30); -x_399 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_399, 0, x_398); -return x_399; +x_11 = lean_box(0); +goto block_16; } } } -else -{ -uint8_t x_400; -lean_dec_ref(x_28); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_400 = !lean_is_exclusive(x_288); -if (x_400 == 0) -{ -return x_288; } -else +block_16: { -lean_object* x_401; lean_object* x_402; -x_401 = lean_ctor_get(x_288, 0); -lean_inc(x_401); -lean_dec(x_288); -x_402 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_402, 0, x_401); -return x_402; +lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_box(0); +x_13 = 1; +x_14 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_14, 0, x_1); +lean_ctor_set(x_14, 1, x_12); +lean_ctor_set_uint8(x_14, sizeof(void*)*2, x_13); +x_15 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_15, 0, x_14); +return x_15; } } } -else -{ -uint8_t x_403; -lean_dec(x_284); -lean_dec_ref(x_28); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_403 = !lean_is_exclusive(x_286); -if (x_403 == 0) +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -return x_286; +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_8 = l___private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0___lam__1___closed__0; +x_9 = l_Lean_Expr_getAppNumArgs(x_2); +lean_inc(x_9); +x_10 = lean_mk_array(x_9, x_8); +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_sub(x_9, x_11); +lean_dec(x_9); +lean_inc_ref(x_2); +x_13 = l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1(x_2, x_1, x_2, x_10, x_12, x_3, x_4, x_5, x_6); +return x_13; } -else -{ -lean_object* x_404; lean_object* x_405; -x_404 = lean_ctor_get(x_286, 0); -lean_inc(x_404); -lean_dec(x_286); -x_405 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_405, 0, x_404); -return x_405; } +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec_ref(x_1); +return x_8; } } -else +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__0___boxed(lean_object* x_1, lean_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_406; -lean_dec_ref(x_28); -lean_dec_ref(x_27); -lean_dec_ref(x_26); -lean_dec_ref(x_25); -lean_dec_ref(x_24); +size_t x_8; size_t x_9; lean_object* x_10; +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = lean_unbox_usize(x_6); lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); +x_10 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__0(x_1, x_2, x_3, x_4, x_8, x_9, x_7); +lean_dec_ref(x_7); +lean_dec_ref(x_4); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_406 = !lean_is_exclusive(x_283); -if (x_406 == 0) +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -return x_283; +lean_object* x_11; +x_11 = l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec_ref(x_2); +return x_11; } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__14___redArg___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: { -lean_object* x_407; lean_object* x_408; -x_407 = lean_ctor_get(x_283, 0); -lean_inc(x_407); -lean_dec(x_283); -x_408 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_408, 0, x_407); -return x_408; +lean_object* x_9; +x_9 = lean_apply_7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, lean_box(0)); +return x_9; } } +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__14___redArg___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__14___redArg___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +return x_9; } } +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__14___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__14___redArg___lam__0___boxed), 8, 3); +lean_closure_set(x_11, 0, x_3); +lean_closure_set(x_11, 1, x_4); +lean_closure_set(x_11, 2, x_5); +x_12 = l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_withUserNamesImpl(lean_box(0), x_1, x_2, x_11, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_12) == 0) +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_ctor_get(x_12, 0); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) +{ +return x_12; } +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_14, 0); +x_17 = lean_ctor_get(x_14, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_14); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +lean_ctor_set(x_12, 0, x_18); +return x_12; } } -block_15: +else { -lean_object* x_11; uint8_t x_12; lean_object* x_13; lean_object* x_14; -x_11 = lean_box(0); -x_12 = 1; -x_13 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_13, 0, x_2); -lean_ctor_set(x_13, 1, x_11); -lean_ctor_set_uint8(x_13, sizeof(void*)*2, x_12); -if (lean_is_scalar(x_10)) { - x_14 = lean_alloc_ctor(0, 1, 0); +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_19 = lean_ctor_get(x_12, 0); +lean_inc(x_19); +lean_dec(x_12); +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +if (lean_is_exclusive(x_19)) { + lean_ctor_release(x_19, 0); + lean_ctor_release(x_19, 1); + x_22 = x_19; } else { - x_14 = x_10; + lean_dec_ref(x_19); + x_22 = lean_box(0); } -lean_ctor_set(x_14, 0, x_13); -return x_14; +if (lean_is_scalar(x_22)) { + x_23 = lean_alloc_ctor(0, 2, 0); +} else { + x_23 = x_22; +} +lean_ctor_set(x_23, 0, x_20); +lean_ctor_set(x_23, 1, x_21); +x_24 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_24, 0, x_23); +return x_24; } } else { -uint8_t x_409; -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_409 = !lean_is_exclusive(x_8); -if (x_409 == 0) +uint8_t x_25; +x_25 = !lean_is_exclusive(x_12); +if (x_25 == 0) { -return x_8; +return x_12; } else { -lean_object* x_410; lean_object* x_411; -x_410 = lean_ctor_get(x_8, 0); -lean_inc(x_410); -lean_dec(x_8); -x_411 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_411, 0, x_410); -return x_411; +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_12, 0); +lean_inc(x_26); +lean_dec(x_12); +x_27 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_27, 0, x_26); +return x_27; } } } } -LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_rwIfWith___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__14(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_8; -x_8 = l_Lean_Tactic_FunInd_rwIfWith(x_1, x_2, x_3, x_4, x_5, x_6); -return x_8; +lean_object* x_12; +x_12 = l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__14___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_12; } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__0___redArg(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__8___redArg___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_4; lean_object* x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; -x_4 = lean_ctor_get(x_2, 2); -x_5 = lean_ctor_get(x_2, 13); -x_6 = l_Lean_checkTraceOption(x_5, x_4, x_1); -x_7 = lean_box(x_6); -x_8 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_8, 0, x_7); -return x_8; +lean_object* x_11; +x_11 = lean_apply_9(x_1, x_4, x_5, x_2, x_3, x_6, x_7, x_8, x_9, lean_box(0)); +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__8___redArg___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_7; -x_7 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__0___redArg(x_1, x_4); -return x_7; +lean_object* x_11; +x_11 = l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__8___redArg___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__8___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_8; lean_object* x_9; uint8_t x_10; -x_8 = lean_ctor_get(x_5, 5); -x_9 = l_Lean_addMessageContextFull___at___00Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__2_spec__7(x_2, x_3, x_4, x_5, x_6); -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = lean_ctor_get(x_9, 0); -x_12 = lean_st_ref_take(x_6); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) +lean_object* x_13; lean_object* x_14; +x_13 = lean_alloc_closure((void*)(l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__8___redArg___lam__0___boxed), 10, 3); +lean_closure_set(x_13, 0, x_3); +lean_closure_set(x_13, 1, x_6); +lean_closure_set(x_13, 2, x_7); +x_14 = l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAux(lean_box(0), x_1, x_2, x_13, x_4, x_5, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_14; uint8_t x_15; -x_14 = lean_ctor_get(x_12, 4); +uint8_t x_15; x_15 = !lean_is_exclusive(x_14); if (x_15 == 0) { -lean_object* x_16; double x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_object* x_16; uint8_t x_17; x_16 = lean_ctor_get(x_14, 0); -x_17 = l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13___redArg___closed__0; -x_18 = 0; -x_19 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__21___closed__3; -x_20 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_20, 0, x_1); -lean_ctor_set(x_20, 1, x_19); -lean_ctor_set_float(x_20, sizeof(void*)*2, x_17); -lean_ctor_set_float(x_20, sizeof(void*)*2 + 8, x_17); -lean_ctor_set_uint8(x_20, sizeof(void*)*2 + 16, x_18); -x_21 = l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__2___closed__0; -x_22 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_11); -lean_ctor_set(x_22, 2, x_21); -lean_inc(x_8); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_8); -lean_ctor_set(x_23, 1, x_22); -x_24 = l_Lean_PersistentArray_push___redArg(x_16, x_23); -lean_ctor_set(x_14, 0, x_24); -x_25 = lean_st_ref_set(x_6, x_12); -x_26 = lean_box(0); -lean_ctor_set(x_9, 0, x_26); -return x_9; +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +return x_14; } else { -uint64_t x_27; lean_object* x_28; double x_29; uint8_t x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_27 = lean_ctor_get_uint64(x_14, sizeof(void*)*1); -x_28 = lean_ctor_get(x_14, 0); -lean_inc(x_28); -lean_dec(x_14); -x_29 = l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13___redArg___closed__0; -x_30 = 0; -x_31 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__21___closed__3; -x_32 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_32, 0, x_1); -lean_ctor_set(x_32, 1, x_31); -lean_ctor_set_float(x_32, sizeof(void*)*2, x_29); -lean_ctor_set_float(x_32, sizeof(void*)*2 + 8, x_29); -lean_ctor_set_uint8(x_32, sizeof(void*)*2 + 16, x_30); -x_33 = l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__2___closed__0; -x_34 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_34, 0, x_32); -lean_ctor_set(x_34, 1, x_11); -lean_ctor_set(x_34, 2, x_33); -lean_inc(x_8); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_8); -lean_ctor_set(x_35, 1, x_34); -x_36 = l_Lean_PersistentArray_push___redArg(x_28, x_35); -x_37 = lean_alloc_ctor(0, 1, 8); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set_uint64(x_37, sizeof(void*)*1, x_27); -lean_ctor_set(x_12, 4, x_37); -x_38 = lean_st_ref_set(x_6, x_12); -x_39 = lean_box(0); -lean_ctor_set(x_9, 0, x_39); -return x_9; +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_16, 0); +x_19 = lean_ctor_get(x_16, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_16); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +lean_ctor_set(x_14, 0, x_20); +return x_14; } } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint64_t x_49; lean_object* x_50; lean_object* x_51; double x_52; uint8_t x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_40 = lean_ctor_get(x_12, 4); -x_41 = lean_ctor_get(x_12, 0); -x_42 = lean_ctor_get(x_12, 1); -x_43 = lean_ctor_get(x_12, 2); -x_44 = lean_ctor_get(x_12, 3); -x_45 = lean_ctor_get(x_12, 5); -x_46 = lean_ctor_get(x_12, 6); -x_47 = lean_ctor_get(x_12, 7); -x_48 = lean_ctor_get(x_12, 8); -lean_inc(x_48); -lean_inc(x_47); -lean_inc(x_46); -lean_inc(x_45); -lean_inc(x_40); -lean_inc(x_44); -lean_inc(x_43); -lean_inc(x_42); -lean_inc(x_41); -lean_dec(x_12); -x_49 = lean_ctor_get_uint64(x_40, sizeof(void*)*1); -x_50 = lean_ctor_get(x_40, 0); -lean_inc_ref(x_50); -if (lean_is_exclusive(x_40)) { - lean_ctor_release(x_40, 0); - x_51 = x_40; +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_21 = lean_ctor_get(x_14, 0); +lean_inc(x_21); +lean_dec(x_14); +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +if (lean_is_exclusive(x_21)) { + lean_ctor_release(x_21, 0); + lean_ctor_release(x_21, 1); + x_24 = x_21; } else { - lean_dec_ref(x_40); - x_51 = lean_box(0); + lean_dec_ref(x_21); + x_24 = lean_box(0); } -x_52 = l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13___redArg___closed__0; -x_53 = 0; -x_54 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__21___closed__3; -x_55 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_55, 0, x_1); -lean_ctor_set(x_55, 1, x_54); -lean_ctor_set_float(x_55, sizeof(void*)*2, x_52); -lean_ctor_set_float(x_55, sizeof(void*)*2 + 8, x_52); -lean_ctor_set_uint8(x_55, sizeof(void*)*2 + 16, x_53); -x_56 = l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__2___closed__0; -x_57 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_57, 0, x_55); -lean_ctor_set(x_57, 1, x_11); -lean_ctor_set(x_57, 2, x_56); -lean_inc(x_8); -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_8); -lean_ctor_set(x_58, 1, x_57); -x_59 = l_Lean_PersistentArray_push___redArg(x_50, x_58); -if (lean_is_scalar(x_51)) { - x_60 = lean_alloc_ctor(0, 1, 8); +if (lean_is_scalar(x_24)) { + x_25 = lean_alloc_ctor(0, 2, 0); } else { - x_60 = x_51; + x_25 = x_24; } -lean_ctor_set(x_60, 0, x_59); -lean_ctor_set_uint64(x_60, sizeof(void*)*1, x_49); -x_61 = lean_alloc_ctor(0, 9, 0); -lean_ctor_set(x_61, 0, x_41); -lean_ctor_set(x_61, 1, x_42); -lean_ctor_set(x_61, 2, x_43); -lean_ctor_set(x_61, 3, x_44); -lean_ctor_set(x_61, 4, x_60); -lean_ctor_set(x_61, 5, x_45); -lean_ctor_set(x_61, 6, x_46); -lean_ctor_set(x_61, 7, x_47); -lean_ctor_set(x_61, 8, x_48); -x_62 = lean_st_ref_set(x_6, x_61); -x_63 = lean_box(0); -lean_ctor_set(x_9, 0, x_63); -return x_9; +lean_ctor_set(x_25, 0, x_22); +lean_ctor_set(x_25, 1, x_23); +x_26 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_26, 0, x_25); +return x_26; } } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; uint64_t x_76; lean_object* x_77; lean_object* x_78; double x_79; uint8_t x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; -x_64 = lean_ctor_get(x_9, 0); -lean_inc(x_64); -lean_dec(x_9); -x_65 = lean_st_ref_take(x_6); -x_66 = lean_ctor_get(x_65, 4); -lean_inc_ref(x_66); -x_67 = lean_ctor_get(x_65, 0); -lean_inc_ref(x_67); -x_68 = lean_ctor_get(x_65, 1); -lean_inc(x_68); -x_69 = lean_ctor_get(x_65, 2); -lean_inc_ref(x_69); -x_70 = lean_ctor_get(x_65, 3); -lean_inc_ref(x_70); -x_71 = lean_ctor_get(x_65, 5); -lean_inc_ref(x_71); -x_72 = lean_ctor_get(x_65, 6); -lean_inc_ref(x_72); -x_73 = lean_ctor_get(x_65, 7); -lean_inc_ref(x_73); -x_74 = lean_ctor_get(x_65, 8); -lean_inc_ref(x_74); -if (lean_is_exclusive(x_65)) { - lean_ctor_release(x_65, 0); - lean_ctor_release(x_65, 1); - lean_ctor_release(x_65, 2); - lean_ctor_release(x_65, 3); - lean_ctor_release(x_65, 4); - lean_ctor_release(x_65, 5); - lean_ctor_release(x_65, 6); - lean_ctor_release(x_65, 7); - lean_ctor_release(x_65, 8); - x_75 = x_65; -} else { - lean_dec_ref(x_65); - x_75 = lean_box(0); -} -x_76 = lean_ctor_get_uint64(x_66, sizeof(void*)*1); -x_77 = lean_ctor_get(x_66, 0); -lean_inc_ref(x_77); -if (lean_is_exclusive(x_66)) { - lean_ctor_release(x_66, 0); - x_78 = x_66; -} else { - lean_dec_ref(x_66); - x_78 = lean_box(0); -} -x_79 = l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13___redArg___closed__0; -x_80 = 0; -x_81 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__21___closed__3; -x_82 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_82, 0, x_1); -lean_ctor_set(x_82, 1, x_81); -lean_ctor_set_float(x_82, sizeof(void*)*2, x_79); -lean_ctor_set_float(x_82, sizeof(void*)*2 + 8, x_79); -lean_ctor_set_uint8(x_82, sizeof(void*)*2 + 16, x_80); -x_83 = l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__2___closed__0; -x_84 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_84, 0, x_82); -lean_ctor_set(x_84, 1, x_64); -lean_ctor_set(x_84, 2, x_83); -lean_inc(x_8); -x_85 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_85, 0, x_8); -lean_ctor_set(x_85, 1, x_84); -x_86 = l_Lean_PersistentArray_push___redArg(x_77, x_85); -if (lean_is_scalar(x_78)) { - x_87 = lean_alloc_ctor(0, 1, 8); -} else { - x_87 = x_78; -} -lean_ctor_set(x_87, 0, x_86); -lean_ctor_set_uint64(x_87, sizeof(void*)*1, x_76); -if (lean_is_scalar(x_75)) { - x_88 = lean_alloc_ctor(0, 9, 0); -} else { - x_88 = x_75; -} -lean_ctor_set(x_88, 0, x_67); -lean_ctor_set(x_88, 1, x_68); -lean_ctor_set(x_88, 2, x_69); -lean_ctor_set(x_88, 3, x_70); -lean_ctor_set(x_88, 4, x_87); -lean_ctor_set(x_88, 5, x_71); -lean_ctor_set(x_88, 6, x_72); -lean_ctor_set(x_88, 7, x_73); -lean_ctor_set(x_88, 8, x_74); -x_89 = lean_st_ref_set(x_6, x_88); -x_90 = lean_box(0); -x_91 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_91, 0, x_90); -return x_91; -} -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___closed__0() { -_start: +uint8_t x_27; +x_27 = !lean_is_exclusive(x_14); +if (x_27 == 0) { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("rwLetWith failed:", 17, 17); -return x_1; -} +return x_14; } -static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___closed__1() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___closed__0; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_14, 0); +lean_inc(x_28); +lean_dec(x_14); +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, x_28); +return x_29; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("not a let expression or `", 25, 25); -return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___closed__3() { +LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___closed__2; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_14; +x_14 = l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__8___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_14; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___closed__4() { +LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__4_spec__6___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("` is not definitionally equal to `", 34, 34); -return x_1; +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_6 = lean_st_ref_get(x_4); +x_7 = lean_ctor_get(x_6, 0); +lean_inc_ref(x_7); +lean_dec(x_6); +x_8 = l_Lean_Meta_Match_Extension_getMatcherInfo_x3f(x_7, x_1); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_2); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_3); +x_11 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_11, 0, x_10); +return x_11; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___closed__5() { +LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__4_spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___closed__4; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_9; +x_9 = l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__4_spec__6___redArg(x_1, x_2, x_3, x_7); +return x_9; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Meta_lambdaBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__9_spec__25___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_8; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_43; -x_43 = l_Lean_Expr_isLet(x_2); -if (x_43 == 0) -{ -x_14 = x_3; -x_15 = x_4; -x_16 = x_5; -x_17 = x_6; -x_18 = lean_box(0); -goto block_42; -} -else -{ -lean_object* x_44; lean_object* x_45; -x_44 = l_Lean_Expr_letValue_x21(x_2); -lean_inc(x_6); -lean_inc_ref(x_5); -lean_inc(x_4); -lean_inc_ref(x_3); -lean_inc_ref(x_1); -x_45 = l_Lean_Meta_isExprDefEq(x_44, x_1, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_45) == 0) +lean_object* x_12; uint8_t x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; +x_12 = lean_alloc_closure((void*)(l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__8___redArg___lam__0___boxed), 10, 3); +lean_closure_set(x_12, 0, x_3); +lean_closure_set(x_12, 1, x_5); +lean_closure_set(x_12, 2, x_6); +x_13 = 1; +x_14 = 0; +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_2); +x_16 = l___private_Lean_Meta_Basic_0__Lean_Meta_lambdaTelescopeImp(lean_box(0), x_1, x_13, x_14, x_13, x_14, x_15, x_12, x_4, x_7, x_8, x_9, x_10); +lean_dec_ref(x_15); +if (lean_obj_tag(x_16) == 0) { -uint8_t x_46; -x_46 = !lean_is_exclusive(x_45); -if (x_46 == 0) +uint8_t x_17; +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) { -lean_object* x_47; uint8_t x_48; -x_47 = lean_ctor_get(x_45, 0); -x_48 = lean_unbox(x_47); -if (x_48 == 0) +lean_object* x_18; uint8_t x_19; +x_18 = lean_ctor_get(x_16, 0); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) { -lean_free_object(x_45); -lean_dec(x_47); -x_14 = x_3; -x_15 = x_4; -x_16 = x_5; -x_17 = x_6; -x_18 = lean_box(0); -goto block_42; +return x_16; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -x_49 = l_Lean_Expr_letBody_x21(x_2); -lean_dec_ref(x_2); -x_50 = lean_expr_instantiate1(x_49, x_1); -lean_dec_ref(x_1); -lean_dec_ref(x_49); -x_51 = lean_box(0); -x_52 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_52, 0, x_50); -lean_ctor_set(x_52, 1, x_51); -x_53 = lean_unbox(x_47); -lean_dec(x_47); -lean_ctor_set_uint8(x_52, sizeof(void*)*2, x_53); -lean_ctor_set(x_45, 0, x_52); -return x_45; +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_18, 0); +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_18); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +lean_ctor_set(x_16, 0, x_22); +return x_16; } } else { -lean_object* x_54; uint8_t x_55; -x_54 = lean_ctor_get(x_45, 0); -lean_inc(x_54); -lean_dec(x_45); -x_55 = lean_unbox(x_54); -if (x_55 == 0) -{ -lean_dec(x_54); -x_14 = x_3; -x_15 = x_4; -x_16 = x_5; -x_17 = x_6; -x_18 = lean_box(0); -goto block_42; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_23 = lean_ctor_get(x_16, 0); +lean_inc(x_23); +lean_dec(x_16); +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +if (lean_is_exclusive(x_23)) { + lean_ctor_release(x_23, 0); + lean_ctor_release(x_23, 1); + x_26 = x_23; +} else { + lean_dec_ref(x_23); + x_26 = lean_box(0); } -else -{ -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; lean_object* x_61; -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -x_56 = l_Lean_Expr_letBody_x21(x_2); -lean_dec_ref(x_2); -x_57 = lean_expr_instantiate1(x_56, x_1); -lean_dec_ref(x_1); -lean_dec_ref(x_56); -x_58 = lean_box(0); -x_59 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_59, 0, x_57); -lean_ctor_set(x_59, 1, x_58); -x_60 = lean_unbox(x_54); -lean_dec(x_54); -lean_ctor_set_uint8(x_59, sizeof(void*)*2, x_60); -x_61 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_61, 0, x_59); -return x_61; +if (lean_is_scalar(x_26)) { + x_27 = lean_alloc_ctor(0, 2, 0); +} else { + x_27 = x_26; } +lean_ctor_set(x_27, 0, x_24); +lean_ctor_set(x_27, 1, x_25); +x_28 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_28, 0, x_27); +return x_28; } } else { -uint8_t x_62; -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_62 = !lean_is_exclusive(x_45); -if (x_62 == 0) +uint8_t x_29; +x_29 = !lean_is_exclusive(x_16); +if (x_29 == 0) { -return x_45; +return x_16; } else { -lean_object* x_63; lean_object* x_64; -x_63 = lean_ctor_get(x_45, 0); -lean_inc(x_63); -lean_dec(x_45); -x_64 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_64, 0, x_63); -return x_64; +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_16, 0); +lean_inc(x_30); +lean_dec(x_16); +x_31 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_31, 0, x_30); +return x_31; } } } -block_13: +} +LEAN_EXPORT lean_object* l_Lean_Meta_lambdaBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__9_spec__25(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_box(0); -x_10 = 1; -x_11 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_11, 0, x_2); -lean_ctor_set(x_11, 1, x_9); -lean_ctor_set_uint8(x_11, sizeof(void*)*2, x_10); -x_12 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_12, 0, x_11); -return x_12; +lean_object* x_13; +x_13 = l_Lean_Meta_lambdaBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__9_spec__25___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; } -block_42: +} +LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__12___redArg(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_19 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__21___closed__2; -x_20 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__0___redArg(x_19, x_16); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -lean_dec_ref(x_20); -x_22 = lean_unbox(x_21); -lean_dec(x_21); -if (x_22 == 0) +lean_object* x_11; uint8_t x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; +x_11 = lean_alloc_closure((void*)(l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__8___redArg___lam__0___boxed), 10, 3); +lean_closure_set(x_11, 0, x_2); +lean_closure_set(x_11, 1, x_4); +lean_closure_set(x_11, 2, x_5); +x_12 = 1; +x_13 = 0; +x_14 = lean_box(0); +x_15 = l___private_Lean_Meta_Basic_0__Lean_Meta_lambdaTelescopeImp(lean_box(0), x_1, x_12, x_13, x_12, x_13, x_14, x_11, x_3, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_15) == 0) { -lean_dec(x_17); -lean_dec_ref(x_16); -lean_dec(x_15); -lean_dec_ref(x_14); -lean_dec_ref(x_1); -x_8 = lean_box(0); -goto block_13; +uint8_t x_16; +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_ctor_get(x_15, 0); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +return x_15; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_23 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___closed__1; -x_24 = lean_unsigned_to_nat(30u); -lean_inc_ref(x_2); -x_25 = l_Lean_inlineExpr(x_2, x_24); -x_26 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_26, 0, x_23); -lean_ctor_set(x_26, 1, x_25); -x_27 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___closed__3; -x_28 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l_Lean_MessageData_ofExpr(x_1); -x_30 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -x_31 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___closed__5; -x_32 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -x_33 = l_Lean_Expr_letValue_x21(x_2); -x_34 = l_Lean_MessageData_ofExpr(x_33); -x_35 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_35, 0, x_32); -lean_ctor_set(x_35, 1, x_34); -x_36 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___redArg___lam__3___closed__3; -x_37 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_37, 0, x_35); -lean_ctor_set(x_37, 1, x_36); -x_38 = l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__1(x_19, x_37, x_14, x_15, x_16, x_17); +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_17, 0); +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_inc(x_19); lean_dec(x_17); -lean_dec_ref(x_16); -lean_dec(x_15); -lean_dec_ref(x_14); -if (lean_obj_tag(x_38) == 0) +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +lean_ctor_set(x_15, 0, x_21); +return x_15; +} +} +else { -lean_dec_ref(x_38); -x_8 = lean_box(0); -goto block_13; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_22 = lean_ctor_get(x_15, 0); +lean_inc(x_22); +lean_dec(x_15); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +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); +} +if (lean_is_scalar(x_25)) { + x_26 = lean_alloc_ctor(0, 2, 0); +} else { + x_26 = x_25; +} +lean_ctor_set(x_26, 0, x_23); +lean_ctor_set(x_26, 1, x_24); +x_27 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_27, 0, x_26); +return x_27; +} } else { -uint8_t x_39; -lean_dec_ref(x_2); -x_39 = !lean_is_exclusive(x_38); -if (x_39 == 0) +uint8_t x_28; +x_28 = !lean_is_exclusive(x_15); +if (x_28 == 0) { -return x_38; +return x_15; } else { -lean_object* x_40; lean_object* x_41; -x_40 = lean_ctor_get(x_38, 0); -lean_inc(x_40); -lean_dec(x_38); -x_41 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_41, 0, x_40); -return x_41; -} -} +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_15, 0); +lean_inc(x_29); +lean_dec(x_15); +x_30 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_30, 0, x_29); +return x_30; } } } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_7; -x_7 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__0(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -return x_7; +lean_object* x_12; +x_12 = l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__12___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_12; } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__1___redArg___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_4; -x_4 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__0___redArg(x_1, x_2); -lean_dec_ref(x_2); -return x_4; +lean_object* x_10; +x_10 = lean_apply_8(x_1, x_4, x_2, x_3, x_5, x_6, x_7, x_8, lean_box(0)); +return x_10; } } -LEAN_EXPORT lean_object* l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__1___redArg___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_8; -x_8 = l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__1(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -return x_8; +lean_object* x_10; +x_10 = l_Lean_Meta_withLetDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__1___redArg___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_10; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__1___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_8; -x_8 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith(x_1, x_2, x_3, x_4, x_5, x_6); -return x_8; -} +lean_object* x_14; lean_object* x_15; +x_14 = lean_alloc_closure((void*)(l_Lean_Meta_withLetDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__1___redArg___lam__0___boxed), 9, 3); +lean_closure_set(x_14, 0, x_4); +lean_closure_set(x_14, 1, x_7); +lean_closure_set(x_14, 2, x_8); +x_15 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLetDeclImp(lean_box(0), x_1, x_2, x_3, x_14, x_5, x_6, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_15) == 0) +{ +uint8_t x_16; +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_ctor_get(x_15, 0); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +return x_15; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwMData___redArg(lean_object* x_1) { -_start: +else { -lean_object* x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; -x_3 = l_Lean_Expr_consumeMData(x_1); -x_4 = lean_box(0); -x_5 = 1; -x_6 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_6, 0, x_3); -lean_ctor_set(x_6, 1, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*2, x_5); -x_7 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_7, 0, x_6); -return x_7; +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_17, 0); +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_17); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +lean_ctor_set(x_15, 0, x_21); +return x_15; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwMData(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +else { -lean_object* x_7; -x_7 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwMData___redArg(x_1); -return x_7; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_22 = lean_ctor_get(x_15, 0); +lean_inc(x_22); +lean_dec(x_15); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +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); } +if (lean_is_scalar(x_25)) { + x_26 = lean_alloc_ctor(0, 2, 0); +} else { + x_26 = x_25; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwMData___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +lean_ctor_set(x_26, 0, x_23); +lean_ctor_set(x_26, 1, x_24); +x_27 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_27, 0, x_26); +return x_27; +} +} +else { -lean_object* x_7; -x_7 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwMData(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -return x_7; +uint8_t x_28; +x_28 = !lean_is_exclusive(x_15); +if (x_28 == 0) +{ +return x_15; +} +else +{ +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_15, 0); +lean_inc(x_29); +lean_dec(x_15); +x_30 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_30, 0, x_29); +return x_30; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwMData___redArg___boxed(lean_object* x_1, lean_object* x_2) { +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_3; -x_3 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwMData___redArg(x_1); -lean_dec_ref(x_1); -return x_3; +lean_object* x_15; +x_15 = l_Lean_Meta_withLetDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__1___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_15; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__2___redArg(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -uint8_t x_8; -x_8 = lean_usize_dec_lt(x_6, x_5); -if (x_8 == 0) +lean_object* x_13; lean_object* x_14; +x_13 = lean_alloc_closure((void*)(l_Lean_Meta_withLetDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__1___redArg___lam__0___boxed), 9, 3); +lean_closure_set(x_13, 0, x_4); +lean_closure_set(x_13, 1, x_6); +lean_closure_set(x_13, 2, x_7); +x_14 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp(lean_box(0), x_1, x_2, x_3, x_13, x_5, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_14) == 0) { -lean_inc_ref(x_7); -return x_7; -} -else +uint8_t x_15; +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) { -lean_object* x_9; uint8_t x_10; -x_9 = lean_array_uget(x_4, x_6); -x_10 = l_Lean_Expr_isConstOf(x_1, x_9); -if (x_10 == 0) +lean_object* x_16; uint8_t x_17; +x_16 = lean_ctor_get(x_14, 0); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) { -size_t x_11; size_t x_12; -lean_dec(x_9); -x_11 = 1; -x_12 = lean_usize_add(x_6, x_11); +return x_14; +} +else { -size_t _tmp_5 = x_12; -lean_object* _tmp_6 = x_2; -x_6 = _tmp_5; -x_7 = _tmp_6; +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_16, 0); +x_19 = lean_ctor_get(x_16, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_16); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +lean_ctor_set(x_14, 0, x_20); +return x_14; } -goto _start; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_14, 0, x_9); -x_15 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_15, 0, x_14); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_3); -return x_16; +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_21 = lean_ctor_get(x_14, 0); +lean_inc(x_21); +lean_dec(x_14); +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +if (lean_is_exclusive(x_21)) { + lean_ctor_release(x_21, 0); + lean_ctor_release(x_21, 1); + x_24 = x_21; +} else { + lean_dec_ref(x_21); + x_24 = lean_box(0); } +if (lean_is_scalar(x_24)) { + x_25 = lean_alloc_ctor(0, 2, 0); +} else { + x_25 = x_24; } +lean_ctor_set(x_25, 0, x_22); +lean_ctor_set(x_25, 1, x_23); +x_26 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_26, 0, x_25); +return x_26; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__0() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Eq", 2, 2); -return x_1; -} +uint8_t x_27; +x_27 = !lean_is_exclusive(x_14); +if (x_27 == 0) +{ +return x_14; } -static lean_object* _init_l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__1() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__0; -x_2 = l_Lean_Name_mkStr1(x_1); -return x_2; +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_14, 0); +lean_inc(x_28); +lean_dec(x_14); +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, x_28); +return x_29; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Not an equality: `", 18, 18); -return x_1; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__3() { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__2(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__2; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_14; +x_14 = l_Lean_Meta_withLocalDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__2___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__4_spec__4_spec__7_spec__18_spec__28_spec__31_spec__33___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_11; -if (lean_obj_tag(x_3) == 5) +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = lean_st_ref_get(x_5); +x_16 = lean_ctor_get(x_15, 0); +lean_inc_ref(x_16); +lean_dec(x_15); +x_17 = l_Lean_Name_isAnonymous(x_2); +if (x_17 == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_17 = lean_ctor_get(x_3, 0); -lean_inc_ref(x_17); -x_18 = lean_ctor_get(x_3, 1); -lean_inc_ref(x_18); -lean_dec_ref(x_3); -x_19 = lean_array_set(x_4, x_5, x_18); -x_20 = lean_unsigned_to_nat(1u); -x_21 = lean_nat_sub(x_5, x_20); -lean_dec(x_5); -x_3 = x_17; -x_4 = x_19; -x_5 = x_21; -goto _start; +uint8_t x_18; +x_18 = lean_ctor_get_uint8(x_16, sizeof(void*)*8); +if (x_18 == 0) +{ +lean_dec_ref(x_16); +lean_dec(x_2); +x_7 = x_1; +x_8 = x_3; +x_9 = x_4; +x_10 = lean_box(0); +goto block_14; } else { -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; -lean_dec(x_5); -x_23 = lean_box(0); -x_24 = lean_box(0); -x_25 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanupAfter_cleanupAfter_x3f___lam__0___closed__0; -x_26 = lean_array_size(x_2); -x_27 = 0; -x_28 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__0(x_3, x_25, x_24, x_2, x_26, x_27, x_25); -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 = lean_ctor_get(x_28, 1); -lean_dec(x_31); -if (lean_obj_tag(x_30) == 0) +lean_object* x_19; uint8_t x_20; +lean_inc_ref(x_16); +x_19 = l_Lean_Environment_setExporting(x_16, x_17); +lean_inc(x_2); +lean_inc_ref(x_19); +x_20 = l_Lean_Environment_contains(x_19, x_2, x_18); +if (x_20 == 0) { -lean_free_object(x_28); -lean_dec(x_9); -lean_dec_ref(x_8); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -x_11 = lean_box(0); -goto block_16; +lean_dec_ref(x_19); +lean_dec_ref(x_16); +lean_dec(x_2); +x_7 = x_1; +x_8 = x_3; +x_9 = x_4; +x_10 = lean_box(0); +goto block_14; } else { -lean_object* x_32; -x_32 = lean_ctor_get(x_30, 0); -lean_inc(x_32); -lean_dec_ref(x_30); -if (lean_obj_tag(x_32) == 1) -{ -lean_object* x_33; uint8_t x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -lean_dec_ref(x_32); -x_34 = 1; -lean_inc(x_9); -lean_inc_ref(x_8); -lean_inc(x_7); -lean_inc_ref(x_6); -x_35 = l_Lean_Meta_getUnfoldEqnFor_x3f(x_33, x_34, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_35) == 0) -{ -uint8_t x_36; -x_36 = !lean_is_exclusive(x_35); -if (x_36 == 0) -{ -lean_object* x_37; -x_37 = lean_ctor_get(x_35, 0); -if (lean_obj_tag(x_37) == 1) -{ -uint8_t x_38; -lean_free_object(x_35); -lean_dec_ref(x_1); -x_38 = !lean_is_exclusive(x_37); -if (x_38 == 0) -{ -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_39 = lean_ctor_get(x_37, 0); -x_40 = l_Lean_Expr_constLevels_x21(x_3); -lean_dec_ref(x_3); -x_41 = l_Lean_mkConst(x_39, x_40); -x_42 = l_Lean_mkAppN(x_41, x_4); -lean_dec_ref(x_4); -lean_inc(x_9); -lean_inc_ref(x_8); -lean_inc(x_7); -lean_inc_ref(x_6); -lean_inc_ref(x_42); -x_43 = lean_infer_type(x_42, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_43) == 0) -{ -uint8_t x_44; -x_44 = !lean_is_exclusive(x_43); -if (x_44 == 0) -{ -lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; -x_45 = lean_ctor_get(x_43, 0); -x_46 = l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__1; -x_47 = lean_unsigned_to_nat(3u); -x_48 = l_Lean_Expr_isAppOfArity(x_45, x_46, x_47); -if (x_48 == 0) +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_21 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg___closed__2; +x_22 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg___closed__6; +x_23 = lean_box(0); +x_24 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_24, 0, x_19); +lean_ctor_set(x_24, 1, x_21); +lean_ctor_set(x_24, 2, x_22); +lean_ctor_set(x_24, 3, x_23); +lean_inc(x_2); +x_25 = l_Lean_MessageData_ofConstName(x_2, x_17); +x_26 = lean_alloc_ctor(3, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +x_27 = l_Lean_Environment_getModuleIdxFor_x3f(x_16, x_2); +if (lean_obj_tag(x_27) == 0) { -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; -lean_free_object(x_43); -lean_dec(x_45); -lean_free_object(x_37); -x_49 = l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__3; -x_50 = l_Lean_MessageData_ofExpr(x_42); -lean_ctor_set_tag(x_28, 7); -lean_ctor_set(x_28, 1, x_50); -lean_ctor_set(x_28, 0, x_49); -x_51 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___redArg___lam__3___closed__3; -x_52 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_52, 0, x_28); -lean_ctor_set(x_52, 1, x_51); -x_53 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___redArg(x_52, x_6, x_7, x_8, x_9); -lean_dec(x_9); -lean_dec_ref(x_8); -lean_dec(x_7); -lean_dec_ref(x_6); -return x_53; +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +lean_dec_ref(x_16); +lean_dec(x_2); +x_28 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg___closed__8; +x_29 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_26); +x_30 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg___closed__10; +x_31 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +x_32 = l_Lean_MessageData_note(x_31); +x_33 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_33, 0, x_1); +lean_ctor_set(x_33, 1, x_32); +x_7 = x_33; +x_8 = x_3; +x_9 = x_4; +x_10 = lean_box(0); +goto block_14; } else { -lean_object* x_54; lean_object* x_55; -lean_free_object(x_28); -lean_dec(x_9); -lean_dec_ref(x_8); -lean_dec(x_7); -lean_dec_ref(x_6); -x_54 = l_Lean_Expr_appArg_x21(x_45); -lean_dec(x_45); -lean_ctor_set(x_37, 0, x_42); -x_55 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_55, 0, x_54); -lean_ctor_set(x_55, 1, x_37); -lean_ctor_set_uint8(x_55, sizeof(void*)*2, x_34); -lean_ctor_set(x_43, 0, x_55); -return x_43; -} -} -else -{ -lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; -x_56 = lean_ctor_get(x_43, 0); -lean_inc(x_56); -lean_dec(x_43); -x_57 = l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__1; -x_58 = lean_unsigned_to_nat(3u); -x_59 = l_Lean_Expr_isAppOfArity(x_56, x_57, x_58); -if (x_59 == 0) -{ -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -lean_dec(x_56); -lean_free_object(x_37); -x_60 = l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__3; -x_61 = l_Lean_MessageData_ofExpr(x_42); -lean_ctor_set_tag(x_28, 7); -lean_ctor_set(x_28, 1, x_61); -lean_ctor_set(x_28, 0, x_60); -x_62 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___redArg___lam__3___closed__3; -x_63 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_63, 0, x_28); -lean_ctor_set(x_63, 1, x_62); -x_64 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___redArg(x_63, x_6, x_7, x_8, x_9); -lean_dec(x_9); -lean_dec_ref(x_8); -lean_dec(x_7); -lean_dec_ref(x_6); -return x_64; -} -else -{ -lean_object* x_65; lean_object* x_66; lean_object* x_67; -lean_free_object(x_28); -lean_dec(x_9); -lean_dec_ref(x_8); -lean_dec(x_7); -lean_dec_ref(x_6); -x_65 = l_Lean_Expr_appArg_x21(x_56); -lean_dec(x_56); -lean_ctor_set(x_37, 0, x_42); -x_66 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_66, 0, x_65); -lean_ctor_set(x_66, 1, x_37); -lean_ctor_set_uint8(x_66, sizeof(void*)*2, x_34); -x_67 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_67, 0, x_66); -return x_67; -} -} -} -else -{ -uint8_t x_68; -lean_dec_ref(x_42); -lean_free_object(x_37); -lean_free_object(x_28); -lean_dec(x_9); -lean_dec_ref(x_8); -lean_dec(x_7); -lean_dec_ref(x_6); -x_68 = !lean_is_exclusive(x_43); -if (x_68 == 0) -{ -return x_43; -} -else -{ -lean_object* x_69; lean_object* x_70; -x_69 = lean_ctor_get(x_43, 0); -lean_inc(x_69); -lean_dec(x_43); -x_70 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_70, 0, x_69); -return x_70; -} -} -} -else -{ -lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_71 = lean_ctor_get(x_37, 0); -lean_inc(x_71); -lean_dec(x_37); -x_72 = l_Lean_Expr_constLevels_x21(x_3); -lean_dec_ref(x_3); -x_73 = l_Lean_mkConst(x_71, x_72); -x_74 = l_Lean_mkAppN(x_73, x_4); -lean_dec_ref(x_4); -lean_inc(x_9); -lean_inc_ref(x_8); -lean_inc(x_7); -lean_inc_ref(x_6); -lean_inc_ref(x_74); -x_75 = lean_infer_type(x_74, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_75) == 0) -{ -lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; uint8_t x_80; -x_76 = lean_ctor_get(x_75, 0); -lean_inc(x_76); -if (lean_is_exclusive(x_75)) { - lean_ctor_release(x_75, 0); - x_77 = x_75; -} else { - lean_dec_ref(x_75); - x_77 = lean_box(0); -} -x_78 = l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__1; -x_79 = lean_unsigned_to_nat(3u); -x_80 = l_Lean_Expr_isAppOfArity(x_76, x_78, x_79); -if (x_80 == 0) -{ -lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -lean_dec(x_77); -lean_dec(x_76); -x_81 = l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__3; -x_82 = l_Lean_MessageData_ofExpr(x_74); -lean_ctor_set_tag(x_28, 7); -lean_ctor_set(x_28, 1, x_82); -lean_ctor_set(x_28, 0, x_81); -x_83 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___redArg___lam__3___closed__3; -x_84 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_84, 0, x_28); -lean_ctor_set(x_84, 1, x_83); -x_85 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___redArg(x_84, x_6, x_7, x_8, x_9); -lean_dec(x_9); -lean_dec_ref(x_8); -lean_dec(x_7); -lean_dec_ref(x_6); -return x_85; -} -else -{ -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -lean_free_object(x_28); -lean_dec(x_9); -lean_dec_ref(x_8); -lean_dec(x_7); -lean_dec_ref(x_6); -x_86 = l_Lean_Expr_appArg_x21(x_76); -lean_dec(x_76); -x_87 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_87, 0, x_74); -x_88 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_88, 0, x_86); -lean_ctor_set(x_88, 1, x_87); -lean_ctor_set_uint8(x_88, sizeof(void*)*2, x_34); -if (lean_is_scalar(x_77)) { - x_89 = lean_alloc_ctor(0, 1, 0); -} else { - x_89 = x_77; -} -lean_ctor_set(x_89, 0, x_88); -return x_89; -} -} -else -{ -lean_object* x_90; lean_object* x_91; lean_object* x_92; -lean_dec_ref(x_74); -lean_free_object(x_28); -lean_dec(x_9); -lean_dec_ref(x_8); -lean_dec(x_7); -lean_dec_ref(x_6); -x_90 = lean_ctor_get(x_75, 0); -lean_inc(x_90); -if (lean_is_exclusive(x_75)) { - lean_ctor_release(x_75, 0); - x_91 = x_75; -} else { - lean_dec_ref(x_75); - x_91 = lean_box(0); -} -if (lean_is_scalar(x_91)) { - x_92 = lean_alloc_ctor(1, 1, 0); -} else { - x_92 = x_91; -} -lean_ctor_set(x_92, 0, x_90); -return x_92; -} -} -} -else -{ -lean_object* x_93; -lean_dec(x_37); -lean_free_object(x_28); -lean_dec(x_9); -lean_dec_ref(x_8); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -x_93 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_93, 0, x_1); -lean_ctor_set(x_93, 1, x_23); -lean_ctor_set_uint8(x_93, sizeof(void*)*2, x_34); -lean_ctor_set(x_35, 0, x_93); -return x_35; -} -} -else -{ -lean_object* x_94; -x_94 = lean_ctor_get(x_35, 0); -lean_inc(x_94); -lean_dec(x_35); -if (lean_obj_tag(x_94) == 1) -{ -lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; -lean_dec_ref(x_1); -x_95 = lean_ctor_get(x_94, 0); -lean_inc(x_95); -if (lean_is_exclusive(x_94)) { - lean_ctor_release(x_94, 0); - x_96 = x_94; -} else { - lean_dec_ref(x_94); - x_96 = lean_box(0); -} -x_97 = l_Lean_Expr_constLevels_x21(x_3); -lean_dec_ref(x_3); -x_98 = l_Lean_mkConst(x_95, x_97); -x_99 = l_Lean_mkAppN(x_98, x_4); -lean_dec_ref(x_4); -lean_inc(x_9); -lean_inc_ref(x_8); -lean_inc(x_7); -lean_inc_ref(x_6); -lean_inc_ref(x_99); -x_100 = lean_infer_type(x_99, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_100) == 0) -{ -lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; uint8_t x_105; -x_101 = lean_ctor_get(x_100, 0); -lean_inc(x_101); -if (lean_is_exclusive(x_100)) { - lean_ctor_release(x_100, 0); - x_102 = x_100; -} else { - lean_dec_ref(x_100); - x_102 = lean_box(0); -} -x_103 = l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__1; -x_104 = lean_unsigned_to_nat(3u); -x_105 = l_Lean_Expr_isAppOfArity(x_101, x_103, x_104); -if (x_105 == 0) -{ -lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; -lean_dec(x_102); -lean_dec(x_101); -lean_dec(x_96); -x_106 = l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__3; -x_107 = l_Lean_MessageData_ofExpr(x_99); -lean_ctor_set_tag(x_28, 7); -lean_ctor_set(x_28, 1, x_107); -lean_ctor_set(x_28, 0, x_106); -x_108 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___redArg___lam__3___closed__3; -x_109 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_109, 0, x_28); -lean_ctor_set(x_109, 1, x_108); -x_110 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___redArg(x_109, x_6, x_7, x_8, x_9); -lean_dec(x_9); -lean_dec_ref(x_8); -lean_dec(x_7); -lean_dec_ref(x_6); -return x_110; -} -else -{ -lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; -lean_free_object(x_28); -lean_dec(x_9); -lean_dec_ref(x_8); -lean_dec(x_7); -lean_dec_ref(x_6); -x_111 = l_Lean_Expr_appArg_x21(x_101); -lean_dec(x_101); -if (lean_is_scalar(x_96)) { - x_112 = lean_alloc_ctor(1, 1, 0); -} else { - x_112 = x_96; -} -lean_ctor_set(x_112, 0, x_99); -x_113 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_113, 0, x_111); -lean_ctor_set(x_113, 1, x_112); -lean_ctor_set_uint8(x_113, sizeof(void*)*2, x_34); -if (lean_is_scalar(x_102)) { - x_114 = lean_alloc_ctor(0, 1, 0); -} else { - x_114 = x_102; -} -lean_ctor_set(x_114, 0, x_113); -return x_114; -} -} -else -{ -lean_object* x_115; lean_object* x_116; lean_object* x_117; -lean_dec_ref(x_99); -lean_dec(x_96); -lean_free_object(x_28); -lean_dec(x_9); -lean_dec_ref(x_8); -lean_dec(x_7); -lean_dec_ref(x_6); -x_115 = lean_ctor_get(x_100, 0); -lean_inc(x_115); -if (lean_is_exclusive(x_100)) { - lean_ctor_release(x_100, 0); - x_116 = x_100; -} else { - lean_dec_ref(x_100); - x_116 = lean_box(0); -} -if (lean_is_scalar(x_116)) { - x_117 = lean_alloc_ctor(1, 1, 0); -} else { - x_117 = x_116; -} -lean_ctor_set(x_117, 0, x_115); -return x_117; -} -} -else -{ -lean_object* x_118; lean_object* x_119; -lean_dec(x_94); -lean_free_object(x_28); -lean_dec(x_9); -lean_dec_ref(x_8); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -x_118 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_118, 0, x_1); -lean_ctor_set(x_118, 1, x_23); -lean_ctor_set_uint8(x_118, sizeof(void*)*2, x_34); -x_119 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_119, 0, x_118); -return x_119; -} -} -} -else -{ -uint8_t x_120; -lean_free_object(x_28); -lean_dec(x_9); -lean_dec_ref(x_8); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -lean_dec_ref(x_1); -x_120 = !lean_is_exclusive(x_35); -if (x_120 == 0) -{ -return x_35; -} -else -{ -lean_object* x_121; lean_object* x_122; -x_121 = lean_ctor_get(x_35, 0); -lean_inc(x_121); -lean_dec(x_35); -x_122 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_122, 0, x_121); -return x_122; -} -} -} -else -{ -lean_dec(x_32); -lean_free_object(x_28); -lean_dec(x_9); -lean_dec_ref(x_8); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -x_11 = lean_box(0); -goto block_16; -} -} -} -else -{ -lean_object* x_123; -x_123 = lean_ctor_get(x_28, 0); -lean_inc(x_123); -lean_dec(x_28); -if (lean_obj_tag(x_123) == 0) -{ -lean_dec(x_9); -lean_dec_ref(x_8); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -x_11 = lean_box(0); -goto block_16; -} -else -{ -lean_object* x_124; -x_124 = lean_ctor_get(x_123, 0); -lean_inc(x_124); -lean_dec_ref(x_123); -if (lean_obj_tag(x_124) == 1) -{ -lean_object* x_125; uint8_t x_126; lean_object* x_127; -x_125 = lean_ctor_get(x_124, 0); -lean_inc(x_125); -lean_dec_ref(x_124); -x_126 = 1; -lean_inc(x_9); -lean_inc_ref(x_8); -lean_inc(x_7); -lean_inc_ref(x_6); -x_127 = l_Lean_Meta_getUnfoldEqnFor_x3f(x_125, x_126, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_127) == 0) -{ -lean_object* x_128; lean_object* x_129; -x_128 = lean_ctor_get(x_127, 0); -lean_inc(x_128); -if (lean_is_exclusive(x_127)) { - lean_ctor_release(x_127, 0); - x_129 = x_127; -} else { - lean_dec_ref(x_127); - x_129 = lean_box(0); -} -if (lean_obj_tag(x_128) == 1) -{ -lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; -lean_dec(x_129); -lean_dec_ref(x_1); -x_130 = lean_ctor_get(x_128, 0); -lean_inc(x_130); -if (lean_is_exclusive(x_128)) { - lean_ctor_release(x_128, 0); - x_131 = x_128; -} else { - lean_dec_ref(x_128); - x_131 = lean_box(0); -} -x_132 = l_Lean_Expr_constLevels_x21(x_3); -lean_dec_ref(x_3); -x_133 = l_Lean_mkConst(x_130, x_132); -x_134 = l_Lean_mkAppN(x_133, x_4); -lean_dec_ref(x_4); -lean_inc(x_9); -lean_inc_ref(x_8); -lean_inc(x_7); -lean_inc_ref(x_6); -lean_inc_ref(x_134); -x_135 = lean_infer_type(x_134, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_135) == 0) -{ -lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; uint8_t x_140; -x_136 = lean_ctor_get(x_135, 0); -lean_inc(x_136); -if (lean_is_exclusive(x_135)) { - lean_ctor_release(x_135, 0); - x_137 = x_135; -} else { - lean_dec_ref(x_135); - x_137 = lean_box(0); -} -x_138 = l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__1; -x_139 = lean_unsigned_to_nat(3u); -x_140 = l_Lean_Expr_isAppOfArity(x_136, x_138, x_139); -if (x_140 == 0) -{ -lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; -lean_dec(x_137); -lean_dec(x_136); -lean_dec(x_131); -x_141 = l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__3; -x_142 = l_Lean_MessageData_ofExpr(x_134); -x_143 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_143, 0, x_141); -lean_ctor_set(x_143, 1, x_142); -x_144 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___redArg___lam__3___closed__3; -x_145 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_145, 0, x_143); -lean_ctor_set(x_145, 1, x_144); -x_146 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___redArg(x_145, x_6, x_7, x_8, x_9); -lean_dec(x_9); -lean_dec_ref(x_8); -lean_dec(x_7); -lean_dec_ref(x_6); -return x_146; -} -else -{ -lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; -lean_dec(x_9); -lean_dec_ref(x_8); -lean_dec(x_7); -lean_dec_ref(x_6); -x_147 = l_Lean_Expr_appArg_x21(x_136); -lean_dec(x_136); -if (lean_is_scalar(x_131)) { - x_148 = lean_alloc_ctor(1, 1, 0); -} else { - x_148 = x_131; -} -lean_ctor_set(x_148, 0, x_134); -x_149 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_149, 0, x_147); -lean_ctor_set(x_149, 1, x_148); -lean_ctor_set_uint8(x_149, sizeof(void*)*2, x_126); -if (lean_is_scalar(x_137)) { - x_150 = lean_alloc_ctor(0, 1, 0); -} else { - x_150 = x_137; -} -lean_ctor_set(x_150, 0, x_149); -return x_150; -} -} -else -{ -lean_object* x_151; lean_object* x_152; lean_object* x_153; -lean_dec_ref(x_134); -lean_dec(x_131); -lean_dec(x_9); -lean_dec_ref(x_8); -lean_dec(x_7); -lean_dec_ref(x_6); -x_151 = lean_ctor_get(x_135, 0); -lean_inc(x_151); -if (lean_is_exclusive(x_135)) { - lean_ctor_release(x_135, 0); - x_152 = x_135; -} else { - lean_dec_ref(x_135); - x_152 = lean_box(0); -} -if (lean_is_scalar(x_152)) { - x_153 = lean_alloc_ctor(1, 1, 0); -} else { - x_153 = x_152; -} -lean_ctor_set(x_153, 0, x_151); -return x_153; -} -} -else -{ -lean_object* x_154; lean_object* x_155; -lean_dec(x_128); -lean_dec(x_9); -lean_dec_ref(x_8); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -x_154 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_154, 0, x_1); -lean_ctor_set(x_154, 1, x_23); -lean_ctor_set_uint8(x_154, sizeof(void*)*2, x_126); -if (lean_is_scalar(x_129)) { - x_155 = lean_alloc_ctor(0, 1, 0); -} else { - x_155 = x_129; -} -lean_ctor_set(x_155, 0, x_154); -return x_155; -} -} -else -{ -lean_object* x_156; lean_object* x_157; lean_object* x_158; -lean_dec(x_9); -lean_dec_ref(x_8); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -lean_dec_ref(x_1); -x_156 = lean_ctor_get(x_127, 0); -lean_inc(x_156); -if (lean_is_exclusive(x_127)) { - lean_ctor_release(x_127, 0); - x_157 = x_127; -} else { - lean_dec_ref(x_127); - x_157 = lean_box(0); -} -if (lean_is_scalar(x_157)) { - x_158 = lean_alloc_ctor(1, 1, 0); -} else { - x_158 = x_157; -} -lean_ctor_set(x_158, 0, x_156); -return x_158; -} -} -else -{ -lean_dec(x_124); -lean_dec(x_9); -lean_dec_ref(x_8); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -x_11 = lean_box(0); -goto block_16; -} -} -} -} -block_16: -{ -lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; -x_12 = lean_box(0); -x_13 = 1; -x_14 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_14, 0, x_1); -lean_ctor_set(x_14, 1, x_12); -lean_ctor_set_uint8(x_14, sizeof(void*)*2, x_13); -x_15 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_15, 0, x_14); -return x_15; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_8 = l___private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0___lam__1___closed__0; -x_9 = l_Lean_Expr_getAppNumArgs(x_2); -lean_inc(x_9); -x_10 = lean_mk_array(x_9, x_8); -x_11 = lean_unsigned_to_nat(1u); -x_12 = lean_nat_sub(x_9, x_11); -lean_dec(x_9); -lean_inc_ref(x_2); -x_13 = l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1(x_2, x_1, x_2, x_10, x_12, x_3, x_4, x_5, x_6); -return x_13; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -x_8 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec_ref(x_1); -return x_8; -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -size_t x_8; size_t x_9; lean_object* x_10; -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_10 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__0(x_1, x_2, x_3, x_4, x_8, x_9, x_7); -lean_dec_ref(x_7); -lean_dec_ref(x_4); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec_ref(x_2); -return x_11; -} -} -static size_t _init_l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3___redArg___closed__0() { -_start: -{ -size_t x_1; size_t x_2; size_t x_3; -x_1 = 5; -x_2 = 1; -x_3 = lean_usize_shift_left(x_2, x_1); -return x_3; -} -} -static size_t _init_l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3___redArg___closed__1() { -_start: -{ -size_t x_1; size_t x_2; size_t x_3; -x_1 = 1; -x_2 = l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3___redArg___closed__0; -x_3 = lean_usize_sub(x_2, x_1); -return x_3; -} -} -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at___00Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3_spec__12___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; uint8_t x_5; -x_4 = lean_array_get_size(x_1); -x_5 = lean_nat_dec_lt(x_2, x_4); -if (x_5 == 0) -{ -lean_dec(x_2); -return x_5; -} -else -{ -lean_object* x_6; uint8_t x_7; -x_6 = lean_array_fget_borrowed(x_1, x_2); -x_7 = l_Lean_instBEqMVarId_beq(x_3, x_6); -if (x_7 == 0) -{ -lean_object* x_8; lean_object* x_9; -x_8 = lean_unsigned_to_nat(1u); -x_9 = lean_nat_add(x_2, x_8); -lean_dec(x_2); -x_2 = x_9; -goto _start; -} -else -{ -lean_dec(x_2); -return x_7; -} -} -} -} -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3___redArg(lean_object* x_1, size_t x_2, lean_object* x_3) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_4; lean_object* x_5; size_t x_6; size_t x_7; size_t x_8; lean_object* x_9; lean_object* x_10; -x_4 = lean_ctor_get(x_1, 0); -lean_inc_ref(x_4); -lean_dec_ref(x_1); -x_5 = lean_box(2); -x_6 = 5; -x_7 = l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3___redArg___closed__1; -x_8 = lean_usize_land(x_2, x_7); -x_9 = lean_usize_to_nat(x_8); -x_10 = lean_array_get(x_5, x_4, x_9); -lean_dec(x_9); -lean_dec_ref(x_4); -switch (lean_obj_tag(x_10)) { -case 0: -{ -lean_object* x_11; uint8_t x_12; -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -lean_dec_ref(x_10); -x_12 = l_Lean_instBEqMVarId_beq(x_3, x_11); -lean_dec(x_11); -return x_12; -} -case 1: -{ -lean_object* x_13; size_t x_14; -x_13 = lean_ctor_get(x_10, 0); -lean_inc(x_13); -lean_dec_ref(x_10); -x_14 = lean_usize_shift_right(x_2, x_6); -x_1 = x_13; -x_2 = x_14; -goto _start; -} -default: -{ -uint8_t x_16; -x_16 = 0; -return x_16; -} -} -} -else -{ -lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_17 = lean_ctor_get(x_1, 0); -lean_inc_ref(x_17); -lean_dec_ref(x_1); -x_18 = lean_unsigned_to_nat(0u); -x_19 = l_Lean_PersistentHashMap_containsAtAux___at___00Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3_spec__12___redArg(x_17, x_18, x_3); -lean_dec_ref(x_17); -return x_19; -} -} -} -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0___redArg(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint64_t x_3; size_t x_4; uint8_t x_5; -x_3 = l_Lean_instHashableMVarId_hash(x_2); -x_4 = lean_uint64_to_usize(x_3); -x_5 = l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3___redArg(x_1, x_4, x_2); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0___redArg(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; lean_object* x_9; -x_4 = lean_st_ref_get(x_2); -x_5 = lean_ctor_get(x_4, 0); -lean_inc_ref(x_5); -lean_dec(x_4); -x_6 = lean_ctor_get(x_5, 7); -lean_inc_ref(x_6); -lean_dec_ref(x_5); -x_7 = l_Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0___redArg(x_6, x_1); -x_8 = lean_box(x_7); -x_9 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_9, 0, x_8); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_7; -x_7 = l_Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0___redArg(x_1, x_3); -return x_7; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_isMatcherApp___at___00Lean_Tactic_FunInd_rwMatcher_spec__1___redArg(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_4; lean_object* x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; -x_4 = lean_st_ref_get(x_2); -x_5 = lean_ctor_get(x_4, 0); -lean_inc_ref(x_5); -lean_dec(x_4); -x_6 = l_Lean_Meta_isMatcherAppCore(x_5, x_1); -x_7 = lean_box(x_6); -x_8 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_8, 0, x_7); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_isMatcherApp___at___00Lean_Tactic_FunInd_rwMatcher_spec__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_7; -x_7 = l_Lean_Meta_isMatcherApp___at___00Lean_Tactic_FunInd_rwMatcher_spec__1___redArg(x_1, x_5); -return x_7; -} -} -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -uint8_t x_4; -x_4 = l_Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0___redArg(x_2, x_3); -return x_4; -} -} -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at___00Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3_spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; -x_7 = l_Lean_PersistentHashMap_containsAtAux___at___00Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3_spec__12___redArg(x_2, x_5, x_6); -return x_7; -} -} -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3(lean_object* x_1, lean_object* x_2, size_t x_3, lean_object* x_4) { -_start: -{ -uint8_t x_5; -x_5 = l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3___redArg(x_2, x_3, x_4); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00Lean_Tactic_FunInd_rwMatcher_spec__3___redArg(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_4; -x_4 = l_Lean_Expr_hasMVar(x_1); -if (x_4 == 0) -{ -lean_object* x_5; -x_5 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_5, 0, x_1); -return x_5; -} -else -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_6 = lean_st_ref_get(x_2); -x_7 = lean_ctor_get(x_6, 0); -lean_inc_ref(x_7); -lean_dec(x_6); -x_8 = l_Lean_instantiateMVarsCore(x_7, x_1); -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_ref(x_8); -x_11 = lean_st_ref_take(x_2); -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = lean_ctor_get(x_11, 0); -lean_dec(x_13); -lean_ctor_set(x_11, 0, x_10); -x_14 = lean_st_ref_set(x_2, x_11); -x_15 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_15, 0, x_9); -return x_15; -} -else -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_16 = lean_ctor_get(x_11, 1); -x_17 = lean_ctor_get(x_11, 2); -x_18 = lean_ctor_get(x_11, 3); -x_19 = lean_ctor_get(x_11, 4); -lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_11); -x_20 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_20, 0, x_10); -lean_ctor_set(x_20, 1, x_16); -lean_ctor_set(x_20, 2, x_17); -lean_ctor_set(x_20, 3, x_18); -lean_ctor_set(x_20, 4, x_19); -x_21 = lean_st_ref_set(x_2, x_20); -x_22 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_22, 0, x_9); -return x_22; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00Lean_Tactic_FunInd_rwMatcher_spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_7; -x_7 = l_Lean_instantiateMVars___at___00Lean_Tactic_FunInd_rwMatcher_spec__3___redArg(x_1, x_3); -return x_7; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -uint8_t x_10; -x_10 = !lean_is_exclusive(x_7); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_11 = lean_ctor_get(x_7, 5); -x_12 = lean_st_ref_get(x_8); -x_13 = lean_ctor_get(x_12, 4); -lean_inc_ref(x_13); -lean_dec(x_12); -x_14 = lean_ctor_get(x_13, 0); -lean_inc_ref(x_14); -lean_dec_ref(x_13); -x_15 = l_Lean_replaceRef(x_3, x_11); -lean_dec(x_11); -lean_ctor_set(x_7, 5, x_15); -x_16 = l_Lean_PersistentArray_toArray___redArg(x_14); -lean_dec_ref(x_14); -x_17 = lean_array_size(x_16); -x_18 = 0; -x_19 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13_spec__35_spec__39(x_17, x_18, x_16); -x_20 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_20, 0, x_2); -lean_ctor_set(x_20, 1, x_4); -lean_ctor_set(x_20, 2, x_19); -x_21 = l_Lean_addMessageContextFull___at___00Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__2_spec__7(x_20, x_5, x_6, x_7, x_8); -lean_dec_ref(x_7); -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) -{ -lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_23 = lean_ctor_get(x_21, 0); -x_24 = lean_st_ref_take(x_8); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) -{ -lean_object* x_26; uint8_t x_27; -x_26 = lean_ctor_get(x_24, 4); -x_27 = !lean_is_exclusive(x_26); -if (x_27 == 0) -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_28 = lean_ctor_get(x_26, 0); -lean_dec(x_28); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_3); -lean_ctor_set(x_29, 1, x_23); -x_30 = l_Lean_PersistentArray_push___redArg(x_1, x_29); -lean_ctor_set(x_26, 0, x_30); -x_31 = lean_st_ref_set(x_8, x_24); -x_32 = lean_box(0); -lean_ctor_set(x_21, 0, x_32); -return x_21; -} -else -{ -uint64_t x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_33 = lean_ctor_get_uint64(x_26, sizeof(void*)*1); -lean_dec(x_26); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_3); -lean_ctor_set(x_34, 1, x_23); -x_35 = l_Lean_PersistentArray_push___redArg(x_1, x_34); -x_36 = lean_alloc_ctor(0, 1, 8); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set_uint64(x_36, sizeof(void*)*1, x_33); -lean_ctor_set(x_24, 4, x_36); -x_37 = lean_st_ref_set(x_8, x_24); -x_38 = lean_box(0); -lean_ctor_set(x_21, 0, x_38); -return x_21; -} -} -else -{ -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint64_t x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_39 = lean_ctor_get(x_24, 4); -x_40 = lean_ctor_get(x_24, 0); -x_41 = lean_ctor_get(x_24, 1); -x_42 = lean_ctor_get(x_24, 2); -x_43 = lean_ctor_get(x_24, 3); -x_44 = lean_ctor_get(x_24, 5); -x_45 = lean_ctor_get(x_24, 6); -x_46 = lean_ctor_get(x_24, 7); -x_47 = lean_ctor_get(x_24, 8); -lean_inc(x_47); -lean_inc(x_46); -lean_inc(x_45); -lean_inc(x_44); -lean_inc(x_39); -lean_inc(x_43); -lean_inc(x_42); -lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_24); -x_48 = lean_ctor_get_uint64(x_39, sizeof(void*)*1); -if (lean_is_exclusive(x_39)) { - lean_ctor_release(x_39, 0); - x_49 = x_39; -} else { - lean_dec_ref(x_39); - x_49 = lean_box(0); -} -x_50 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_50, 0, x_3); -lean_ctor_set(x_50, 1, x_23); -x_51 = l_Lean_PersistentArray_push___redArg(x_1, x_50); -if (lean_is_scalar(x_49)) { - x_52 = lean_alloc_ctor(0, 1, 8); -} else { - x_52 = x_49; -} -lean_ctor_set(x_52, 0, x_51); -lean_ctor_set_uint64(x_52, sizeof(void*)*1, x_48); -x_53 = lean_alloc_ctor(0, 9, 0); -lean_ctor_set(x_53, 0, x_40); -lean_ctor_set(x_53, 1, x_41); -lean_ctor_set(x_53, 2, x_42); -lean_ctor_set(x_53, 3, x_43); -lean_ctor_set(x_53, 4, x_52); -lean_ctor_set(x_53, 5, x_44); -lean_ctor_set(x_53, 6, x_45); -lean_ctor_set(x_53, 7, x_46); -lean_ctor_set(x_53, 8, x_47); -x_54 = lean_st_ref_set(x_8, x_53); -x_55 = lean_box(0); -lean_ctor_set(x_21, 0, x_55); -return x_21; -} -} -else -{ -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; uint64_t x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_56 = lean_ctor_get(x_21, 0); -lean_inc(x_56); -lean_dec(x_21); -x_57 = lean_st_ref_take(x_8); -x_58 = lean_ctor_get(x_57, 4); -lean_inc_ref(x_58); -x_59 = lean_ctor_get(x_57, 0); -lean_inc_ref(x_59); -x_60 = lean_ctor_get(x_57, 1); -lean_inc(x_60); -x_61 = lean_ctor_get(x_57, 2); -lean_inc_ref(x_61); -x_62 = lean_ctor_get(x_57, 3); -lean_inc_ref(x_62); -x_63 = lean_ctor_get(x_57, 5); -lean_inc_ref(x_63); -x_64 = lean_ctor_get(x_57, 6); -lean_inc_ref(x_64); -x_65 = lean_ctor_get(x_57, 7); -lean_inc_ref(x_65); -x_66 = lean_ctor_get(x_57, 8); -lean_inc_ref(x_66); -if (lean_is_exclusive(x_57)) { - lean_ctor_release(x_57, 0); - lean_ctor_release(x_57, 1); - lean_ctor_release(x_57, 2); - lean_ctor_release(x_57, 3); - lean_ctor_release(x_57, 4); - lean_ctor_release(x_57, 5); - lean_ctor_release(x_57, 6); - lean_ctor_release(x_57, 7); - lean_ctor_release(x_57, 8); - x_67 = x_57; -} else { - lean_dec_ref(x_57); - x_67 = lean_box(0); -} -x_68 = lean_ctor_get_uint64(x_58, sizeof(void*)*1); -if (lean_is_exclusive(x_58)) { - lean_ctor_release(x_58, 0); - x_69 = x_58; -} else { - lean_dec_ref(x_58); - x_69 = lean_box(0); -} -x_70 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_70, 0, x_3); -lean_ctor_set(x_70, 1, x_56); -x_71 = l_Lean_PersistentArray_push___redArg(x_1, x_70); -if (lean_is_scalar(x_69)) { - x_72 = lean_alloc_ctor(0, 1, 8); -} else { - x_72 = x_69; -} -lean_ctor_set(x_72, 0, x_71); -lean_ctor_set_uint64(x_72, sizeof(void*)*1, x_68); -if (lean_is_scalar(x_67)) { - x_73 = lean_alloc_ctor(0, 9, 0); -} else { - x_73 = x_67; -} -lean_ctor_set(x_73, 0, x_59); -lean_ctor_set(x_73, 1, x_60); -lean_ctor_set(x_73, 2, x_61); -lean_ctor_set(x_73, 3, x_62); -lean_ctor_set(x_73, 4, x_72); -lean_ctor_set(x_73, 5, x_63); -lean_ctor_set(x_73, 6, x_64); -lean_ctor_set(x_73, 7, x_65); -lean_ctor_set(x_73, 8, x_66); -x_74 = lean_st_ref_set(x_8, x_73); -x_75 = lean_box(0); -x_76 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_76, 0, x_75); -return x_76; -} -} -else -{ -lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; lean_object* x_90; uint8_t x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; size_t x_99; size_t x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; uint64_t x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; -x_77 = lean_ctor_get(x_7, 0); -x_78 = lean_ctor_get(x_7, 1); -x_79 = lean_ctor_get(x_7, 2); -x_80 = lean_ctor_get(x_7, 3); -x_81 = lean_ctor_get(x_7, 4); -x_82 = lean_ctor_get(x_7, 5); -x_83 = lean_ctor_get(x_7, 6); -x_84 = lean_ctor_get(x_7, 7); -x_85 = lean_ctor_get(x_7, 8); -x_86 = lean_ctor_get(x_7, 9); -x_87 = lean_ctor_get(x_7, 10); -x_88 = lean_ctor_get(x_7, 11); -x_89 = lean_ctor_get_uint8(x_7, sizeof(void*)*14); -x_90 = lean_ctor_get(x_7, 12); -x_91 = lean_ctor_get_uint8(x_7, sizeof(void*)*14 + 1); -x_92 = lean_ctor_get(x_7, 13); -lean_inc(x_92); -lean_inc(x_90); -lean_inc(x_88); -lean_inc(x_87); -lean_inc(x_86); -lean_inc(x_85); -lean_inc(x_84); -lean_inc(x_83); -lean_inc(x_82); -lean_inc(x_81); -lean_inc(x_80); -lean_inc(x_79); -lean_inc(x_78); -lean_inc(x_77); -lean_dec(x_7); -x_93 = lean_st_ref_get(x_8); -x_94 = lean_ctor_get(x_93, 4); -lean_inc_ref(x_94); -lean_dec(x_93); -x_95 = lean_ctor_get(x_94, 0); -lean_inc_ref(x_95); -lean_dec_ref(x_94); -x_96 = l_Lean_replaceRef(x_3, x_82); -lean_dec(x_82); -x_97 = lean_alloc_ctor(0, 14, 2); -lean_ctor_set(x_97, 0, x_77); -lean_ctor_set(x_97, 1, x_78); -lean_ctor_set(x_97, 2, x_79); -lean_ctor_set(x_97, 3, x_80); -lean_ctor_set(x_97, 4, x_81); -lean_ctor_set(x_97, 5, x_96); -lean_ctor_set(x_97, 6, x_83); -lean_ctor_set(x_97, 7, x_84); -lean_ctor_set(x_97, 8, x_85); -lean_ctor_set(x_97, 9, x_86); -lean_ctor_set(x_97, 10, x_87); -lean_ctor_set(x_97, 11, x_88); -lean_ctor_set(x_97, 12, x_90); -lean_ctor_set(x_97, 13, x_92); -lean_ctor_set_uint8(x_97, sizeof(void*)*14, x_89); -lean_ctor_set_uint8(x_97, sizeof(void*)*14 + 1, x_91); -x_98 = l_Lean_PersistentArray_toArray___redArg(x_95); -lean_dec_ref(x_95); -x_99 = lean_array_size(x_98); -x_100 = 0; -x_101 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13_spec__35_spec__39(x_99, x_100, x_98); -x_102 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_102, 0, x_2); -lean_ctor_set(x_102, 1, x_4); -lean_ctor_set(x_102, 2, x_101); -x_103 = l_Lean_addMessageContextFull___at___00Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__2_spec__7(x_102, x_5, x_6, x_97, x_8); -lean_dec_ref(x_97); -x_104 = lean_ctor_get(x_103, 0); -lean_inc(x_104); -if (lean_is_exclusive(x_103)) { - lean_ctor_release(x_103, 0); - x_105 = x_103; -} else { - lean_dec_ref(x_103); - x_105 = lean_box(0); -} -x_106 = lean_st_ref_take(x_8); -x_107 = lean_ctor_get(x_106, 4); -lean_inc_ref(x_107); -x_108 = lean_ctor_get(x_106, 0); -lean_inc_ref(x_108); -x_109 = lean_ctor_get(x_106, 1); -lean_inc(x_109); -x_110 = lean_ctor_get(x_106, 2); -lean_inc_ref(x_110); -x_111 = lean_ctor_get(x_106, 3); -lean_inc_ref(x_111); -x_112 = lean_ctor_get(x_106, 5); -lean_inc_ref(x_112); -x_113 = lean_ctor_get(x_106, 6); -lean_inc_ref(x_113); -x_114 = lean_ctor_get(x_106, 7); -lean_inc_ref(x_114); -x_115 = lean_ctor_get(x_106, 8); -lean_inc_ref(x_115); -if (lean_is_exclusive(x_106)) { - lean_ctor_release(x_106, 0); - lean_ctor_release(x_106, 1); - lean_ctor_release(x_106, 2); - lean_ctor_release(x_106, 3); - lean_ctor_release(x_106, 4); - lean_ctor_release(x_106, 5); - lean_ctor_release(x_106, 6); - lean_ctor_release(x_106, 7); - lean_ctor_release(x_106, 8); - x_116 = x_106; -} else { - lean_dec_ref(x_106); - x_116 = lean_box(0); -} -x_117 = lean_ctor_get_uint64(x_107, sizeof(void*)*1); -if (lean_is_exclusive(x_107)) { - lean_ctor_release(x_107, 0); - x_118 = x_107; -} else { - lean_dec_ref(x_107); - x_118 = lean_box(0); -} -x_119 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_119, 0, x_3); -lean_ctor_set(x_119, 1, x_104); -x_120 = l_Lean_PersistentArray_push___redArg(x_1, x_119); -if (lean_is_scalar(x_118)) { - x_121 = lean_alloc_ctor(0, 1, 8); -} else { - x_121 = x_118; -} -lean_ctor_set(x_121, 0, x_120); -lean_ctor_set_uint64(x_121, sizeof(void*)*1, x_117); -if (lean_is_scalar(x_116)) { - x_122 = lean_alloc_ctor(0, 9, 0); -} else { - x_122 = x_116; -} -lean_ctor_set(x_122, 0, x_108); -lean_ctor_set(x_122, 1, x_109); -lean_ctor_set(x_122, 2, x_110); -lean_ctor_set(x_122, 3, x_111); -lean_ctor_set(x_122, 4, x_121); -lean_ctor_set(x_122, 5, x_112); -lean_ctor_set(x_122, 6, x_113); -lean_ctor_set(x_122, 7, x_114); -lean_ctor_set(x_122, 8, x_115); -x_123 = lean_st_ref_set(x_8, x_122); -x_124 = lean_box(0); -if (lean_is_scalar(x_105)) { - x_125 = lean_alloc_ctor(0, 1, 0); -} else { - x_125 = x_105; -} -lean_ctor_set(x_125, 0, x_124); -return x_125; -} -} -} -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__10___redArg(lean_object* x_1) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -uint8_t x_3; -x_3 = !lean_is_exclusive(x_1); -if (x_3 == 0) -{ -lean_ctor_set_tag(x_1, 1); -return x_1; -} -else -{ -lean_object* x_4; lean_object* x_5; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); -lean_dec(x_1); -x_5 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_5, 0, x_4); -return x_5; -} -} -else -{ -uint8_t x_6; -x_6 = !lean_is_exclusive(x_1); -if (x_6 == 0) -{ -lean_ctor_set_tag(x_1, 0); -return x_1; -} -else -{ -lean_object* x_7; lean_object* x_8; -x_7 = lean_ctor_get(x_1, 0); -lean_inc(x_7); -lean_dec(x_1); -x_8 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_8, 0, x_7); -return x_8; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__8___redArg(lean_object* x_1) { -_start: -{ -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_3 = lean_st_ref_get(x_1); -x_4 = lean_ctor_get(x_3, 4); -lean_inc_ref(x_4); -lean_dec(x_3); -x_5 = lean_ctor_get(x_4, 0); -lean_inc_ref(x_5); -lean_dec_ref(x_4); -x_6 = lean_st_ref_take(x_1); -x_7 = !lean_is_exclusive(x_6); -if (x_7 == 0) -{ -lean_object* x_8; uint8_t x_9; -x_8 = lean_ctor_get(x_6, 4); -x_9 = !lean_is_exclusive(x_8); -if (x_9 == 0) -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = lean_ctor_get(x_8, 0); -lean_dec(x_10); -x_11 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13_spec__33___redArg___closed__2; -lean_ctor_set(x_8, 0, x_11); -x_12 = lean_st_ref_set(x_1, x_6); -x_13 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_13, 0, x_5); -return x_13; -} -else -{ -uint64_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_14 = lean_ctor_get_uint64(x_8, sizeof(void*)*1); -lean_dec(x_8); -x_15 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13_spec__33___redArg___closed__2; -x_16 = lean_alloc_ctor(0, 1, 8); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set_uint64(x_16, sizeof(void*)*1, x_14); -lean_ctor_set(x_6, 4, x_16); -x_17 = lean_st_ref_set(x_1, x_6); -x_18 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_18, 0, x_5); -return x_18; -} -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint64_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_19 = lean_ctor_get(x_6, 4); -x_20 = lean_ctor_get(x_6, 0); -x_21 = lean_ctor_get(x_6, 1); -x_22 = lean_ctor_get(x_6, 2); -x_23 = lean_ctor_get(x_6, 3); -x_24 = lean_ctor_get(x_6, 5); -x_25 = lean_ctor_get(x_6, 6); -x_26 = lean_ctor_get(x_6, 7); -x_27 = lean_ctor_get(x_6, 8); -lean_inc(x_27); -lean_inc(x_26); -lean_inc(x_25); -lean_inc(x_24); -lean_inc(x_19); -lean_inc(x_23); -lean_inc(x_22); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_6); -x_28 = lean_ctor_get_uint64(x_19, sizeof(void*)*1); -if (lean_is_exclusive(x_19)) { - lean_ctor_release(x_19, 0); - x_29 = x_19; -} else { - lean_dec_ref(x_19); - x_29 = lean_box(0); -} -x_30 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13_spec__33___redArg___closed__2; -if (lean_is_scalar(x_29)) { - x_31 = lean_alloc_ctor(0, 1, 8); -} else { - x_31 = x_29; -} -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set_uint64(x_31, sizeof(void*)*1, x_28); -x_32 = lean_alloc_ctor(0, 9, 0); -lean_ctor_set(x_32, 0, x_20); -lean_ctor_set(x_32, 1, x_21); -lean_ctor_set(x_32, 2, x_22); -lean_ctor_set(x_32, 3, x_23); -lean_ctor_set(x_32, 4, x_31); -lean_ctor_set(x_32, 5, x_24); -lean_ctor_set(x_32, 6, x_25); -lean_ctor_set(x_32, 7, x_26); -lean_ctor_set(x_32, 8, x_27); -x_33 = lean_st_ref_set(x_1, x_32); -x_34 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_34, 0, x_5); -return x_34; -} -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; double x_28; lean_object* x_29; uint8_t x_30; double x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; double x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; double x_43; lean_object* x_44; 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; double x_64; double x_65; lean_object* x_66; uint8_t x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; double x_75; double x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; lean_object* x_80; lean_object* x_85; lean_object* x_86; double x_87; lean_object* x_88; lean_object* x_89; uint8_t x_90; double x_91; lean_object* x_92; uint8_t x_93; double x_127; lean_object* x_128; lean_object* x_129; uint8_t x_130; double x_131; lean_object* x_132; double x_133; lean_object* x_137; uint8_t x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; double x_159; double x_160; lean_object* x_161; lean_object* x_162; uint8_t x_163; lean_object* x_164; uint8_t x_165; double x_199; double x_200; lean_object* x_201; lean_object* x_202; uint8_t x_203; lean_object* x_204; double x_205; lean_object* x_209; lean_object* x_210; uint8_t x_211; lean_object* x_212; lean_object* x_213; uint8_t x_249; -x_11 = lean_ctor_get(x_8, 2); -x_12 = lean_ctor_get(x_8, 5); -lean_inc(x_1); -x_85 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__0___redArg(x_1, x_8); -x_86 = lean_ctor_get(x_85, 0); -lean_inc(x_86); -lean_dec_ref(x_85); -x_249 = lean_unbox(x_86); -if (x_249 == 0) -{ -lean_object* x_250; uint8_t x_251; -x_250 = l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13___redArg___closed__4; -x_251 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13_spec__34(x_11, x_250); -if (x_251 == 0) -{ -lean_object* x_252; -lean_dec(x_86); -lean_dec_ref(x_5); -lean_dec_ref(x_2); -lean_dec(x_1); -x_252 = lean_apply_5(x_3, x_6, x_7, x_8, x_9, lean_box(0)); -return x_252; -} -else -{ -lean_inc(x_12); -goto block_248; -} -} -else -{ -lean_inc(x_12); -goto block_248; -} -block_27: -{ -lean_object* x_22; -x_22 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__9(x_13, x_16, x_12, x_14, x_17, x_18, x_19, x_20); -lean_dec(x_20); -lean_dec(x_18); -lean_dec_ref(x_17); -if (lean_obj_tag(x_22) == 0) -{ -lean_object* x_23; -lean_dec_ref(x_22); -x_23 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__10___redArg(x_15); -return x_23; -} -else -{ -uint8_t x_24; -lean_dec_ref(x_15); -x_24 = !lean_is_exclusive(x_22); -if (x_24 == 0) -{ -return x_22; -} -else -{ -lean_object* x_25; lean_object* x_26; -x_25 = lean_ctor_get(x_22, 0); -lean_inc(x_25); -lean_dec(x_22); -x_26 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_26, 0, x_25); -return x_26; -} -} -} -block_38: -{ -if (x_30 == 0) -{ -double x_35; lean_object* x_36; -x_35 = l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13___redArg___closed__0; -x_36 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_36, 0, x_1); -lean_ctor_set(x_36, 1, x_5); -lean_ctor_set_float(x_36, sizeof(void*)*2, x_35); -lean_ctor_set_float(x_36, sizeof(void*)*2 + 8, x_35); -lean_ctor_set_uint8(x_36, sizeof(void*)*2 + 16, x_4); -x_13 = x_29; -x_14 = x_33; -x_15 = x_32; -x_16 = x_36; -x_17 = x_6; -x_18 = x_7; -x_19 = x_8; -x_20 = x_9; -x_21 = lean_box(0); -goto block_27; -} -else -{ -lean_object* x_37; -x_37 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_37, 0, x_1); -lean_ctor_set(x_37, 1, x_5); -lean_ctor_set_float(x_37, sizeof(void*)*2, x_31); -lean_ctor_set_float(x_37, sizeof(void*)*2 + 8, x_28); -lean_ctor_set_uint8(x_37, sizeof(void*)*2 + 16, x_4); -x_13 = x_29; -x_14 = x_33; -x_15 = x_32; -x_16 = x_37; -x_17 = x_6; -x_18 = x_7; -x_19 = x_8; -x_20 = x_9; -x_21 = lean_box(0); -goto block_27; -} -} -block_48: -{ -lean_object* x_45; -lean_inc(x_9); -lean_inc_ref(x_8); -lean_inc(x_7); -lean_inc_ref(x_6); -lean_inc_ref(x_44); -x_45 = lean_apply_6(x_2, x_44, x_6, x_7, x_8, x_9, lean_box(0)); -if (lean_obj_tag(x_45) == 0) -{ -lean_object* x_46; -x_46 = lean_ctor_get(x_45, 0); -lean_inc(x_46); -lean_dec_ref(x_45); -x_28 = x_39; -x_29 = x_40; -x_30 = x_42; -x_31 = x_43; -x_32 = x_44; -x_33 = x_46; -x_34 = lean_box(0); -goto block_38; -} -else -{ -lean_object* x_47; -lean_dec_ref(x_45); -x_47 = l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13___redArg___closed__2; -x_28 = x_39; -x_29 = x_40; -x_30 = x_42; -x_31 = x_43; -x_32 = x_44; -x_33 = x_47; -x_34 = lean_box(0); -goto block_38; -} -} -block_63: -{ -lean_object* x_58; -x_58 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__9(x_49, x_52, x_12, x_51, x_53, x_54, x_55, x_56); -lean_dec(x_56); -lean_dec(x_54); -lean_dec_ref(x_53); -if (lean_obj_tag(x_58) == 0) -{ -lean_object* x_59; -lean_dec_ref(x_58); -x_59 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__10___redArg(x_50); -return x_59; -} -else -{ -uint8_t x_60; -lean_dec_ref(x_50); -x_60 = !lean_is_exclusive(x_58); -if (x_60 == 0) -{ -return x_58; -} -else -{ -lean_object* x_61; lean_object* x_62; -x_61 = lean_ctor_get(x_58, 0); -lean_inc(x_61); -lean_dec(x_58); -x_62 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_62, 0, x_61); -return x_62; -} -} -} -block_74: -{ -if (x_67 == 0) -{ -double x_71; lean_object* x_72; -x_71 = l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13___redArg___closed__0; -x_72 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_72, 0, x_1); -lean_ctor_set(x_72, 1, x_5); -lean_ctor_set_float(x_72, sizeof(void*)*2, x_71); -lean_ctor_set_float(x_72, sizeof(void*)*2 + 8, x_71); -lean_ctor_set_uint8(x_72, sizeof(void*)*2 + 16, x_4); -x_49 = x_66; -x_50 = x_68; -x_51 = x_69; -x_52 = x_72; -x_53 = x_6; -x_54 = x_7; -x_55 = x_8; -x_56 = x_9; -x_57 = lean_box(0); -goto block_63; -} -else -{ -lean_object* x_73; -x_73 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_73, 0, x_1); -lean_ctor_set(x_73, 1, x_5); -lean_ctor_set_float(x_73, sizeof(void*)*2, x_65); -lean_ctor_set_float(x_73, sizeof(void*)*2 + 8, x_64); -lean_ctor_set_uint8(x_73, sizeof(void*)*2 + 16, x_4); -x_49 = x_66; -x_50 = x_68; -x_51 = x_69; -x_52 = x_73; -x_53 = x_6; -x_54 = x_7; -x_55 = x_8; -x_56 = x_9; -x_57 = lean_box(0); -goto block_63; -} -} -block_84: -{ -lean_object* x_81; -lean_inc(x_9); -lean_inc_ref(x_8); -lean_inc(x_7); -lean_inc_ref(x_6); -lean_inc_ref(x_80); -x_81 = lean_apply_6(x_2, x_80, x_6, x_7, x_8, x_9, lean_box(0)); -if (lean_obj_tag(x_81) == 0) -{ -lean_object* x_82; -x_82 = lean_ctor_get(x_81, 0); -lean_inc(x_82); -lean_dec_ref(x_81); -x_64 = x_75; -x_65 = x_76; -x_66 = x_78; -x_67 = x_79; -x_68 = x_80; -x_69 = x_82; -x_70 = lean_box(0); -goto block_74; -} -else -{ -lean_object* x_83; -lean_dec_ref(x_81); -x_83 = l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13___redArg___closed__2; -x_64 = x_75; -x_65 = x_76; -x_66 = x_78; -x_67 = x_79; -x_68 = x_80; -x_69 = x_83; -x_70 = lean_box(0); -goto block_74; -} -} -block_126: -{ -uint8_t x_94; -x_94 = lean_unbox(x_86); -lean_dec(x_86); -if (x_94 == 0) -{ -if (x_93 == 0) -{ -lean_object* x_95; uint8_t x_96; -lean_dec(x_12); -lean_dec_ref(x_8); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec_ref(x_5); -lean_dec_ref(x_2); -lean_dec(x_1); -x_95 = lean_st_ref_take(x_9); -x_96 = !lean_is_exclusive(x_95); -if (x_96 == 0) -{ -lean_object* x_97; uint8_t x_98; -x_97 = lean_ctor_get(x_95, 4); -x_98 = !lean_is_exclusive(x_97); -if (x_98 == 0) -{ -lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; -x_99 = lean_ctor_get(x_97, 0); -x_100 = l_Lean_PersistentArray_append___redArg(x_88, x_99); -lean_dec_ref(x_99); -lean_ctor_set(x_97, 0, x_100); -x_101 = lean_st_ref_set(x_9, x_95); -lean_dec(x_9); -x_102 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__10___redArg(x_92); -return x_102; -} -else -{ -uint64_t x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_103 = lean_ctor_get_uint64(x_97, sizeof(void*)*1); -x_104 = lean_ctor_get(x_97, 0); -lean_inc(x_104); -lean_dec(x_97); -x_105 = l_Lean_PersistentArray_append___redArg(x_88, x_104); -lean_dec_ref(x_104); -x_106 = lean_alloc_ctor(0, 1, 8); -lean_ctor_set(x_106, 0, x_105); -lean_ctor_set_uint64(x_106, sizeof(void*)*1, x_103); -lean_ctor_set(x_95, 4, x_106); -x_107 = lean_st_ref_set(x_9, x_95); -lean_dec(x_9); -x_108 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__10___redArg(x_92); -return x_108; -} -} -else -{ -lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; uint64_t x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; -x_109 = lean_ctor_get(x_95, 4); -x_110 = lean_ctor_get(x_95, 0); -x_111 = lean_ctor_get(x_95, 1); -x_112 = lean_ctor_get(x_95, 2); -x_113 = lean_ctor_get(x_95, 3); -x_114 = lean_ctor_get(x_95, 5); -x_115 = lean_ctor_get(x_95, 6); -x_116 = lean_ctor_get(x_95, 7); -x_117 = lean_ctor_get(x_95, 8); -lean_inc(x_117); -lean_inc(x_116); -lean_inc(x_115); -lean_inc(x_114); -lean_inc(x_109); -lean_inc(x_113); -lean_inc(x_112); -lean_inc(x_111); -lean_inc(x_110); -lean_dec(x_95); -x_118 = lean_ctor_get_uint64(x_109, sizeof(void*)*1); -x_119 = lean_ctor_get(x_109, 0); -lean_inc_ref(x_119); -if (lean_is_exclusive(x_109)) { - lean_ctor_release(x_109, 0); - x_120 = x_109; -} else { - lean_dec_ref(x_109); - x_120 = lean_box(0); -} -x_121 = l_Lean_PersistentArray_append___redArg(x_88, x_119); -lean_dec_ref(x_119); -if (lean_is_scalar(x_120)) { - x_122 = lean_alloc_ctor(0, 1, 8); -} else { - x_122 = x_120; -} -lean_ctor_set(x_122, 0, x_121); -lean_ctor_set_uint64(x_122, sizeof(void*)*1, x_118); -x_123 = lean_alloc_ctor(0, 9, 0); -lean_ctor_set(x_123, 0, x_110); -lean_ctor_set(x_123, 1, x_111); -lean_ctor_set(x_123, 2, x_112); -lean_ctor_set(x_123, 3, x_113); -lean_ctor_set(x_123, 4, x_122); -lean_ctor_set(x_123, 5, x_114); -lean_ctor_set(x_123, 6, x_115); -lean_ctor_set(x_123, 7, x_116); -lean_ctor_set(x_123, 8, x_117); -x_124 = lean_st_ref_set(x_9, x_123); -lean_dec(x_9); -x_125 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__10___redArg(x_92); -return x_125; -} -} -else -{ -x_39 = x_87; -x_40 = x_88; -x_41 = lean_box(0); -x_42 = x_90; -x_43 = x_91; -x_44 = x_92; -goto block_48; -} -} -else -{ -x_39 = x_87; -x_40 = x_88; -x_41 = lean_box(0); -x_42 = x_90; -x_43 = x_91; -x_44 = x_92; -goto block_48; -} -} -block_136: -{ -double x_134; uint8_t x_135; -x_134 = lean_float_sub(x_127, x_131); -x_135 = lean_float_decLt(x_133, x_134); -x_87 = x_127; -x_88 = x_128; -x_89 = lean_box(0); -x_90 = x_130; -x_91 = x_131; -x_92 = x_132; -x_93 = x_135; -goto block_126; -} -block_158: -{ -lean_object* x_142; double x_143; double x_144; double x_145; double x_146; double x_147; lean_object* x_148; uint8_t x_149; -x_142 = lean_io_mono_nanos_now(); -x_143 = lean_float_of_nat(x_139); -x_144 = l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13___redArg___closed__3; -x_145 = lean_float_div(x_143, x_144); -x_146 = lean_float_of_nat(x_142); -x_147 = lean_float_div(x_146, x_144); -x_148 = l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13___redArg___closed__4; -x_149 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13_spec__34(x_11, x_148); -if (x_149 == 0) -{ -x_87 = x_147; -x_88 = x_137; -x_89 = lean_box(0); -x_90 = x_149; -x_91 = x_145; -x_92 = x_140; -x_93 = x_149; -goto block_126; -} -else -{ -if (x_138 == 0) -{ -lean_object* x_150; lean_object* x_151; double x_152; double x_153; double x_154; -x_150 = l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13___redArg___closed__5; -x_151 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13_spec__37(x_11, x_150); -x_152 = lean_float_of_nat(x_151); -x_153 = l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13___redArg___closed__6; -x_154 = lean_float_div(x_152, x_153); -x_127 = x_147; -x_128 = x_137; -x_129 = lean_box(0); -x_130 = x_149; -x_131 = x_145; -x_132 = x_140; -x_133 = x_154; -goto block_136; -} -else -{ -lean_object* x_155; lean_object* x_156; double x_157; -x_155 = l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13___redArg___closed__5; -x_156 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13_spec__37(x_11, x_155); -x_157 = lean_float_of_nat(x_156); -x_127 = x_147; -x_128 = x_137; -x_129 = lean_box(0); -x_130 = x_149; -x_131 = x_145; -x_132 = x_140; -x_133 = x_157; -goto block_136; -} -} -} -block_198: -{ -uint8_t x_166; -x_166 = lean_unbox(x_86); -lean_dec(x_86); -if (x_166 == 0) -{ -if (x_165 == 0) -{ -lean_object* x_167; uint8_t x_168; -lean_dec(x_12); -lean_dec_ref(x_8); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec_ref(x_5); -lean_dec_ref(x_2); -lean_dec(x_1); -x_167 = lean_st_ref_take(x_9); -x_168 = !lean_is_exclusive(x_167); -if (x_168 == 0) -{ -lean_object* x_169; uint8_t x_170; -x_169 = lean_ctor_get(x_167, 4); -x_170 = !lean_is_exclusive(x_169); -if (x_170 == 0) -{ -lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; -x_171 = lean_ctor_get(x_169, 0); -x_172 = l_Lean_PersistentArray_append___redArg(x_161, x_171); -lean_dec_ref(x_171); -lean_ctor_set(x_169, 0, x_172); -x_173 = lean_st_ref_set(x_9, x_167); -lean_dec(x_9); -x_174 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__10___redArg(x_164); -return x_174; -} -else -{ -uint64_t x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; -x_175 = lean_ctor_get_uint64(x_169, sizeof(void*)*1); -x_176 = lean_ctor_get(x_169, 0); -lean_inc(x_176); -lean_dec(x_169); -x_177 = l_Lean_PersistentArray_append___redArg(x_161, x_176); -lean_dec_ref(x_176); -x_178 = lean_alloc_ctor(0, 1, 8); -lean_ctor_set(x_178, 0, x_177); -lean_ctor_set_uint64(x_178, sizeof(void*)*1, x_175); -lean_ctor_set(x_167, 4, x_178); -x_179 = lean_st_ref_set(x_9, x_167); -lean_dec(x_9); -x_180 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__10___redArg(x_164); -return x_180; -} -} -else -{ -lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; uint64_t x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; -x_181 = lean_ctor_get(x_167, 4); -x_182 = lean_ctor_get(x_167, 0); -x_183 = lean_ctor_get(x_167, 1); -x_184 = lean_ctor_get(x_167, 2); -x_185 = lean_ctor_get(x_167, 3); -x_186 = lean_ctor_get(x_167, 5); -x_187 = lean_ctor_get(x_167, 6); -x_188 = lean_ctor_get(x_167, 7); -x_189 = lean_ctor_get(x_167, 8); -lean_inc(x_189); -lean_inc(x_188); -lean_inc(x_187); -lean_inc(x_186); -lean_inc(x_181); -lean_inc(x_185); -lean_inc(x_184); -lean_inc(x_183); -lean_inc(x_182); -lean_dec(x_167); -x_190 = lean_ctor_get_uint64(x_181, sizeof(void*)*1); -x_191 = lean_ctor_get(x_181, 0); -lean_inc_ref(x_191); -if (lean_is_exclusive(x_181)) { - lean_ctor_release(x_181, 0); - x_192 = x_181; -} else { - lean_dec_ref(x_181); - x_192 = lean_box(0); -} -x_193 = l_Lean_PersistentArray_append___redArg(x_161, x_191); -lean_dec_ref(x_191); -if (lean_is_scalar(x_192)) { - x_194 = lean_alloc_ctor(0, 1, 8); -} else { - x_194 = x_192; -} -lean_ctor_set(x_194, 0, x_193); -lean_ctor_set_uint64(x_194, sizeof(void*)*1, x_190); -x_195 = lean_alloc_ctor(0, 9, 0); -lean_ctor_set(x_195, 0, x_182); -lean_ctor_set(x_195, 1, x_183); -lean_ctor_set(x_195, 2, x_184); -lean_ctor_set(x_195, 3, x_185); -lean_ctor_set(x_195, 4, x_194); -lean_ctor_set(x_195, 5, x_186); -lean_ctor_set(x_195, 6, x_187); -lean_ctor_set(x_195, 7, x_188); -lean_ctor_set(x_195, 8, x_189); -x_196 = lean_st_ref_set(x_9, x_195); -lean_dec(x_9); -x_197 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__10___redArg(x_164); -return x_197; -} -} -else -{ -x_75 = x_159; -x_76 = x_160; -x_77 = lean_box(0); -x_78 = x_161; -x_79 = x_163; -x_80 = x_164; -goto block_84; -} -} -else -{ -x_75 = x_159; -x_76 = x_160; -x_77 = lean_box(0); -x_78 = x_161; -x_79 = x_163; -x_80 = x_164; -goto block_84; -} -} -block_208: -{ -double x_206; uint8_t x_207; -x_206 = lean_float_sub(x_199, x_200); -x_207 = lean_float_decLt(x_205, x_206); -x_159 = x_199; -x_160 = x_200; -x_161 = x_202; -x_162 = lean_box(0); -x_163 = x_203; -x_164 = x_204; -x_165 = x_207; -goto block_198; -} -block_227: -{ -lean_object* x_214; double x_215; double x_216; lean_object* x_217; uint8_t x_218; -x_214 = lean_io_get_num_heartbeats(); -x_215 = lean_float_of_nat(x_209); -x_216 = lean_float_of_nat(x_214); -x_217 = l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13___redArg___closed__4; -x_218 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13_spec__34(x_11, x_217); -if (x_218 == 0) -{ -x_159 = x_216; -x_160 = x_215; -x_161 = x_210; -x_162 = lean_box(0); -x_163 = x_218; -x_164 = x_212; -x_165 = x_218; -goto block_198; -} -else -{ -if (x_211 == 0) -{ -lean_object* x_219; lean_object* x_220; double x_221; double x_222; double x_223; -x_219 = l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13___redArg___closed__5; -x_220 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13_spec__37(x_11, x_219); -x_221 = lean_float_of_nat(x_220); -x_222 = l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13___redArg___closed__6; -x_223 = lean_float_div(x_221, x_222); -x_199 = x_216; -x_200 = x_215; -x_201 = lean_box(0); -x_202 = x_210; -x_203 = x_218; -x_204 = x_212; -x_205 = x_223; -goto block_208; -} -else -{ -lean_object* x_224; lean_object* x_225; double x_226; -x_224 = l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13___redArg___closed__5; -x_225 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13_spec__37(x_11, x_224); -x_226 = lean_float_of_nat(x_225); -x_199 = x_216; -x_200 = x_215; -x_201 = lean_box(0); -x_202 = x_210; -x_203 = x_218; -x_204 = x_212; -x_205 = x_226; -goto block_208; -} -} -} -block_248: -{ -lean_object* x_228; lean_object* x_229; lean_object* x_230; uint8_t x_231; -x_228 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__8___redArg(x_9); -x_229 = lean_ctor_get(x_228, 0); -lean_inc(x_229); -lean_dec_ref(x_228); -x_230 = l_Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13___redArg___closed__7; -x_231 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13_spec__34(x_11, x_230); -if (x_231 == 0) -{ -lean_object* x_232; lean_object* x_233; -x_232 = lean_io_mono_nanos_now(); -lean_inc(x_9); -lean_inc_ref(x_8); -lean_inc(x_7); -lean_inc_ref(x_6); -x_233 = lean_apply_5(x_3, x_6, x_7, x_8, x_9, lean_box(0)); -if (lean_obj_tag(x_233) == 0) -{ -uint8_t x_234; -x_234 = !lean_is_exclusive(x_233); -if (x_234 == 0) -{ -lean_ctor_set_tag(x_233, 1); -x_137 = x_229; -x_138 = x_231; -x_139 = x_232; -x_140 = x_233; -x_141 = lean_box(0); -goto block_158; -} -else -{ -lean_object* x_235; lean_object* x_236; -x_235 = lean_ctor_get(x_233, 0); -lean_inc(x_235); -lean_dec(x_233); -x_236 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_236, 0, x_235); -x_137 = x_229; -x_138 = x_231; -x_139 = x_232; -x_140 = x_236; -x_141 = lean_box(0); -goto block_158; -} -} -else -{ -uint8_t x_237; -x_237 = !lean_is_exclusive(x_233); -if (x_237 == 0) -{ -lean_ctor_set_tag(x_233, 0); -x_137 = x_229; -x_138 = x_231; -x_139 = x_232; -x_140 = x_233; -x_141 = lean_box(0); -goto block_158; -} -else -{ -lean_object* x_238; lean_object* x_239; -x_238 = lean_ctor_get(x_233, 0); -lean_inc(x_238); -lean_dec(x_233); -x_239 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_239, 0, x_238); -x_137 = x_229; -x_138 = x_231; -x_139 = x_232; -x_140 = x_239; -x_141 = lean_box(0); -goto block_158; -} -} -} -else -{ -lean_object* x_240; lean_object* x_241; -x_240 = lean_io_get_num_heartbeats(); -lean_inc(x_9); -lean_inc_ref(x_8); -lean_inc(x_7); -lean_inc_ref(x_6); -x_241 = lean_apply_5(x_3, x_6, x_7, x_8, x_9, lean_box(0)); -if (lean_obj_tag(x_241) == 0) -{ -uint8_t x_242; -x_242 = !lean_is_exclusive(x_241); -if (x_242 == 0) -{ -lean_ctor_set_tag(x_241, 1); -x_209 = x_240; -x_210 = x_229; -x_211 = x_231; -x_212 = x_241; -x_213 = lean_box(0); -goto block_227; -} -else -{ -lean_object* x_243; lean_object* x_244; -x_243 = lean_ctor_get(x_241, 0); -lean_inc(x_243); -lean_dec(x_241); -x_244 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_244, 0, x_243); -x_209 = x_240; -x_210 = x_229; -x_211 = x_231; -x_212 = x_244; -x_213 = lean_box(0); -goto block_227; -} -} -else -{ -uint8_t x_245; -x_245 = !lean_is_exclusive(x_241); -if (x_245 == 0) -{ -lean_ctor_set_tag(x_241, 0); -x_209 = x_240; -x_210 = x_229; -x_211 = x_231; -x_212 = x_241; -x_213 = lean_box(0); -goto block_227; -} -else -{ -lean_object* x_246; lean_object* x_247; -x_246 = lean_ctor_get(x_241, 0); -lean_inc(x_246); -lean_dec(x_241); -x_247 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_247, 0, x_246); -x_209 = x_240; -x_210 = x_229; -x_211 = x_231; -x_212 = x_247; -x_213 = lean_box(0); -goto block_227; -} -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_12; -x_12 = l_Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_12; -} -} -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_8; -x_8 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__10___redArg(x_2); -return x_8; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_6; -x_6 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__8___redArg(x_4); -return x_6; -} -} -LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__2(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_9 = lean_box(0); -x_10 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_10, 0, x_1); -lean_ctor_set(x_10, 1, x_9); -lean_ctor_set_uint8(x_10, sizeof(void*)*2, x_2); -x_11 = lean_box(0); -x_12 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_12, 0, x_10); -lean_ctor_set(x_12, 1, x_11); -x_13 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_13, 0, x_12); -return x_13; -} -} -LEAN_EXPORT lean_object* l_List_mapTR_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__4(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_3; -x_3 = l_List_reverse___redArg(x_2); -return x_3; -} -else -{ -uint8_t x_4; -x_4 = !lean_is_exclusive(x_1); -if (x_4 == 0) -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = lean_ctor_get(x_1, 0); -x_6 = lean_ctor_get(x_1, 1); -x_7 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_7, 0, x_5); -lean_ctor_set(x_1, 1, x_2); -lean_ctor_set(x_1, 0, x_7); -{ -lean_object* _tmp_0 = x_6; -lean_object* _tmp_1 = x_1; -x_1 = _tmp_0; -x_2 = _tmp_1; -} -goto _start; -} -else -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_ctor_get(x_1, 0); -x_10 = lean_ctor_get(x_1, 1); -lean_inc(x_10); -lean_inc(x_9); -lean_dec(x_1); -x_11 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_11, 0, x_9); -x_12 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_2); -x_1 = x_10; -x_2 = x_12; -goto _start; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Tactic_FunInd_rwMatcher_spec__2(size_t x_1, size_t x_2, lean_object* x_3) { -_start: -{ -uint8_t x_4; -x_4 = lean_usize_dec_lt(x_2, x_1); -if (x_4 == 0) -{ -return x_3; -} -else -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; 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 = l_Lean_Expr_mvarId_x21(x_5); -lean_dec(x_5); -x_9 = 1; -x_10 = lean_usize_add(x_2, x_9); -x_11 = lean_array_uset(x_7, x_2, x_8); -x_2 = x_10; -x_3 = x_11; -goto _start; -} -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___lam__0___closed__0() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked(" rewriting with ", 16, 16); -return x_1; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___lam__0___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Tactic_FunInd_rwMatcher___lam__0___closed__0; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___lam__0___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked(" in", 3, 3); -return x_1; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___lam__0___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Tactic_FunInd_rwMatcher___lam__0___closed__2; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__0(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_10 = l_Lean_exceptEmoji___redArg(x_4); -x_11 = l_Lean_stringToMessageData(x_10); -x_12 = l_Lean_Tactic_FunInd_rwMatcher___lam__0___closed__1; -x_13 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_13, 0, x_11); -lean_ctor_set(x_13, 1, x_12); -x_14 = l_Lean_MessageData_ofConstName(x_1, x_2); -x_15 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_15, 0, x_13); -lean_ctor_set(x_15, 1, x_14); -x_16 = l_Lean_Tactic_FunInd_rwMatcher___lam__0___closed__3; -x_17 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_16); -x_18 = l_Lean_indentExpr(x_3); -x_19 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_19, 1, x_18); -x_20 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_20, 0, x_19); -return x_20; -} -} -LEAN_EXPORT lean_object* l___private_Init_While_0__Lean_Loop_forIn_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_7; -lean_inc(x_5); -lean_inc_ref(x_4); -lean_inc(x_3); -lean_inc_ref(x_2); -lean_inc_ref(x_1); -x_7 = l_Lean_Meta_reduceRecMatcher_x3f(x_1, x_2, x_3, x_4, x_5); -if (lean_obj_tag(x_7) == 0) -{ -uint8_t x_8; -x_8 = !lean_is_exclusive(x_7); -if (x_8 == 0) -{ -lean_object* x_9; -x_9 = lean_ctor_get(x_7, 0); -if (lean_obj_tag(x_9) == 1) -{ -lean_object* x_10; lean_object* x_11; -lean_free_object(x_7); -lean_dec_ref(x_1); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -lean_dec_ref(x_9); -x_11 = l_Lean_Expr_headBeta(x_10); -x_1 = x_11; -goto _start; -} -else -{ -lean_object* x_13; uint8_t x_14; -lean_dec(x_9); -lean_inc_ref(x_1); -x_13 = l_Lean_Expr_headBeta(x_1); -x_14 = lean_expr_eqv(x_1, x_13); -if (x_14 == 0) -{ -lean_free_object(x_7); -lean_dec_ref(x_1); -x_1 = x_13; -goto _start; -} -else -{ -lean_dec_ref(x_13); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -lean_ctor_set(x_7, 0, x_1); -return x_7; -} -} -} -else -{ -lean_object* x_16; -x_16 = lean_ctor_get(x_7, 0); -lean_inc(x_16); -lean_dec(x_7); -if (lean_obj_tag(x_16) == 1) -{ -lean_object* x_17; lean_object* x_18; -lean_dec_ref(x_1); -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec_ref(x_16); -x_18 = l_Lean_Expr_headBeta(x_17); -x_1 = x_18; -goto _start; -} -else -{ -lean_object* x_20; uint8_t x_21; -lean_dec(x_16); -lean_inc_ref(x_1); -x_20 = l_Lean_Expr_headBeta(x_1); -x_21 = lean_expr_eqv(x_1, x_20); -if (x_21 == 0) -{ -lean_dec_ref(x_1); -x_1 = x_20; -goto _start; -} -else -{ -lean_object* x_23; -lean_dec_ref(x_20); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -x_23 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_23, 0, x_1); -return x_23; -} -} -} -} -else -{ -uint8_t x_24; -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_24 = !lean_is_exclusive(x_7); -if (x_24 == 0) -{ -return x_7; -} -else -{ -lean_object* x_25; lean_object* x_26; -x_25 = lean_ctor_get(x_7, 0); -lean_inc(x_25); -lean_dec(x_7); -x_26 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_26, 0, x_25); -return x_26; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Tactic_FunInd_rwMatcher_spec__6(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_10; lean_object* x_11; uint8_t x_16; -x_16 = lean_usize_dec_eq(x_2, x_3); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_21; -x_17 = lean_array_uget(x_1, x_2); -x_21 = l_Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0___redArg(x_17, x_6); -if (lean_obj_tag(x_21) == 0) -{ -lean_object* x_22; uint8_t x_23; -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -lean_dec_ref(x_21); -x_23 = lean_unbox(x_22); -lean_dec(x_22); -if (x_23 == 0) -{ -x_18 = lean_box(0); -goto block_20; -} -else -{ -lean_dec(x_17); -x_10 = x_4; -x_11 = lean_box(0); -goto block_15; -} -} -else -{ -if (lean_obj_tag(x_21) == 0) -{ -lean_object* x_24; uint8_t x_25; -x_24 = lean_ctor_get(x_21, 0); -lean_inc(x_24); -lean_dec_ref(x_21); -x_25 = lean_unbox(x_24); -lean_dec(x_24); -if (x_25 == 0) -{ -lean_dec(x_17); -x_10 = x_4; -x_11 = lean_box(0); -goto block_15; -} -else -{ -x_18 = lean_box(0); -goto block_20; -} -} -else -{ -uint8_t x_26; -lean_dec(x_17); -lean_dec_ref(x_4); -x_26 = !lean_is_exclusive(x_21); -if (x_26 == 0) -{ -return x_21; -} -else -{ -lean_object* x_27; lean_object* x_28; -x_27 = lean_ctor_get(x_21, 0); -lean_inc(x_27); -lean_dec(x_21); -x_28 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_28, 0, x_27); -return x_28; -} -} -} -block_20: -{ -lean_object* x_19; -x_19 = lean_array_push(x_4, x_17); -x_10 = x_19; -x_11 = lean_box(0); -goto block_15; -} -} -else -{ -lean_object* x_29; -x_29 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_29, 0, x_4); -return x_29; -} -block_15: -{ -size_t x_12; size_t x_13; -x_12 = 1; -x_13 = lean_usize_add(x_2, x_12); -x_2 = x_13; -x_4 = x_10; -goto _start; -} -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___closed__0() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Failed to apply ", 16, 16); -return x_1; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Tactic_FunInd_rwMatcher___closed__0; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked(":", 1, 1); -return x_1; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Tactic_FunInd_rwMatcher___closed__2; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Not a matcher application:", 26, 26); -return x_1; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Tactic_FunInd_rwMatcher___closed__4; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___closed__6() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("When trying to reduce arm ", 26, 26); -return x_1; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___closed__7() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Tactic_FunInd_rwMatcher___closed__6; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___closed__8() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked(", only ", 7, 7); -return x_1; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___closed__9() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Tactic_FunInd_rwMatcher___closed__8; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___closed__10() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked(" equations for ", 15, 15); -return x_1; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___closed__11() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Tactic_FunInd_rwMatcher___closed__10; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___closed__12() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("PSum", 4, 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___closed__13() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("casesOn", 7, 7); -return x_1; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___closed__14() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Tactic_FunInd_rwMatcher___closed__13; -x_2 = l_Lean_Tactic_FunInd_rwMatcher___closed__12; -x_3 = l_Lean_Name_mkStr2(x_2, x_1); -return x_3; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___closed__15() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("PSigma", 6, 6); -return x_1; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___closed__16() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Tactic_FunInd_rwMatcher___closed__13; -x_2 = l_Lean_Tactic_FunInd_rwMatcher___closed__15; -x_3 = l_Lean_Name_mkStr2(x_2, x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -uint8_t x_10; lean_object* x_11; -x_10 = lean_unbox(x_2); -x_11 = l_Lean_Tactic_FunInd_rwMatcher___lam__0(x_1, x_10, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec_ref(x_4); -return x_11; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__0() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Could not un-HEq `", 18, 18); -return x_1; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__0; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("`:", 2, 2); -return x_1; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__2; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked(" ", 1, 1); -return x_1; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__4; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__6() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Not all hypotheses of `", 23, 23); -return x_1; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__7() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__6; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__8() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("` could be discharged: ", 23, 23); -return x_1; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__9() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__8; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5___closed__0() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Failed to resolve `", 19, 19); -return x_1; -} -} -static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5___closed__0; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Failed to discharge `", 21, 21); -return x_1; -} -} -static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5___closed__2; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5(lean_object* x_1, uint8_t x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_12; lean_object* x_13; lean_object* x_18; lean_object* x_20; lean_object* x_22; uint8_t x_24; -x_24 = lean_usize_dec_lt(x_5, x_4); -if (x_24 == 0) -{ -lean_object* x_25; -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec(x_8); -lean_dec_ref(x_7); -x_25 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_25, 0, x_6); -return x_25; -} -else -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; lean_object* x_65; -x_26 = lean_array_uget(x_3, x_5); -x_65 = l_Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0___redArg(x_26, x_8); -if (lean_obj_tag(x_65) == 0) -{ -lean_object* x_66; uint8_t x_67; -x_66 = lean_ctor_get(x_65, 0); -lean_inc(x_66); -lean_dec_ref(x_65); -x_67 = lean_unbox(x_66); -lean_dec(x_66); -if (x_67 == 0) -{ -lean_object* x_68; -lean_inc(x_26); -x_68 = l_Lean_MVarId_getType(x_26, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_68) == 0) -{ -lean_object* x_69; uint8_t x_70; -x_69 = lean_ctor_get(x_68, 0); -lean_inc(x_69); -lean_dec_ref(x_68); -lean_inc(x_69); -x_70 = l_Lean_Meta_Simp_isEqnThmHypothesis(x_69); -if (x_70 == 0) -{ -uint8_t x_71; -x_71 = l_Lean_Expr_isEq(x_69); -if (x_71 == 0) -{ -uint8_t x_72; -x_72 = l_Lean_Expr_isHEq(x_69); -lean_dec(x_69); -if (x_72 == 0) -{ -lean_dec(x_26); -x_12 = x_1; -x_13 = lean_box(0); -goto block_17; -} -else -{ -lean_object* x_73; -x_73 = l_Lean_Meta_saveState___redArg(x_8, x_10); -if (lean_obj_tag(x_73) == 0) -{ -lean_object* x_74; lean_object* x_75; -x_74 = lean_ctor_get(x_73, 0); -lean_inc(x_74); -lean_dec_ref(x_73); -lean_inc(x_10); -lean_inc_ref(x_9); -lean_inc(x_8); -lean_inc_ref(x_7); -lean_inc(x_26); -x_75 = l_Lean_MVarId_assumption(x_26, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_75) == 0) -{ -lean_dec(x_74); -lean_dec(x_26); -x_20 = x_75; -goto block_21; -} -else -{ -lean_object* x_76; uint8_t x_77; uint8_t x_89; -x_76 = lean_ctor_get(x_75, 0); -lean_inc(x_76); -x_89 = l_Lean_Exception_isInterrupt(x_76); -if (x_89 == 0) -{ -uint8_t x_90; -x_90 = l_Lean_Exception_isRuntime(x_76); -x_77 = x_90; -goto block_88; -} -else -{ -lean_dec(x_76); -x_77 = x_89; -goto block_88; -} -block_88: -{ -if (x_77 == 0) -{ -lean_object* x_78; -lean_dec_ref(x_75); -x_78 = l_Lean_Meta_SavedState_restore___redArg(x_74, x_8, x_10); -lean_dec(x_74); -if (lean_obj_tag(x_78) == 0) -{ -lean_object* x_79; -lean_dec_ref(x_78); -x_79 = l_Lean_Meta_saveState___redArg(x_8, x_10); -if (lean_obj_tag(x_79) == 0) -{ -lean_object* x_80; lean_object* x_81; -x_80 = lean_ctor_get(x_79, 0); -lean_inc(x_80); -lean_dec_ref(x_79); -lean_inc(x_10); -lean_inc_ref(x_9); -lean_inc(x_8); -lean_inc_ref(x_7); -lean_inc(x_26); -x_81 = l_Lean_MVarId_hrefl(x_26, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_81) == 0) -{ -lean_dec(x_80); -lean_dec(x_26); -x_20 = x_81; -goto block_21; -} -else -{ -lean_object* x_82; uint8_t x_83; -x_82 = lean_ctor_get(x_81, 0); -lean_inc(x_82); -x_83 = l_Lean_Exception_isInterrupt(x_82); -if (x_83 == 0) -{ -uint8_t x_84; -x_84 = l_Lean_Exception_isRuntime(x_82); -x_46 = x_80; -x_47 = lean_box(0); -x_48 = x_81; -x_49 = x_84; -goto block_64; -} -else -{ -lean_dec(x_82); -x_46 = x_80; -x_47 = lean_box(0); -x_48 = x_81; -x_49 = x_83; -goto block_64; -} -} -} -else -{ -uint8_t x_85; -lean_dec(x_26); -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec(x_8); -lean_dec_ref(x_7); -x_85 = !lean_is_exclusive(x_79); -if (x_85 == 0) -{ -return x_79; -} -else -{ -lean_object* x_86; lean_object* x_87; -x_86 = lean_ctor_get(x_79, 0); -lean_inc(x_86); -lean_dec(x_79); -x_87 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_87, 0, x_86); -return x_87; -} -} -} -else -{ -lean_dec(x_26); -x_20 = x_78; -goto block_21; -} -} -else -{ -lean_dec(x_74); -lean_dec(x_26); -x_20 = x_75; -goto block_21; -} -} -} -} -else -{ -uint8_t x_91; -lean_dec(x_26); -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec(x_8); -lean_dec_ref(x_7); -x_91 = !lean_is_exclusive(x_73); -if (x_91 == 0) -{ -return x_73; -} -else -{ -lean_object* x_92; lean_object* x_93; -x_92 = lean_ctor_get(x_73, 0); -lean_inc(x_92); -lean_dec(x_73); -x_93 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_93, 0, x_92); -return x_93; -} -} -} -} -else -{ -lean_object* x_94; -lean_dec(x_69); -x_94 = l_Lean_Meta_saveState___redArg(x_8, x_10); -if (lean_obj_tag(x_94) == 0) -{ -lean_object* x_95; lean_object* x_96; -x_95 = lean_ctor_get(x_94, 0); -lean_inc(x_95); -lean_dec_ref(x_94); -lean_inc(x_10); -lean_inc_ref(x_9); -lean_inc(x_8); -lean_inc_ref(x_7); -lean_inc(x_26); -x_96 = l_Lean_MVarId_assumption(x_26, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_96) == 0) -{ -lean_dec(x_95); -lean_dec(x_26); -x_18 = x_96; -goto block_19; -} -else -{ -lean_object* x_97; uint8_t x_98; uint8_t x_110; -x_97 = lean_ctor_get(x_96, 0); -lean_inc(x_97); -x_110 = l_Lean_Exception_isInterrupt(x_97); -if (x_110 == 0) -{ -uint8_t x_111; -x_111 = l_Lean_Exception_isRuntime(x_97); -x_98 = x_111; -goto block_109; -} -else -{ -lean_dec(x_97); -x_98 = x_110; -goto block_109; -} -block_109: -{ -if (x_98 == 0) -{ -lean_object* x_99; -lean_dec_ref(x_96); -x_99 = l_Lean_Meta_SavedState_restore___redArg(x_95, x_8, x_10); -lean_dec(x_95); -if (lean_obj_tag(x_99) == 0) -{ -lean_object* x_100; -lean_dec_ref(x_99); -x_100 = l_Lean_Meta_saveState___redArg(x_8, x_10); -if (lean_obj_tag(x_100) == 0) -{ -lean_object* x_101; lean_object* x_102; -x_101 = lean_ctor_get(x_100, 0); -lean_inc(x_101); -lean_dec_ref(x_100); -lean_inc(x_10); -lean_inc_ref(x_9); -lean_inc(x_8); -lean_inc_ref(x_7); -lean_inc(x_26); -x_102 = l_Lean_MVarId_refl(x_26, x_2, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_102) == 0) -{ -lean_dec(x_101); -lean_dec(x_26); -x_18 = x_102; -goto block_19; -} -else -{ -lean_object* x_103; uint8_t x_104; -x_103 = lean_ctor_get(x_102, 0); -lean_inc(x_103); -x_104 = l_Lean_Exception_isInterrupt(x_103); -if (x_104 == 0) -{ -uint8_t x_105; -x_105 = l_Lean_Exception_isRuntime(x_103); -x_27 = x_101; -x_28 = lean_box(0); -x_29 = x_102; -x_30 = x_105; -goto block_45; -} -else -{ -lean_dec(x_103); -x_27 = x_101; -x_28 = lean_box(0); -x_29 = x_102; -x_30 = x_104; -goto block_45; -} -} -} -else -{ -uint8_t x_106; -lean_dec(x_26); -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec(x_8); -lean_dec_ref(x_7); -x_106 = !lean_is_exclusive(x_100); -if (x_106 == 0) -{ -return x_100; -} -else -{ -lean_object* x_107; lean_object* x_108; -x_107 = lean_ctor_get(x_100, 0); -lean_inc(x_107); -lean_dec(x_100); -x_108 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_108, 0, x_107); -return x_108; -} -} -} -else -{ -lean_dec(x_26); -x_18 = x_99; -goto block_19; -} -} -else -{ -lean_dec(x_95); -lean_dec(x_26); -x_18 = x_96; -goto block_19; -} -} -} -} -else -{ -uint8_t x_112; -lean_dec(x_26); -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec(x_8); -lean_dec_ref(x_7); -x_112 = !lean_is_exclusive(x_94); -if (x_112 == 0) -{ -return x_94; -} -else -{ -lean_object* x_113; lean_object* x_114; -x_113 = lean_ctor_get(x_94, 0); -lean_inc(x_113); -lean_dec(x_94); -x_114 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_114, 0, x_113); -return x_114; -} -} -} -} -else -{ -lean_object* x_115; -lean_dec(x_69); -x_115 = l_Lean_Meta_saveState___redArg(x_8, x_10); -if (lean_obj_tag(x_115) == 0) -{ -lean_object* x_116; lean_object* x_117; -x_116 = lean_ctor_get(x_115, 0); -lean_inc(x_116); -lean_dec_ref(x_115); -lean_inc(x_10); -lean_inc_ref(x_9); -lean_inc(x_8); -lean_inc_ref(x_7); -lean_inc(x_26); -x_117 = l_Lean_MVarId_assumption(x_26, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_117) == 0) -{ -lean_dec(x_116); -lean_dec(x_26); -x_22 = x_117; -goto block_23; -} -else -{ -lean_object* x_118; uint8_t x_119; uint8_t x_135; -x_118 = lean_ctor_get(x_117, 0); -lean_inc(x_118); -x_135 = l_Lean_Exception_isInterrupt(x_118); -if (x_135 == 0) -{ -uint8_t x_136; -x_136 = l_Lean_Exception_isRuntime(x_118); -x_119 = x_136; -goto block_134; -} -else -{ -lean_dec(x_118); -x_119 = x_135; -goto block_134; -} -block_134: -{ -if (x_119 == 0) -{ -lean_object* x_120; -lean_dec_ref(x_117); -x_120 = l_Lean_Meta_SavedState_restore___redArg(x_116, x_8, x_10); -lean_dec(x_116); -if (lean_obj_tag(x_120) == 0) -{ -uint8_t x_121; -x_121 = !lean_is_exclusive(x_120); -if (x_121 == 0) -{ -lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; -x_122 = lean_ctor_get(x_120, 0); -lean_dec(x_122); -x_123 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5___closed__3; -lean_ctor_set_tag(x_120, 1); -lean_ctor_set(x_120, 0, x_26); -x_124 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_124, 0, x_123); -lean_ctor_set(x_124, 1, x_120); -x_125 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___redArg___lam__3___closed__3; -x_126 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_126, 0, x_124); -lean_ctor_set(x_126, 1, x_125); -x_127 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___redArg(x_126, x_7, x_8, x_9, x_10); -x_22 = x_127; -goto block_23; -} -else -{ -lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; -lean_dec(x_120); -x_128 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5___closed__3; -x_129 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_129, 0, x_26); -x_130 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_130, 0, x_128); -lean_ctor_set(x_130, 1, x_129); -x_131 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___redArg___lam__3___closed__3; -x_132 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_132, 0, x_130); -lean_ctor_set(x_132, 1, x_131); -x_133 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___redArg(x_132, x_7, x_8, x_9, x_10); -x_22 = x_133; -goto block_23; -} -} -else -{ -lean_dec(x_26); -x_22 = x_120; -goto block_23; -} -} -else -{ -lean_dec(x_116); -lean_dec(x_26); -x_22 = x_117; -goto block_23; -} -} -} -} -else -{ -uint8_t x_137; -lean_dec(x_26); -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec(x_8); -lean_dec_ref(x_7); -x_137 = !lean_is_exclusive(x_115); -if (x_137 == 0) -{ -return x_115; -} -else -{ -lean_object* x_138; lean_object* x_139; -x_138 = lean_ctor_get(x_115, 0); -lean_inc(x_138); -lean_dec(x_115); -x_139 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_139, 0, x_138); -return x_139; -} -} -} -} -else -{ -uint8_t x_140; -lean_dec(x_26); -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec(x_8); -lean_dec_ref(x_7); -x_140 = !lean_is_exclusive(x_68); -if (x_140 == 0) -{ -return x_68; -} -else -{ -lean_object* x_141; lean_object* x_142; -x_141 = lean_ctor_get(x_68, 0); -lean_inc(x_141); -lean_dec(x_68); -x_142 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_142, 0, x_141); -return x_142; -} -} -} -else -{ -lean_dec(x_26); -x_12 = x_1; -x_13 = lean_box(0); -goto block_17; -} -} -else -{ -uint8_t x_143; -lean_dec(x_26); -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec(x_8); -lean_dec_ref(x_7); -x_143 = !lean_is_exclusive(x_65); -if (x_143 == 0) -{ -return x_65; -} -else -{ -lean_object* x_144; lean_object* x_145; -x_144 = lean_ctor_get(x_65, 0); -lean_inc(x_144); -lean_dec(x_65); -x_145 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_145, 0, x_144); -return x_145; -} -} -block_45: -{ -if (x_30 == 0) -{ -lean_object* x_31; -lean_dec_ref(x_29); -x_31 = l_Lean_Meta_SavedState_restore___redArg(x_27, x_8, x_10); -lean_dec_ref(x_27); -if (lean_obj_tag(x_31) == 0) -{ -uint8_t x_32; -x_32 = !lean_is_exclusive(x_31); -if (x_32 == 0) -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_33 = lean_ctor_get(x_31, 0); -lean_dec(x_33); -x_34 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5___closed__1; -lean_ctor_set_tag(x_31, 1); -lean_ctor_set(x_31, 0, x_26); -x_35 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_31); -x_36 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___redArg___lam__3___closed__3; -x_37 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_37, 0, x_35); -lean_ctor_set(x_37, 1, x_36); -x_38 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___redArg(x_37, x_7, x_8, x_9, x_10); -x_18 = x_38; -goto block_19; -} -else -{ -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -lean_dec(x_31); -x_39 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5___closed__1; -x_40 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_40, 0, x_26); -x_41 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); -x_42 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___redArg___lam__3___closed__3; -x_43 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_43, 0, x_41); -lean_ctor_set(x_43, 1, x_42); -x_44 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___redArg(x_43, x_7, x_8, x_9, x_10); -x_18 = x_44; -goto block_19; -} -} -else -{ -lean_dec(x_26); -x_18 = x_31; -goto block_19; -} -} -else -{ -lean_dec_ref(x_27); -lean_dec(x_26); -x_18 = x_29; -goto block_19; -} -} -block_64: -{ -if (x_49 == 0) -{ -lean_object* x_50; -lean_dec_ref(x_48); -x_50 = l_Lean_Meta_SavedState_restore___redArg(x_46, x_8, x_10); -lean_dec_ref(x_46); -if (lean_obj_tag(x_50) == 0) -{ -uint8_t x_51; -x_51 = !lean_is_exclusive(x_50); -if (x_51 == 0) -{ -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_52 = lean_ctor_get(x_50, 0); -lean_dec(x_52); -x_53 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5___closed__1; -lean_ctor_set_tag(x_50, 1); -lean_ctor_set(x_50, 0, x_26); -x_54 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_50); -x_55 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___redArg___lam__3___closed__3; -x_56 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_56, 0, x_54); -lean_ctor_set(x_56, 1, x_55); -x_57 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___redArg(x_56, x_7, x_8, x_9, x_10); -x_20 = x_57; -goto block_21; -} -else -{ -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; -lean_dec(x_50); -x_58 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5___closed__1; -x_59 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_59, 0, x_26); -x_60 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_60, 0, x_58); -lean_ctor_set(x_60, 1, x_59); -x_61 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___redArg___lam__3___closed__3; -x_62 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_62, 0, x_60); -lean_ctor_set(x_62, 1, x_61); -x_63 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___redArg(x_62, x_7, x_8, x_9, x_10); -x_20 = x_63; -goto block_21; -} -} -else -{ -lean_dec(x_26); -x_20 = x_50; -goto block_21; -} -} -else -{ -lean_dec_ref(x_46); -lean_dec(x_26); -x_20 = x_48; -goto block_21; -} -} -} -block_17: -{ -size_t x_14; size_t x_15; -x_14 = 1; -x_15 = lean_usize_add(x_5, x_14); -x_5 = x_15; -x_6 = x_12; -goto _start; -} -block_19: -{ -if (lean_obj_tag(x_18) == 0) -{ -lean_dec_ref(x_18); -x_12 = x_1; -x_13 = lean_box(0); -goto block_17; -} -else -{ -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec(x_8); -lean_dec_ref(x_7); -return x_18; -} -} -block_21: -{ -if (lean_obj_tag(x_20) == 0) -{ -lean_dec_ref(x_20); -x_12 = x_1; -x_13 = lean_box(0); -goto block_17; -} -else -{ -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec(x_8); -lean_dec_ref(x_7); -return x_20; -} -} -block_23: -{ -if (lean_obj_tag(x_22) == 0) -{ -lean_dec_ref(x_22); -x_12 = x_1; -x_13 = lean_box(0); -goto block_17; -} -else -{ -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec(x_8); -lean_dec_ref(x_7); -return x_22; -} -} -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__10() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Left-hand side `", 16, 16); -return x_1; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__11() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__10; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__12() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("` of `", 6, 6); -return x_1; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__13() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__12; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__14() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("` does not apply to `", 21, 21); -return x_1; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__15() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__14; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__16() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("HEq", 3, 3); -return x_1; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__17() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__16; -x_2 = l_Lean_Name_mkStr1(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__18() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Type of `", 9, 9); -return x_1; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__19() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__18; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__20() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("` is not an equality", 20, 20); -return x_1; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__21() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__20; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__22() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("eqProof has type", 16, 16); -return x_1; -} -} -static lean_object* _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__23() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__22; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_19; lean_object* x_20; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; lean_object* x_48; uint8_t x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; uint8_t x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; uint8_t x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; size_t x_104; lean_object* x_105; uint8_t x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; size_t x_130; lean_object* x_131; lean_object* x_132; uint8_t x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_165; -lean_inc(x_10); -lean_inc_ref(x_9); -lean_inc(x_8); -lean_inc_ref(x_7); -lean_inc_ref(x_1); -x_165 = lean_infer_type(x_1, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_165) == 0) -{ -lean_object* x_166; uint8_t x_167; lean_object* x_168; -x_166 = lean_ctor_get(x_165, 0); -lean_inc(x_166); -lean_dec_ref(x_165); -x_167 = 0; -lean_inc(x_10); -lean_inc_ref(x_9); -lean_inc(x_8); -lean_inc_ref(x_7); -x_168 = l_Lean_Meta_forallMetaTelescope(x_166, x_167, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_168) == 0) -{ -lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_207; lean_object* x_208; uint8_t x_209; -x_169 = lean_ctor_get(x_168, 0); -lean_inc(x_169); -lean_dec_ref(x_168); -x_170 = lean_ctor_get(x_169, 1); -lean_inc(x_170); -x_171 = lean_ctor_get(x_169, 0); -lean_inc(x_171); -if (lean_is_exclusive(x_169)) { - lean_ctor_release(x_169, 0); - lean_ctor_release(x_169, 1); - x_172 = x_169; -} else { - lean_dec_ref(x_169); - x_172 = lean_box(0); -} -x_173 = lean_ctor_get(x_170, 1); -lean_inc(x_173); -if (lean_is_exclusive(x_170)) { - lean_ctor_release(x_170, 0); - lean_ctor_release(x_170, 1); - x_174 = x_170; -} else { - lean_dec_ref(x_170); - x_174 = lean_box(0); -} -lean_inc(x_2); -x_207 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__0___redArg(x_2, x_9); -x_208 = lean_ctor_get(x_207, 0); -lean_inc(x_208); -lean_dec_ref(x_207); -x_209 = lean_unbox(x_208); -lean_dec(x_208); -if (x_209 == 0) -{ -lean_dec(x_2); -x_175 = x_7; -x_176 = x_8; -x_177 = x_9; -x_178 = x_10; -x_179 = lean_box(0); -goto block_206; -} -else -{ -lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; -x_210 = l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__23; -lean_inc(x_173); -x_211 = l_Lean_indentExpr(x_173); -x_212 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_212, 0, x_210); -lean_ctor_set(x_212, 1, x_211); -x_213 = l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__1(x_2, x_212, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_213) == 0) -{ -lean_dec_ref(x_213); -x_175 = x_7; -x_176 = x_8; -x_177 = x_9; -x_178 = x_10; -x_179 = lean_box(0); -goto block_206; -} -else -{ -uint8_t x_214; -lean_dec(x_174); -lean_dec(x_173); -lean_dec(x_172); -lean_dec(x_171); -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_1); -x_214 = !lean_is_exclusive(x_213); -if (x_214 == 0) -{ -return x_213; -} -else -{ -lean_object* x_215; lean_object* x_216; -x_215 = lean_ctor_get(x_213, 0); -lean_inc(x_215); -lean_dec(x_213); -x_216 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_216, 0, x_215); -return x_216; -} -} -} -block_206: -{ -lean_object* x_180; size_t x_181; size_t x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; uint8_t x_186; -x_180 = l_Lean_mkAppN(x_1, x_171); -x_181 = lean_array_size(x_171); -x_182 = 0; -x_183 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Tactic_FunInd_rwMatcher_spec__2(x_181, x_182, x_171); -x_184 = l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__17; -x_185 = lean_unsigned_to_nat(4u); -x_186 = l_Lean_Expr_isAppOfArity(x_173, x_184, x_185); -if (x_186 == 0) -{ -lean_object* x_187; lean_object* x_188; uint8_t x_189; -x_187 = l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__1; -x_188 = lean_unsigned_to_nat(3u); -x_189 = l_Lean_Expr_isAppOfArity(x_173, x_187, x_188); -if (x_189 == 0) -{ -lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; uint8_t x_196; -lean_dec_ref(x_183); -lean_dec_ref(x_180); -lean_dec(x_173); -lean_dec_ref(x_5); -x_190 = l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__19; -x_191 = l_Lean_MessageData_ofConstName(x_4, x_189); -if (lean_is_scalar(x_174)) { - x_192 = lean_alloc_ctor(7, 2, 0); -} else { - x_192 = x_174; - lean_ctor_set_tag(x_192, 7); -} -lean_ctor_set(x_192, 0, x_190); -lean_ctor_set(x_192, 1, x_191); -x_193 = l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__21; -if (lean_is_scalar(x_172)) { - x_194 = lean_alloc_ctor(7, 2, 0); -} else { - x_194 = x_172; - lean_ctor_set_tag(x_194, 7); -} -lean_ctor_set(x_194, 0, x_192); -lean_ctor_set(x_194, 1, x_193); -x_195 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___redArg(x_194, x_175, x_176, x_177, x_178); -lean_dec(x_178); -lean_dec_ref(x_177); -lean_dec(x_176); -lean_dec_ref(x_175); -x_196 = !lean_is_exclusive(x_195); -if (x_196 == 0) -{ -return x_195; -} -else -{ -lean_object* x_197; lean_object* x_198; -x_197 = lean_ctor_get(x_195, 0); -lean_inc(x_197); -lean_dec(x_195); -x_198 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_198, 0, x_197); -return x_198; -} -} -else -{ -lean_object* x_199; lean_object* x_200; lean_object* x_201; -lean_dec(x_174); -lean_dec(x_172); -x_199 = l_Lean_Expr_appFn_x21(x_173); -x_200 = l_Lean_Expr_appArg_x21(x_199); -lean_dec_ref(x_199); -x_201 = l_Lean_Expr_appArg_x21(x_173); -lean_dec(x_173); -x_130 = x_182; -x_131 = x_183; -x_132 = x_180; -x_133 = x_186; -x_134 = x_200; -x_135 = x_201; -x_136 = x_175; -x_137 = x_176; -x_138 = x_177; -x_139 = x_178; -x_140 = lean_box(0); -goto block_164; -} -} -else -{ -lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; -lean_dec(x_174); -lean_dec(x_172); -x_202 = l_Lean_Expr_appFn_x21(x_173); -x_203 = l_Lean_Expr_appFn_x21(x_202); -lean_dec_ref(x_202); -x_204 = l_Lean_Expr_appArg_x21(x_203); -lean_dec_ref(x_203); -x_205 = l_Lean_Expr_appArg_x21(x_173); -lean_dec(x_173); -x_130 = x_182; -x_131 = x_183; -x_132 = x_180; -x_133 = x_3; -x_134 = x_204; -x_135 = x_205; -x_136 = x_175; -x_137 = x_176; -x_138 = x_177; -x_139 = x_178; -x_140 = lean_box(0); -goto block_164; -} -} -} -else -{ -uint8_t x_217; -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec(x_2); -lean_dec_ref(x_1); -x_217 = !lean_is_exclusive(x_168); -if (x_217 == 0) -{ -return x_168; -} -else -{ -lean_object* x_218; lean_object* x_219; -x_218 = lean_ctor_get(x_168, 0); -lean_inc(x_218); -lean_dec(x_168); -x_219 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_219, 0, x_218); -return x_219; -} -} -} -else -{ -uint8_t x_220; -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec(x_2); -lean_dec_ref(x_1); -x_220 = !lean_is_exclusive(x_165); -if (x_220 == 0) -{ -return x_165; -} -else -{ -lean_object* x_221; lean_object* x_222; -x_221 = lean_ctor_get(x_165, 0); -lean_inc(x_221); -lean_dec(x_165); -x_222 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_222, 0, x_221); -return x_222; -} -} -block_18: -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_15, 0, x_13); -x_16 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_16, 0, x_12); -lean_ctor_set(x_16, 1, x_15); -lean_ctor_set_uint8(x_16, sizeof(void*)*2, x_3); -x_17 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_17, 0, x_16); -return x_17; -} -block_25: -{ -if (lean_obj_tag(x_20) == 0) -{ -lean_object* x_21; -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -lean_dec_ref(x_20); -x_12 = x_19; -x_13 = x_21; -x_14 = lean_box(0); -goto block_18; -} -else -{ -uint8_t x_22; -lean_dec_ref(x_19); -x_22 = !lean_is_exclusive(x_20); -if (x_22 == 0) -{ -return x_20; -} -else -{ -lean_object* x_23; lean_object* x_24; -x_23 = lean_ctor_get(x_20, 0); -lean_inc(x_23); -lean_dec(x_20); -x_24 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_24, 0, x_23); -return x_24; -} -} -} -block_47: -{ -if (x_35 == 0) -{ -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -lean_dec_ref(x_26); -x_36 = l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__1; -x_37 = l_Lean_MessageData_ofExpr(x_34); -x_38 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -x_39 = l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__3; -x_40 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -x_41 = l_Lean_Exception_toMessageData(x_27); -x_42 = l_Lean_indentD(x_41); -x_43 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_43, 0, x_40); -lean_ctor_set(x_43, 1, x_42); -x_44 = l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__5; -x_45 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_45, 0, x_43); -lean_ctor_set(x_45, 1, x_44); -x_46 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___redArg(x_45, x_30, x_28, x_29, x_32); -lean_dec(x_32); -lean_dec_ref(x_29); -lean_dec(x_28); -lean_dec_ref(x_30); -x_19 = x_33; -x_20 = x_46; -goto block_25; -} -else -{ -lean_dec_ref(x_34); -lean_dec(x_32); -lean_dec_ref(x_30); -lean_dec_ref(x_29); -lean_dec(x_28); -lean_dec_ref(x_27); -x_19 = x_33; -x_20 = x_26; -goto block_25; -} -} -block_65: -{ -lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_56 = l_Lean_instantiateMVars___at___00Lean_Tactic_FunInd_rwMatcher_spec__3___redArg(x_48, x_52); -x_57 = lean_ctor_get(x_56, 0); -lean_inc(x_57); -lean_dec_ref(x_56); -x_58 = l_Lean_instantiateMVars___at___00Lean_Tactic_FunInd_rwMatcher_spec__3___redArg(x_50, x_52); -if (x_49 == 0) -{ -lean_object* x_59; -lean_dec(x_54); -lean_dec_ref(x_53); -lean_dec(x_52); -lean_dec_ref(x_51); -x_59 = lean_ctor_get(x_58, 0); -lean_inc(x_59); -lean_dec_ref(x_58); -x_12 = x_57; -x_13 = x_59; -x_14 = lean_box(0); -goto block_18; -} -else -{ -lean_object* x_60; lean_object* x_61; -x_60 = lean_ctor_get(x_58, 0); -lean_inc(x_60); -lean_dec_ref(x_58); -lean_inc(x_54); -lean_inc_ref(x_53); -lean_inc(x_52); -lean_inc_ref(x_51); -lean_inc(x_60); -x_61 = l_Lean_Meta_mkEqOfHEq(x_60, x_3, x_51, x_52, x_53, x_54); -if (lean_obj_tag(x_61) == 0) -{ -lean_dec(x_60); -lean_dec(x_54); -lean_dec_ref(x_53); -lean_dec(x_52); -lean_dec_ref(x_51); -x_19 = x_57; -x_20 = x_61; -goto block_25; -} -else -{ -lean_object* x_62; uint8_t x_63; -x_62 = lean_ctor_get(x_61, 0); -lean_inc(x_62); -x_63 = l_Lean_Exception_isInterrupt(x_62); -if (x_63 == 0) -{ -uint8_t x_64; -lean_inc(x_62); -x_64 = l_Lean_Exception_isRuntime(x_62); -x_26 = x_61; -x_27 = x_62; -x_28 = x_52; -x_29 = x_53; -x_30 = x_51; -x_31 = lean_box(0); -x_32 = x_54; -x_33 = x_57; -x_34 = x_60; -x_35 = x_64; -goto block_47; -} -else -{ -x_26 = x_61; -x_27 = x_62; -x_28 = x_52; -x_29 = x_53; -x_30 = x_51; -x_31 = lean_box(0); -x_32 = x_54; -x_33 = x_57; -x_34 = x_60; -x_35 = x_63; -goto block_47; -} -} -} -} -block_90: -{ -uint8_t x_75; -x_75 = l_Array_isEmpty___redArg(x_73); -if (x_75 == 0) -{ -lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; uint8_t x_87; -lean_dec_ref(x_72); -lean_dec_ref(x_68); -x_76 = l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__7; -x_77 = l_Lean_MessageData_ofConstName(x_4, x_75); -x_78 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_78, 0, x_76); -lean_ctor_set(x_78, 1, x_77); -x_79 = l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__9; -x_80 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_80, 0, x_78); -lean_ctor_set(x_80, 1, x_79); -x_81 = lean_array_to_list(x_73); -x_82 = lean_box(0); -x_83 = l_List_mapTR_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__4(x_81, x_82); -x_84 = l_Lean_MessageData_ofList(x_83); -x_85 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_85, 0, x_80); -lean_ctor_set(x_85, 1, x_84); -x_86 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___redArg(x_85, x_66, x_69, x_67, x_71); -lean_dec(x_71); -lean_dec_ref(x_67); -lean_dec(x_69); -lean_dec_ref(x_66); -x_87 = !lean_is_exclusive(x_86); -if (x_87 == 0) -{ -return x_86; -} -else -{ -lean_object* x_88; lean_object* x_89; -x_88 = lean_ctor_get(x_86, 0); -lean_inc(x_88); -lean_dec(x_86); -x_89 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_89, 0, x_88); -return x_89; -} -} -else -{ -lean_dec_ref(x_73); -lean_dec(x_4); -x_48 = x_68; -x_49 = x_70; -x_50 = x_72; -x_51 = x_66; -x_52 = x_69; -x_53 = x_67; -x_54 = x_71; -x_55 = lean_box(0); -goto block_65; -} -} -block_103: -{ -if (lean_obj_tag(x_98) == 0) -{ -lean_object* x_99; -x_99 = lean_ctor_get(x_98, 0); -lean_inc(x_99); -lean_dec_ref(x_98); -x_66 = x_91; -x_67 = x_92; -x_68 = x_93; -x_69 = x_94; -x_70 = x_95; -x_71 = x_96; -x_72 = x_97; -x_73 = x_99; -x_74 = lean_box(0); -goto block_90; -} -else -{ -uint8_t x_100; -lean_dec_ref(x_97); -lean_dec(x_96); -lean_dec(x_94); -lean_dec_ref(x_93); -lean_dec_ref(x_92); -lean_dec_ref(x_91); -lean_dec(x_4); -x_100 = !lean_is_exclusive(x_98); -if (x_100 == 0) -{ -return x_98; -} -else -{ -lean_object* x_101; lean_object* x_102; -x_101 = lean_ctor_get(x_98, 0); -lean_inc(x_101); -lean_dec(x_98); -x_102 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_102, 0, x_101); -return x_102; -} -} -} -block_129: -{ -lean_object* x_114; size_t x_115; lean_object* x_116; -x_114 = lean_box(0); -x_115 = lean_array_size(x_107); -lean_inc(x_112); -lean_inc_ref(x_111); -lean_inc(x_110); -lean_inc_ref(x_109); -x_116 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5(x_114, x_3, x_107, x_115, x_104, x_114, x_109, x_110, x_111, x_112); -if (lean_obj_tag(x_116) == 0) -{ -lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; -lean_dec_ref(x_116); -x_117 = lean_unsigned_to_nat(0u); -x_118 = lean_array_get_size(x_107); -x_119 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M2_run___redArg___closed__0; -x_120 = lean_nat_dec_lt(x_117, x_118); -if (x_120 == 0) -{ -lean_dec_ref(x_107); -x_66 = x_109; -x_67 = x_111; -x_68 = x_105; -x_69 = x_110; -x_70 = x_106; -x_71 = x_112; -x_72 = x_108; -x_73 = x_119; -x_74 = lean_box(0); -goto block_90; -} -else -{ -uint8_t x_121; -x_121 = lean_nat_dec_le(x_118, x_118); -if (x_121 == 0) -{ -if (x_120 == 0) -{ -lean_dec_ref(x_107); -x_66 = x_109; -x_67 = x_111; -x_68 = x_105; -x_69 = x_110; -x_70 = x_106; -x_71 = x_112; -x_72 = x_108; -x_73 = x_119; -x_74 = lean_box(0); -goto block_90; -} -else -{ -size_t x_122; lean_object* x_123; -x_122 = lean_usize_of_nat(x_118); -x_123 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Tactic_FunInd_rwMatcher_spec__6(x_107, x_104, x_122, x_119, x_109, x_110, x_111, x_112); -lean_dec_ref(x_107); -x_91 = x_109; -x_92 = x_111; -x_93 = x_105; -x_94 = x_110; -x_95 = x_106; -x_96 = x_112; -x_97 = x_108; -x_98 = x_123; -goto block_103; -} -} -else -{ -size_t x_124; lean_object* x_125; -x_124 = lean_usize_of_nat(x_118); -x_125 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Tactic_FunInd_rwMatcher_spec__6(x_107, x_104, x_124, x_119, x_109, x_110, x_111, x_112); -lean_dec_ref(x_107); -x_91 = x_109; -x_92 = x_111; -x_93 = x_105; -x_94 = x_110; -x_95 = x_106; -x_96 = x_112; -x_97 = x_108; -x_98 = x_125; -goto block_103; -} -} -} -else -{ -uint8_t x_126; -lean_dec(x_112); -lean_dec_ref(x_111); -lean_dec(x_110); -lean_dec_ref(x_109); -lean_dec_ref(x_108); -lean_dec_ref(x_107); -lean_dec_ref(x_105); -lean_dec(x_4); -x_126 = !lean_is_exclusive(x_116); -if (x_126 == 0) -{ -return x_116; -} -else -{ -lean_object* x_127; lean_object* x_128; -x_127 = lean_ctor_get(x_116, 0); -lean_inc(x_127); -lean_dec(x_116); -x_128 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_128, 0, x_127); -return x_128; -} -} -} -block_164: -{ -lean_object* x_141; -lean_inc(x_139); -lean_inc_ref(x_138); -lean_inc(x_137); -lean_inc_ref(x_136); -lean_inc_ref(x_134); -lean_inc_ref(x_5); -x_141 = l_Lean_Meta_isExprDefEq(x_5, x_134, x_136, x_137, x_138, x_139); -if (lean_obj_tag(x_141) == 0) -{ -lean_object* x_142; uint8_t x_143; -x_142 = lean_ctor_get(x_141, 0); -lean_inc(x_142); -lean_dec_ref(x_141); -x_143 = lean_unbox(x_142); -lean_dec(x_142); -if (x_143 == 0) -{ -lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; uint8_t x_158; -lean_dec_ref(x_135); -lean_dec_ref(x_132); -lean_dec_ref(x_131); -x_144 = l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__11; -x_145 = l_Lean_MessageData_ofExpr(x_134); -x_146 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_146, 0, x_144); -lean_ctor_set(x_146, 1, x_145); -x_147 = l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__13; -x_148 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_148, 0, x_146); -lean_ctor_set(x_148, 1, x_147); -x_149 = l_Lean_MessageData_ofConstName(x_4, x_6); -x_150 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_150, 0, x_148); -lean_ctor_set(x_150, 1, x_149); -x_151 = l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__15; -x_152 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_152, 0, x_150); -lean_ctor_set(x_152, 1, x_151); -x_153 = l_Lean_MessageData_ofExpr(x_5); -x_154 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_154, 0, x_152); -lean_ctor_set(x_154, 1, x_153); -x_155 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___redArg___lam__3___closed__3; -x_156 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_156, 0, x_154); -lean_ctor_set(x_156, 1, x_155); -x_157 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___redArg(x_156, x_136, x_137, x_138, x_139); -lean_dec(x_139); -lean_dec_ref(x_138); -lean_dec(x_137); -lean_dec_ref(x_136); -x_158 = !lean_is_exclusive(x_157); -if (x_158 == 0) -{ -return x_157; -} -else -{ -lean_object* x_159; lean_object* x_160; -x_159 = lean_ctor_get(x_157, 0); -lean_inc(x_159); -lean_dec(x_157); -x_160 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_160, 0, x_159); -return x_160; -} -} -else -{ -lean_dec_ref(x_134); -lean_dec_ref(x_5); -x_104 = x_130; -x_105 = x_135; -x_106 = x_133; -x_107 = x_131; -x_108 = x_132; -x_109 = x_136; -x_110 = x_137; -x_111 = x_138; -x_112 = x_139; -x_113 = lean_box(0); -goto block_129; -} -} -else -{ -uint8_t x_161; -lean_dec(x_139); -lean_dec_ref(x_138); -lean_dec(x_137); -lean_dec_ref(x_136); -lean_dec_ref(x_135); -lean_dec_ref(x_134); -lean_dec_ref(x_132); -lean_dec_ref(x_131); -lean_dec_ref(x_5); -lean_dec(x_4); -x_161 = !lean_is_exclusive(x_141); -if (x_161 == 0) -{ -return x_141; -} -else -{ -lean_object* x_162; lean_object* x_163; -x_162 = lean_ctor_get(x_141, 0); -lean_inc(x_162); -lean_dec(x_141); -x_163 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_163, 0, x_162); -return x_163; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -uint8_t x_12; uint8_t x_13; lean_object* x_14; -x_12 = lean_unbox(x_3); -x_13 = lean_unbox(x_6); -x_14 = l_Lean_Tactic_FunInd_rwMatcher___lam__1(x_1, x_2, x_12, x_4, x_5, x_13, x_7, x_8, x_9, x_10); -return x_14; -} -} -LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_rwMatcher___lam__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -uint8_t x_9; lean_object* x_10; -x_9 = lean_unbox(x_2); -x_10 = l_Lean_Tactic_FunInd_rwMatcher___lam__2(x_1, x_9, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec(x_5); -lean_dec_ref(x_4); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_rwMatcher(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_8; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; uint8_t x_46; lean_object* x_47; uint8_t x_52; lean_object* x_53; uint8_t x_58; lean_object* x_233; uint8_t x_234; -x_233 = l_Lean_Tactic_FunInd_rwMatcher___closed__14; -x_234 = l_Lean_Expr_isAppOf(x_2, x_233); -if (x_234 == 0) -{ -lean_object* x_235; uint8_t x_236; -x_235 = l_Lean_Tactic_FunInd_rwMatcher___closed__16; -x_236 = l_Lean_Expr_isAppOf(x_2, x_235); -x_58 = x_236; -goto block_232; -} -else -{ -x_58 = x_234; -goto block_232; -} -block_18: -{ -if (lean_obj_tag(x_8) == 0) -{ -uint8_t x_9; -x_9 = !lean_is_exclusive(x_8); -if (x_9 == 0) -{ -lean_object* x_10; lean_object* x_11; -x_10 = lean_ctor_get(x_8, 0); -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -lean_dec(x_10); -lean_ctor_set(x_8, 0, x_11); -return x_8; -} -else -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_8, 0); -lean_inc(x_12); -lean_dec(x_8); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -lean_dec(x_12); -x_14 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_14, 0, x_13); -return x_14; -} -} -else -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_8); -if (x_15 == 0) -{ -return x_8; -} -else -{ -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_8, 0); -lean_inc(x_16); -lean_dec(x_8); -x_17 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_17, 0, x_16); -return x_17; -} -} -} -block_45: -{ -if (x_25 == 0) -{ -lean_object* x_26; lean_object* x_27; uint8_t x_28; -lean_dec_ref(x_24); -lean_inc(x_19); -x_26 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__0___redArg(x_19, x_5); -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -lean_dec_ref(x_26); -x_28 = lean_unbox(x_27); -lean_dec(x_27); -if (x_28 == 0) -{ -lean_object* x_29; lean_object* x_30; -lean_dec(x_22); -lean_dec_ref(x_21); -lean_dec(x_19); -x_29 = lean_box(0); -x_30 = lean_apply_6(x_20, x_29, x_3, x_4, x_5, x_6, lean_box(0)); -x_8 = x_30; -goto block_18; -} -else -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_31 = l_Lean_Tactic_FunInd_rwMatcher___closed__1; -x_32 = l_Lean_MessageData_ofConstName(x_22, x_25); -x_33 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -x_34 = l_Lean_Tactic_FunInd_rwMatcher___closed__3; -x_35 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -x_36 = l_Lean_Exception_toMessageData(x_21); -x_37 = l_Lean_indentD(x_36); -x_38 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_38, 0, x_35); -lean_ctor_set(x_38, 1, x_37); -x_39 = l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__1(x_19, x_38, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_39) == 0) -{ -lean_object* x_40; lean_object* x_41; -x_40 = lean_ctor_get(x_39, 0); -lean_inc(x_40); -lean_dec_ref(x_39); -x_41 = lean_apply_6(x_20, x_40, x_3, x_4, x_5, x_6, lean_box(0)); -x_8 = x_41; -goto block_18; -} -else -{ -uint8_t x_42; -lean_dec_ref(x_20); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -x_42 = !lean_is_exclusive(x_39); -if (x_42 == 0) -{ -return x_39; -} -else -{ -lean_object* x_43; lean_object* x_44; -x_43 = lean_ctor_get(x_39, 0); -lean_inc(x_43); -lean_dec(x_39); -x_44 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_44, 0, x_43); -return x_44; -} -} -} -} -else -{ -lean_dec(x_22); -lean_dec_ref(x_21); -lean_dec_ref(x_20); -lean_dec(x_19); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -return x_24; -} -} -block_51: -{ -lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_48 = lean_box(0); -x_49 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_49, 0, x_2); -lean_ctor_set(x_49, 1, x_48); -lean_ctor_set_uint8(x_49, sizeof(void*)*2, x_46); -x_50 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_50, 0, x_49); -return x_50; -} -block_57: -{ -lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_54 = lean_box(0); -x_55 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_55, 0, x_2); -lean_ctor_set(x_55, 1, x_54); -lean_ctor_set_uint8(x_55, sizeof(void*)*2, x_52); -x_56 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_56, 0, x_55); -return x_56; -} -block_232: -{ -uint8_t x_59; -x_59 = 1; -if (x_58 == 0) -{ -lean_object* x_60; uint8_t x_61; -x_60 = l_Lean_Meta_isMatcherApp___at___00Lean_Tactic_FunInd_rwMatcher_spec__1___redArg(x_2, x_6); -x_61 = !lean_is_exclusive(x_60); -if (x_61 == 0) -{ -lean_object* x_62; uint8_t x_63; -x_62 = lean_ctor_get(x_60, 0); -x_63 = lean_unbox(x_62); -lean_dec(x_62); -if (x_63 == 0) -{ -lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; -lean_free_object(x_60); -lean_dec(x_1); -x_64 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__21___closed__2; -x_65 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__0___redArg(x_64, x_5); -x_66 = lean_ctor_get(x_65, 0); -lean_inc(x_66); -lean_dec_ref(x_65); -x_67 = lean_unbox(x_66); -lean_dec(x_66); -if (x_67 == 0) -{ -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -x_46 = x_59; -x_47 = lean_box(0); -goto block_51; -} -else -{ -lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_68 = l_Lean_Tactic_FunInd_rwMatcher___closed__5; -lean_inc_ref(x_2); -x_69 = l_Lean_indentExpr(x_2); -x_70 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_69); -x_71 = l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__1(x_64, x_70, x_3, x_4, x_5, x_6); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -if (lean_obj_tag(x_71) == 0) -{ -lean_dec_ref(x_71); -x_46 = x_59; -x_47 = lean_box(0); -goto block_51; -} -else -{ -uint8_t x_72; -lean_dec_ref(x_2); -x_72 = !lean_is_exclusive(x_71); -if (x_72 == 0) -{ -return x_71; -} -else -{ -lean_object* x_73; lean_object* x_74; -x_73 = lean_ctor_get(x_71, 0); -lean_inc(x_73); -lean_dec(x_71); -x_74 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_74, 0, x_73); -return x_74; -} -} -} -} -else -{ -lean_object* x_75; lean_object* x_76; lean_object* x_77; -x_75 = l_Lean_Expr_getAppFn(x_2); -x_76 = l_Lean_Expr_constName_x21(x_75); -lean_inc(x_6); -lean_inc_ref(x_5); -lean_inc(x_4); -lean_inc_ref(x_3); -lean_inc(x_76); -x_77 = lean_get_congr_match_equations_for(x_76, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_77) == 0) -{ -lean_object* x_78; lean_object* x_79; uint8_t x_80; -x_78 = lean_ctor_get(x_77, 0); -lean_inc(x_78); -lean_dec_ref(x_77); -x_79 = lean_array_get_size(x_78); -x_80 = lean_nat_dec_lt(x_1, x_79); -if (x_80 == 0) -{ -lean_object* x_81; lean_object* x_82; uint8_t x_83; -lean_dec(x_78); -lean_dec_ref(x_75); -x_81 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__21___closed__2; -x_82 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__0___redArg(x_81, x_5); -x_83 = !lean_is_exclusive(x_82); -if (x_83 == 0) -{ -lean_object* x_84; uint8_t x_85; -x_84 = lean_ctor_get(x_82, 0); -x_85 = lean_unbox(x_84); -lean_dec(x_84); -if (x_85 == 0) -{ -lean_free_object(x_82); -lean_dec(x_76); -lean_free_object(x_60); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -lean_dec(x_1); -x_52 = x_59; -x_53 = lean_box(0); -goto block_57; -} -else -{ -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_86 = l_Lean_Tactic_FunInd_rwMatcher___closed__7; -x_87 = l_Nat_reprFast(x_1); -lean_ctor_set_tag(x_82, 3); -lean_ctor_set(x_82, 0, x_87); -x_88 = l_Lean_MessageData_ofFormat(x_82); -x_89 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_89, 0, x_86); -lean_ctor_set(x_89, 1, x_88); -x_90 = l_Lean_Tactic_FunInd_rwMatcher___closed__9; -x_91 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_91, 0, x_89); -lean_ctor_set(x_91, 1, x_90); -x_92 = l_Nat_reprFast(x_79); -lean_ctor_set_tag(x_60, 3); -lean_ctor_set(x_60, 0, x_92); -x_93 = l_Lean_MessageData_ofFormat(x_60); -x_94 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_94, 0, x_91); -lean_ctor_set(x_94, 1, x_93); -x_95 = l_Lean_Tactic_FunInd_rwMatcher___closed__11; -x_96 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_96, 0, x_94); -lean_ctor_set(x_96, 1, x_95); -x_97 = l_Lean_MessageData_ofConstName(x_76, x_58); -x_98 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_98, 0, x_96); -lean_ctor_set(x_98, 1, x_97); -x_99 = l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__1(x_81, x_98, x_3, x_4, x_5, x_6); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -if (lean_obj_tag(x_99) == 0) -{ -lean_dec_ref(x_99); -x_52 = x_59; -x_53 = lean_box(0); -goto block_57; -} -else -{ -uint8_t x_100; -lean_dec_ref(x_2); -x_100 = !lean_is_exclusive(x_99); -if (x_100 == 0) -{ -return x_99; -} -else -{ -lean_object* x_101; lean_object* x_102; -x_101 = lean_ctor_get(x_99, 0); -lean_inc(x_101); -lean_dec(x_99); -x_102 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_102, 0, x_101); -return x_102; -} -} -} -} -else -{ -lean_object* x_103; uint8_t x_104; -x_103 = lean_ctor_get(x_82, 0); -lean_inc(x_103); -lean_dec(x_82); -x_104 = lean_unbox(x_103); -lean_dec(x_103); -if (x_104 == 0) -{ -lean_dec(x_76); -lean_free_object(x_60); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -lean_dec(x_1); -x_52 = x_59; -x_53 = lean_box(0); -goto block_57; -} -else -{ -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; -x_105 = l_Lean_Tactic_FunInd_rwMatcher___closed__7; -x_106 = l_Nat_reprFast(x_1); -x_107 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_107, 0, x_106); -x_108 = l_Lean_MessageData_ofFormat(x_107); -x_109 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_109, 0, x_105); -lean_ctor_set(x_109, 1, x_108); -x_110 = l_Lean_Tactic_FunInd_rwMatcher___closed__9; -x_111 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_111, 0, x_109); -lean_ctor_set(x_111, 1, x_110); -x_112 = l_Nat_reprFast(x_79); -lean_ctor_set_tag(x_60, 3); -lean_ctor_set(x_60, 0, x_112); -x_113 = l_Lean_MessageData_ofFormat(x_60); -x_114 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_114, 0, x_111); -lean_ctor_set(x_114, 1, x_113); -x_115 = l_Lean_Tactic_FunInd_rwMatcher___closed__11; -x_116 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_116, 0, x_114); -lean_ctor_set(x_116, 1, x_115); -x_117 = l_Lean_MessageData_ofConstName(x_76, x_58); -x_118 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_118, 0, x_116); -lean_ctor_set(x_118, 1, x_117); -x_119 = l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__1(x_81, x_118, x_3, x_4, x_5, x_6); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -if (lean_obj_tag(x_119) == 0) -{ -lean_dec_ref(x_119); -x_52 = x_59; -x_53 = lean_box(0); -goto block_57; -} -else -{ -lean_object* x_120; lean_object* x_121; lean_object* x_122; -lean_dec_ref(x_2); -x_120 = lean_ctor_get(x_119, 0); -lean_inc(x_120); -if (lean_is_exclusive(x_119)) { - lean_ctor_release(x_119, 0); - x_121 = x_119; -} else { - lean_dec_ref(x_119); - x_121 = lean_box(0); -} -if (lean_is_scalar(x_121)) { - x_122 = lean_alloc_ctor(1, 1, 0); -} else { - x_122 = x_121; -} -lean_ctor_set(x_122, 0, x_120); -return x_122; -} -} -} -} -else -{ -lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; -lean_dec(x_76); -lean_free_object(x_60); -x_123 = lean_box(0); -x_124 = lean_array_get(x_123, x_78, x_1); -lean_dec(x_1); -lean_dec(x_78); -x_125 = lean_box(x_58); -lean_inc_ref(x_2); -lean_inc(x_124); -x_126 = lean_alloc_closure((void*)(l_Lean_Tactic_FunInd_rwMatcher___lam__0___boxed), 9, 3); -lean_closure_set(x_126, 0, x_124); -lean_closure_set(x_126, 1, x_125); -lean_closure_set(x_126, 2, x_2); -x_127 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__21___closed__2; -x_128 = l_Lean_Expr_constLevels_x21(x_75); -lean_dec_ref(x_75); -lean_inc(x_124); -x_129 = l_Lean_mkConst(x_124, x_128); -x_130 = l___private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0___lam__1___closed__0; -x_131 = l_Lean_Expr_getAppNumArgs(x_2); -lean_inc(x_131); -x_132 = lean_mk_array(x_131, x_130); -x_133 = lean_unsigned_to_nat(1u); -x_134 = lean_nat_sub(x_131, x_133); -lean_dec(x_131); -lean_inc_ref(x_2); -x_135 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_2, x_132, x_134); -x_136 = l_Lean_mkAppN(x_129, x_135); -lean_dec_ref(x_135); -x_137 = lean_box(x_59); -x_138 = lean_box(x_58); -lean_inc_ref(x_2); -lean_inc(x_124); -x_139 = lean_alloc_closure((void*)(l_Lean_Tactic_FunInd_rwMatcher___lam__1___boxed), 11, 6); -lean_closure_set(x_139, 0, x_136); -lean_closure_set(x_139, 1, x_127); -lean_closure_set(x_139, 2, x_137); -lean_closure_set(x_139, 3, x_124); -lean_closure_set(x_139, 4, x_2); -lean_closure_set(x_139, 5, x_138); -x_140 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__21___closed__3; -lean_inc(x_6); -lean_inc_ref(x_5); -lean_inc(x_4); -lean_inc_ref(x_3); -x_141 = l_Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7___redArg(x_127, x_126, x_139, x_59, x_140, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_141) == 0) -{ -lean_dec(x_124); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -lean_dec_ref(x_2); -return x_141; -} -else -{ -lean_object* x_142; lean_object* x_143; lean_object* x_144; uint8_t x_145; -x_142 = lean_ctor_get(x_141, 0); -lean_inc(x_142); -x_143 = lean_box(x_59); -x_144 = lean_alloc_closure((void*)(l_Lean_Tactic_FunInd_rwMatcher___lam__2___boxed), 8, 2); -lean_closure_set(x_144, 0, x_2); -lean_closure_set(x_144, 1, x_143); -x_145 = l_Lean_Exception_isInterrupt(x_142); -if (x_145 == 0) -{ -uint8_t x_146; -lean_inc(x_142); -x_146 = l_Lean_Exception_isRuntime(x_142); -x_19 = x_127; -x_20 = x_144; -x_21 = x_142; -x_22 = x_124; -x_23 = lean_box(0); -x_24 = x_141; -x_25 = x_146; -goto block_45; -} -else -{ -x_19 = x_127; -x_20 = x_144; -x_21 = x_142; -x_22 = x_124; -x_23 = lean_box(0); -x_24 = x_141; -x_25 = x_145; -goto block_45; -} -} -} -} -else -{ -uint8_t x_147; -lean_dec(x_76); -lean_dec_ref(x_75); -lean_free_object(x_60); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -lean_dec_ref(x_2); -lean_dec(x_1); -x_147 = !lean_is_exclusive(x_77); -if (x_147 == 0) -{ -return x_77; -} -else -{ -lean_object* x_148; lean_object* x_149; -x_148 = lean_ctor_get(x_77, 0); -lean_inc(x_148); -lean_dec(x_77); -x_149 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_149, 0, x_148); -return x_149; -} -} -} -} -else -{ -lean_object* x_150; uint8_t x_151; -x_150 = lean_ctor_get(x_60, 0); -lean_inc(x_150); -lean_dec(x_60); -x_151 = lean_unbox(x_150); -lean_dec(x_150); -if (x_151 == 0) -{ -lean_object* x_152; lean_object* x_153; lean_object* x_154; uint8_t x_155; -lean_dec(x_1); -x_152 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__21___closed__2; -x_153 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__0___redArg(x_152, x_5); -x_154 = lean_ctor_get(x_153, 0); -lean_inc(x_154); -lean_dec_ref(x_153); -x_155 = lean_unbox(x_154); -lean_dec(x_154); -if (x_155 == 0) -{ -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -x_46 = x_59; -x_47 = lean_box(0); -goto block_51; -} -else -{ -lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; -x_156 = l_Lean_Tactic_FunInd_rwMatcher___closed__5; -lean_inc_ref(x_2); -x_157 = l_Lean_indentExpr(x_2); -x_158 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_158, 0, x_156); -lean_ctor_set(x_158, 1, x_157); -x_159 = l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__1(x_152, x_158, x_3, x_4, x_5, x_6); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -if (lean_obj_tag(x_159) == 0) -{ -lean_dec_ref(x_159); -x_46 = x_59; -x_47 = lean_box(0); -goto block_51; -} -else -{ -lean_object* x_160; lean_object* x_161; lean_object* x_162; -lean_dec_ref(x_2); -x_160 = lean_ctor_get(x_159, 0); -lean_inc(x_160); -if (lean_is_exclusive(x_159)) { - lean_ctor_release(x_159, 0); - x_161 = x_159; -} else { - lean_dec_ref(x_159); - x_161 = lean_box(0); -} -if (lean_is_scalar(x_161)) { - x_162 = lean_alloc_ctor(1, 1, 0); -} else { - x_162 = x_161; -} -lean_ctor_set(x_162, 0, x_160); -return x_162; -} -} -} -else -{ -lean_object* x_163; lean_object* x_164; lean_object* x_165; -x_163 = l_Lean_Expr_getAppFn(x_2); -x_164 = l_Lean_Expr_constName_x21(x_163); -lean_inc(x_6); -lean_inc_ref(x_5); -lean_inc(x_4); -lean_inc_ref(x_3); -lean_inc(x_164); -x_165 = lean_get_congr_match_equations_for(x_164, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_165) == 0) -{ -lean_object* x_166; lean_object* x_167; uint8_t x_168; -x_166 = lean_ctor_get(x_165, 0); -lean_inc(x_166); -lean_dec_ref(x_165); -x_167 = lean_array_get_size(x_166); -x_168 = lean_nat_dec_lt(x_1, x_167); -if (x_168 == 0) -{ -lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; uint8_t x_173; -lean_dec(x_166); -lean_dec_ref(x_163); -x_169 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__21___closed__2; -x_170 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__0___redArg(x_169, x_5); -x_171 = lean_ctor_get(x_170, 0); -lean_inc(x_171); -if (lean_is_exclusive(x_170)) { - lean_ctor_release(x_170, 0); - x_172 = x_170; -} else { - lean_dec_ref(x_170); - x_172 = lean_box(0); -} -x_173 = lean_unbox(x_171); -lean_dec(x_171); -if (x_173 == 0) -{ -lean_dec(x_172); -lean_dec(x_164); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -lean_dec(x_1); -x_52 = x_59; -x_53 = lean_box(0); -goto block_57; -} -else -{ -lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; -x_174 = l_Lean_Tactic_FunInd_rwMatcher___closed__7; -x_175 = l_Nat_reprFast(x_1); -if (lean_is_scalar(x_172)) { - x_176 = lean_alloc_ctor(3, 1, 0); -} else { - x_176 = x_172; - lean_ctor_set_tag(x_176, 3); -} -lean_ctor_set(x_176, 0, x_175); -x_177 = l_Lean_MessageData_ofFormat(x_176); -x_178 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_178, 0, x_174); -lean_ctor_set(x_178, 1, x_177); -x_179 = l_Lean_Tactic_FunInd_rwMatcher___closed__9; -x_180 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_180, 0, x_178); -lean_ctor_set(x_180, 1, x_179); -x_181 = l_Nat_reprFast(x_167); -x_182 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_182, 0, x_181); -x_183 = l_Lean_MessageData_ofFormat(x_182); -x_184 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_184, 0, x_180); -lean_ctor_set(x_184, 1, x_183); -x_185 = l_Lean_Tactic_FunInd_rwMatcher___closed__11; -x_186 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_186, 0, x_184); -lean_ctor_set(x_186, 1, x_185); -x_187 = l_Lean_MessageData_ofConstName(x_164, x_58); -x_188 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_188, 0, x_186); -lean_ctor_set(x_188, 1, x_187); -x_189 = l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith_spec__1(x_169, x_188, x_3, x_4, x_5, x_6); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -if (lean_obj_tag(x_189) == 0) -{ -lean_dec_ref(x_189); -x_52 = x_59; -x_53 = lean_box(0); -goto block_57; -} -else -{ -lean_object* x_190; lean_object* x_191; lean_object* x_192; -lean_dec_ref(x_2); -x_190 = lean_ctor_get(x_189, 0); -lean_inc(x_190); -if (lean_is_exclusive(x_189)) { - lean_ctor_release(x_189, 0); - x_191 = x_189; -} else { - lean_dec_ref(x_189); - x_191 = lean_box(0); -} -if (lean_is_scalar(x_191)) { - x_192 = lean_alloc_ctor(1, 1, 0); -} else { - x_192 = x_191; -} -lean_ctor_set(x_192, 0, x_190); -return x_192; -} -} -} -else -{ -lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; -lean_dec(x_164); -x_193 = lean_box(0); -x_194 = lean_array_get(x_193, x_166, x_1); -lean_dec(x_1); -lean_dec(x_166); -x_195 = lean_box(x_58); -lean_inc_ref(x_2); -lean_inc(x_194); -x_196 = lean_alloc_closure((void*)(l_Lean_Tactic_FunInd_rwMatcher___lam__0___boxed), 9, 3); -lean_closure_set(x_196, 0, x_194); -lean_closure_set(x_196, 1, x_195); -lean_closure_set(x_196, 2, x_2); -x_197 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__21___closed__2; -x_198 = l_Lean_Expr_constLevels_x21(x_163); -lean_dec_ref(x_163); -lean_inc(x_194); -x_199 = l_Lean_mkConst(x_194, x_198); -x_200 = l___private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0___lam__1___closed__0; -x_201 = l_Lean_Expr_getAppNumArgs(x_2); -lean_inc(x_201); -x_202 = lean_mk_array(x_201, x_200); -x_203 = lean_unsigned_to_nat(1u); -x_204 = lean_nat_sub(x_201, x_203); -lean_dec(x_201); -lean_inc_ref(x_2); -x_205 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_2, x_202, x_204); -x_206 = l_Lean_mkAppN(x_199, x_205); -lean_dec_ref(x_205); -x_207 = lean_box(x_59); -x_208 = lean_box(x_58); -lean_inc_ref(x_2); -lean_inc(x_194); -x_209 = lean_alloc_closure((void*)(l_Lean_Tactic_FunInd_rwMatcher___lam__1___boxed), 11, 6); -lean_closure_set(x_209, 0, x_206); -lean_closure_set(x_209, 1, x_197); -lean_closure_set(x_209, 2, x_207); -lean_closure_set(x_209, 3, x_194); -lean_closure_set(x_209, 4, x_2); -lean_closure_set(x_209, 5, x_208); -x_210 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__21___closed__3; -lean_inc(x_6); -lean_inc_ref(x_5); -lean_inc(x_4); -lean_inc_ref(x_3); -x_211 = l_Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7___redArg(x_197, x_196, x_209, x_59, x_210, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_211) == 0) -{ -lean_dec(x_194); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -lean_dec_ref(x_2); -return x_211; -} -else -{ -lean_object* x_212; lean_object* x_213; lean_object* x_214; uint8_t x_215; -x_212 = lean_ctor_get(x_211, 0); -lean_inc(x_212); -x_213 = lean_box(x_59); -x_214 = lean_alloc_closure((void*)(l_Lean_Tactic_FunInd_rwMatcher___lam__2___boxed), 8, 2); -lean_closure_set(x_214, 0, x_2); -lean_closure_set(x_214, 1, x_213); -x_215 = l_Lean_Exception_isInterrupt(x_212); -if (x_215 == 0) -{ -uint8_t x_216; -lean_inc(x_212); -x_216 = l_Lean_Exception_isRuntime(x_212); -x_19 = x_197; -x_20 = x_214; -x_21 = x_212; -x_22 = x_194; -x_23 = lean_box(0); -x_24 = x_211; -x_25 = x_216; -goto block_45; -} -else -{ -x_19 = x_197; -x_20 = x_214; -x_21 = x_212; -x_22 = x_194; -x_23 = lean_box(0); -x_24 = x_211; -x_25 = x_215; -goto block_45; -} -} -} -} -else -{ -lean_object* x_217; lean_object* x_218; lean_object* x_219; -lean_dec(x_164); -lean_dec_ref(x_163); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -lean_dec_ref(x_2); -lean_dec(x_1); -x_217 = lean_ctor_get(x_165, 0); -lean_inc(x_217); -if (lean_is_exclusive(x_165)) { - lean_ctor_release(x_165, 0); - x_218 = x_165; -} else { - lean_dec_ref(x_165); - x_218 = lean_box(0); -} -if (lean_is_scalar(x_218)) { - x_219 = lean_alloc_ctor(1, 1, 0); -} else { - x_219 = x_218; -} -lean_ctor_set(x_219, 0, x_217); -return x_219; -} -} -} -} -else -{ -lean_object* x_220; -lean_dec(x_1); -x_220 = l___private_Init_While_0__Lean_Loop_forIn_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__8(x_2, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_220) == 0) -{ -uint8_t x_221; -x_221 = !lean_is_exclusive(x_220); -if (x_221 == 0) -{ -lean_object* x_222; lean_object* x_223; lean_object* x_224; -x_222 = lean_ctor_get(x_220, 0); -x_223 = lean_box(0); -x_224 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_224, 0, x_222); -lean_ctor_set(x_224, 1, x_223); -lean_ctor_set_uint8(x_224, sizeof(void*)*2, x_59); -lean_ctor_set(x_220, 0, x_224); -return x_220; -} -else -{ -lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; -x_225 = lean_ctor_get(x_220, 0); -lean_inc(x_225); -lean_dec(x_220); -x_226 = lean_box(0); -x_227 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_227, 0, x_225); -lean_ctor_set(x_227, 1, x_226); -lean_ctor_set_uint8(x_227, sizeof(void*)*2, x_59); -x_228 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_228, 0, x_227); -return x_228; -} -} -else -{ -uint8_t x_229; -x_229 = !lean_is_exclusive(x_220); -if (x_229 == 0) -{ -return x_220; -} -else -{ -lean_object* x_230; lean_object* x_231; -x_230 = lean_ctor_get(x_220, 0); -lean_inc(x_230); -lean_dec(x_220); -x_231 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_231, 0, x_230); -return x_231; -} -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0___boxed(lean_object* x_1, 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_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec(x_1); -return x_7; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_isMatcherApp___at___00Lean_Tactic_FunInd_rwMatcher_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; -x_7 = l_Lean_Meta_isMatcherApp___at___00Lean_Tactic_FunInd_rwMatcher_spec__1(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -return x_7; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -uint8_t x_4; lean_object* x_5; -x_4 = l_Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0(x_1, x_2, x_3); -lean_dec(x_3); -x_5 = lean_box(x_4); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at___00Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3_spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; lean_object* x_8; -x_7 = l_Lean_PersistentHashMap_containsAtAux___at___00Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3_spec__12(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_6); -lean_dec_ref(x_3); -lean_dec_ref(x_2); -x_8 = lean_box(x_7); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -size_t x_5; uint8_t x_6; lean_object* x_7; -x_5 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_6 = l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3(x_1, x_2, x_5, x_4); -lean_dec(x_4); -x_7 = lean_box(x_6); -return x_7; -} -} -LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00Lean_Tactic_FunInd_rwMatcher_spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; -x_7 = l_Lean_instantiateMVars___at___00Lean_Tactic_FunInd_rwMatcher_spec__3(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_5); -lean_dec_ref(x_4); -lean_dec(x_3); -lean_dec_ref(x_2); -return x_7; -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -uint8_t x_12; lean_object* x_13; -x_12 = lean_unbox(x_5); -x_13 = l_Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7(x_1, x_2, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10); -return x_13; -} -} -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -x_8 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__10(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec_ref(x_3); -return x_8; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_6; -x_6 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__8(x_1, x_2, x_3, x_4); -lean_dec(x_4); -lean_dec_ref(x_3); -lean_dec(x_2); -lean_dec_ref(x_1); -return x_6; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; lean_object* x_4; -x_3 = l_Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0___redArg(x_1, x_2); -lean_dec(x_2); -x_4 = lean_box(x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__10___redArg___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_MonadExcept_ofExcept___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__10___redArg(x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_isMatcherApp___at___00Lean_Tactic_FunInd_rwMatcher_spec__1___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l_Lean_Meta_isMatcherApp___at___00Lean_Tactic_FunInd_rwMatcher_spec__1___redArg(x_1, x_2); -lean_dec(x_2); -lean_dec_ref(x_1); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l_Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0___redArg(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -return x_4; -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Tactic_FunInd_rwMatcher_spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -size_t x_4; size_t x_5; lean_object* x_6; -x_4 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_6 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Tactic_FunInd_rwMatcher_spec__2(x_4, x_5, x_3); -return x_6; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at___00Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3_spec__12___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -uint8_t x_4; lean_object* x_5; -x_4 = l_Lean_PersistentHashMap_containsAtAux___at___00Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3_spec__12___redArg(x_1, x_2, x_3); -lean_dec(x_3); -lean_dec_ref(x_1); -x_5 = lean_box(x_4); -return x_5; -} -} -LEAN_EXPORT lean_object* l___private_Init_While_0__Lean_Loop_forIn_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; -x_7 = l___private_Init_While_0__Lean_Loop_forIn_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__8(x_1, x_2, x_3, x_4, x_5); -return x_7; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -size_t x_4; uint8_t x_5; lean_object* x_6; -x_4 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_5 = l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3___redArg(x_1, x_4, x_3); -lean_dec(x_3); -x_6 = lean_box(x_5); -return x_6; -} -} -LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00Lean_Tactic_FunInd_rwMatcher_spec__3___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l_Lean_instantiateMVars___at___00Lean_Tactic_FunInd_rwMatcher_spec__3___redArg(x_1, x_2); -lean_dec(x_2); -return x_4; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__8___redArg___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__8___redArg(x_1); -lean_dec(x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Tactic_FunInd_rwMatcher_spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -size_t x_10; size_t x_11; lean_object* x_12; -x_10 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_11 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_12 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Tactic_FunInd_rwMatcher_spec__6(x_1, x_10, x_11, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec_ref(x_1); -return x_12; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; -x_10 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at___00Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7_spec__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_8); -lean_dec(x_6); -lean_dec_ref(x_5); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_rwMatcher___boxed(lean_object* x_1, lean_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_Tactic_FunInd_rwMatcher(x_1, x_2, x_3, x_4, x_5, x_6); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -uint8_t x_11; lean_object* x_12; -x_11 = lean_unbox(x_4); -x_12 = l_Lean_withTraceNode___at___00Lean_Tactic_FunInd_rwMatcher_spec__7___redArg(x_1, x_2, x_3, x_11, x_5, x_6, x_7, x_8, x_9); -return x_12; -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -uint8_t x_12; size_t x_13; size_t x_14; lean_object* x_15; -x_12 = lean_unbox(x_2); -x_13 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_14 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_15 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5(x_1, x_12, x_3, x_13, x_14, x_6, x_7, x_8, x_9, x_10); -lean_dec_ref(x_3); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__14___redArg___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_9; -x_9 = lean_apply_7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, lean_box(0)); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__14___redArg___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__14___redArg___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__14___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_11; lean_object* x_12; -x_11 = lean_alloc_closure((void*)(l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__14___redArg___lam__0___boxed), 8, 3); -lean_closure_set(x_11, 0, x_3); -lean_closure_set(x_11, 1, x_4); -lean_closure_set(x_11, 2, x_5); -x_12 = l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_withUserNamesImpl(lean_box(0), x_1, x_2, x_11, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_12) == 0) -{ -uint8_t x_13; -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) -{ -lean_object* x_14; uint8_t x_15; -x_14 = lean_ctor_get(x_12, 0); -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) -{ -return x_12; -} -else -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_14, 0); -x_17 = lean_ctor_get(x_14, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_14); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_16); -lean_ctor_set(x_18, 1, x_17); -lean_ctor_set(x_12, 0, x_18); -return x_12; -} -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_19 = lean_ctor_get(x_12, 0); -lean_inc(x_19); -lean_dec(x_12); -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 1); -lean_inc(x_21); -if (lean_is_exclusive(x_19)) { - lean_ctor_release(x_19, 0); - lean_ctor_release(x_19, 1); - x_22 = x_19; -} else { - lean_dec_ref(x_19); - x_22 = lean_box(0); -} -if (lean_is_scalar(x_22)) { - x_23 = lean_alloc_ctor(0, 2, 0); -} else { - x_23 = x_22; -} -lean_ctor_set(x_23, 0, x_20); -lean_ctor_set(x_23, 1, x_21); -x_24 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_24, 0, x_23); -return x_24; -} -} -else -{ -uint8_t x_25; -x_25 = !lean_is_exclusive(x_12); -if (x_25 == 0) -{ -return x_12; -} -else -{ -lean_object* x_26; lean_object* x_27; -x_26 = lean_ctor_get(x_12, 0); -lean_inc(x_26); -lean_dec(x_12); -x_27 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_27, 0, x_26); -return x_27; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__14(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_12; -x_12 = l_Lean_Meta_MatcherApp_withUserNames___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__14___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__8___redArg___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_11; -x_11 = lean_apply_9(x_1, x_4, x_5, x_2, x_3, x_6, x_7, x_8, x_9, lean_box(0)); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__8___redArg___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_forallBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__8___redArg___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__8___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_13; lean_object* x_14; -x_13 = lean_alloc_closure((void*)(l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__8___redArg___lam__0___boxed), 10, 3); -lean_closure_set(x_13, 0, x_3); -lean_closure_set(x_13, 1, x_6); -lean_closure_set(x_13, 2, x_7); -x_14 = l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAux(lean_box(0), x_1, x_2, x_13, x_4, x_5, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_14) == 0) -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) -{ -lean_object* x_16; uint8_t x_17; -x_16 = lean_ctor_get(x_14, 0); -x_17 = !lean_is_exclusive(x_16); -if (x_17 == 0) -{ -return x_14; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_16, 0); -x_19 = lean_ctor_get(x_16, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_16); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_18); -lean_ctor_set(x_20, 1, x_19); -lean_ctor_set(x_14, 0, x_20); -return x_14; -} -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_21 = lean_ctor_get(x_14, 0); -lean_inc(x_21); -lean_dec(x_14); -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_21, 1); -lean_inc(x_23); -if (lean_is_exclusive(x_21)) { - lean_ctor_release(x_21, 0); - lean_ctor_release(x_21, 1); - x_24 = x_21; -} else { - lean_dec_ref(x_21); - x_24 = lean_box(0); -} -if (lean_is_scalar(x_24)) { - x_25 = lean_alloc_ctor(0, 2, 0); -} else { - x_25 = x_24; -} -lean_ctor_set(x_25, 0, x_22); -lean_ctor_set(x_25, 1, x_23); -x_26 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_26, 0, x_25); -return x_26; -} -} -else -{ -uint8_t x_27; -x_27 = !lean_is_exclusive(x_14); -if (x_27 == 0) -{ -return x_14; -} -else -{ -lean_object* x_28; lean_object* x_29; -x_28 = lean_ctor_get(x_14, 0); -lean_inc(x_28); -lean_dec(x_14); -x_29 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_29, 0, x_28); -return x_29; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_14; -x_14 = l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__8___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_14; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__4_spec__6___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_6 = lean_st_ref_get(x_4); -x_7 = lean_ctor_get(x_6, 0); -lean_inc_ref(x_7); -lean_dec(x_6); -x_8 = l_Lean_Meta_Match_Extension_getMatcherInfo_x3f(x_7, x_1); -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_8); -lean_ctor_set(x_9, 1, x_2); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_3); -x_11 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_11, 0, x_10); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__4_spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_9; -x_9 = l_Lean_Meta_getMatcherInfo_x3f___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__4_spec__6___redArg(x_1, x_2, x_3, x_7); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_lambdaBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__9_spec__25___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_12; uint8_t x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; -x_12 = lean_alloc_closure((void*)(l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__8___redArg___lam__0___boxed), 10, 3); -lean_closure_set(x_12, 0, x_3); -lean_closure_set(x_12, 1, x_5); -lean_closure_set(x_12, 2, x_6); -x_13 = 1; -x_14 = 0; -x_15 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_15, 0, x_2); -x_16 = l___private_Lean_Meta_Basic_0__Lean_Meta_lambdaTelescopeImp(lean_box(0), x_1, x_13, x_14, x_13, x_14, x_15, x_12, x_4, x_7, x_8, x_9, x_10); -lean_dec_ref(x_15); -if (lean_obj_tag(x_16) == 0) -{ -uint8_t x_17; -x_17 = !lean_is_exclusive(x_16); -if (x_17 == 0) -{ -lean_object* x_18; uint8_t x_19; -x_18 = lean_ctor_get(x_16, 0); -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 0) -{ -return x_16; -} -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_18, 0); -x_21 = lean_ctor_get(x_18, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_18); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -lean_ctor_set(x_16, 0, x_22); -return x_16; -} -} -else -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_23 = lean_ctor_get(x_16, 0); -lean_inc(x_23); -lean_dec(x_16); -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); -lean_inc(x_25); -if (lean_is_exclusive(x_23)) { - lean_ctor_release(x_23, 0); - lean_ctor_release(x_23, 1); - x_26 = x_23; -} else { - lean_dec_ref(x_23); - x_26 = lean_box(0); -} -if (lean_is_scalar(x_26)) { - x_27 = lean_alloc_ctor(0, 2, 0); -} else { - x_27 = x_26; -} -lean_ctor_set(x_27, 0, x_24); -lean_ctor_set(x_27, 1, x_25); -x_28 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_28, 0, x_27); -return x_28; -} -} -else -{ -uint8_t x_29; -x_29 = !lean_is_exclusive(x_16); -if (x_29 == 0) -{ -return x_16; -} -else -{ -lean_object* x_30; lean_object* x_31; -x_30 = lean_ctor_get(x_16, 0); -lean_inc(x_30); -lean_dec(x_16); -x_31 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_31, 0, x_30); -return x_31; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_lambdaBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__9_spec__25(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_13; -x_13 = l_Lean_Meta_lambdaBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__9_spec__25___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__12___redArg(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_11; uint8_t x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; -x_11 = lean_alloc_closure((void*)(l_Lean_Meta_forallBoundedTelescope___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__8___redArg___lam__0___boxed), 10, 3); -lean_closure_set(x_11, 0, x_2); -lean_closure_set(x_11, 1, x_4); -lean_closure_set(x_11, 2, x_5); -x_12 = 1; -x_13 = 0; -x_14 = lean_box(0); -x_15 = l___private_Lean_Meta_Basic_0__Lean_Meta_lambdaTelescopeImp(lean_box(0), x_1, x_12, x_13, x_12, x_13, x_14, x_11, x_3, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_15) == 0) -{ -uint8_t x_16; -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) -{ -lean_object* x_17; uint8_t x_18; -x_17 = lean_ctor_get(x_15, 0); -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 0) -{ -return x_15; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_17, 0); -x_20 = lean_ctor_get(x_17, 1); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_17); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -lean_ctor_set(x_15, 0, x_21); -return x_15; -} -} -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_22 = lean_ctor_get(x_15, 0); -lean_inc(x_22); -lean_dec(x_15); -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_22, 1); -lean_inc(x_24); -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); -} -if (lean_is_scalar(x_25)) { - x_26 = lean_alloc_ctor(0, 2, 0); -} else { - x_26 = x_25; -} -lean_ctor_set(x_26, 0, x_23); -lean_ctor_set(x_26, 1, x_24); -x_27 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_27, 0, x_26); -return x_27; -} -} -else -{ -uint8_t x_28; -x_28 = !lean_is_exclusive(x_15); -if (x_28 == 0) -{ -return x_15; -} -else -{ -lean_object* x_29; lean_object* x_30; -x_29 = lean_ctor_get(x_15, 0); -lean_inc(x_29); -lean_dec(x_15); -x_30 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_30, 0, x_29); -return x_30; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_12; -x_12 = l_Lean_Meta_lambdaTelescope___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__12___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__1___redArg___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_10; -x_10 = lean_apply_8(x_1, x_4, x_2, x_3, x_5, x_6, x_7, x_8, lean_box(0)); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__1___redArg___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; -x_10 = l_Lean_Meta_withLetDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__1___redArg___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__1___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_14; lean_object* x_15; -x_14 = lean_alloc_closure((void*)(l_Lean_Meta_withLetDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__1___redArg___lam__0___boxed), 9, 3); -lean_closure_set(x_14, 0, x_4); -lean_closure_set(x_14, 1, x_7); -lean_closure_set(x_14, 2, x_8); -x_15 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLetDeclImp(lean_box(0), x_1, x_2, x_3, x_14, x_5, x_6, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_15) == 0) -{ -uint8_t x_16; -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) -{ -lean_object* x_17; uint8_t x_18; -x_17 = lean_ctor_get(x_15, 0); -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 0) -{ -return x_15; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_17, 0); -x_20 = lean_ctor_get(x_17, 1); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_17); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -lean_ctor_set(x_15, 0, x_21); -return x_15; -} -} -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_22 = lean_ctor_get(x_15, 0); -lean_inc(x_22); -lean_dec(x_15); -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_22, 1); -lean_inc(x_24); -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); -} -if (lean_is_scalar(x_25)) { - x_26 = lean_alloc_ctor(0, 2, 0); -} else { - x_26 = x_25; -} -lean_ctor_set(x_26, 0, x_23); -lean_ctor_set(x_26, 1, x_24); -x_27 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_27, 0, x_26); -return x_27; -} -} -else -{ -uint8_t x_28; -x_28 = !lean_is_exclusive(x_15); -if (x_28 == 0) -{ -return x_15; -} -else -{ -lean_object* x_29; lean_object* x_30; -x_29 = lean_ctor_get(x_15, 0); -lean_inc(x_29); -lean_dec(x_15); -x_30 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_30, 0, x_29); -return x_30; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: -{ -lean_object* x_15; -x_15 = l_Lean_Meta_withLetDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__1___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__2___redArg(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_13; lean_object* x_14; -x_13 = lean_alloc_closure((void*)(l_Lean_Meta_withLetDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__1___redArg___lam__0___boxed), 9, 3); -lean_closure_set(x_13, 0, x_4); -lean_closure_set(x_13, 1, x_6); -lean_closure_set(x_13, 2, x_7); -x_14 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp(lean_box(0), x_1, x_2, x_3, x_13, x_5, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_14) == 0) -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) -{ -lean_object* x_16; uint8_t x_17; -x_16 = lean_ctor_get(x_14, 0); -x_17 = !lean_is_exclusive(x_16); -if (x_17 == 0) -{ -return x_14; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_16, 0); -x_19 = lean_ctor_get(x_16, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_16); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_18); -lean_ctor_set(x_20, 1, x_19); -lean_ctor_set(x_14, 0, x_20); -return x_14; -} -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_21 = lean_ctor_get(x_14, 0); -lean_inc(x_21); -lean_dec(x_14); -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_21, 1); -lean_inc(x_23); -if (lean_is_exclusive(x_21)) { - lean_ctor_release(x_21, 0); - lean_ctor_release(x_21, 1); - x_24 = x_21; -} else { - lean_dec_ref(x_21); - x_24 = lean_box(0); -} -if (lean_is_scalar(x_24)) { - x_25 = lean_alloc_ctor(0, 2, 0); -} else { - x_25 = x_24; -} -lean_ctor_set(x_25, 0, x_22); -lean_ctor_set(x_25, 1, x_23); -x_26 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_26, 0, x_25); -return x_26; -} -} -else -{ -uint8_t x_27; -x_27 = !lean_is_exclusive(x_14); -if (x_27 == 0) -{ -return x_14; -} -else -{ -lean_object* x_28; lean_object* x_29; -x_28 = lean_ctor_get(x_14, 0); -lean_inc(x_28); -lean_dec(x_14); -x_29 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_29, 0, x_28); -return x_29; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__2(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_14; -x_14 = l_Lean_Meta_withLocalDecl___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__2___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_14; -} -} -LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__4_spec__4_spec__7_spec__18_spec__28_spec__31_spec__33___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_15 = lean_st_ref_get(x_5); -x_16 = lean_ctor_get(x_15, 0); -lean_inc_ref(x_16); -lean_dec(x_15); -x_17 = l_Lean_Name_isAnonymous(x_2); -if (x_17 == 0) -{ -uint8_t x_18; -x_18 = lean_ctor_get_uint8(x_16, sizeof(void*)*8); -if (x_18 == 0) -{ -lean_dec_ref(x_16); -lean_dec(x_2); -x_7 = x_1; -x_8 = x_3; -x_9 = x_4; -x_10 = lean_box(0); -goto block_14; -} -else -{ -lean_object* x_19; uint8_t x_20; -lean_inc_ref(x_16); -x_19 = l_Lean_Environment_setExporting(x_16, x_17); -lean_inc(x_2); -lean_inc_ref(x_19); -x_20 = l_Lean_Environment_contains(x_19, x_2, x_18); -if (x_20 == 0) -{ -lean_dec_ref(x_19); -lean_dec_ref(x_16); -lean_dec(x_2); -x_7 = x_1; -x_8 = x_3; -x_9 = x_4; -x_10 = lean_box(0); -goto block_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_21 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg___closed__2; -x_22 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg___closed__6; -x_23 = lean_box(0); -x_24 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_24, 0, x_19); -lean_ctor_set(x_24, 1, x_21); -lean_ctor_set(x_24, 2, x_22); -lean_ctor_set(x_24, 3, x_23); -lean_inc(x_2); -x_25 = l_Lean_MessageData_ofConstName(x_2, x_17); -x_26 = lean_alloc_ctor(3, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -x_27 = l_Lean_Environment_getModuleIdxFor_x3f(x_16, x_2); -if (lean_obj_tag(x_27) == 0) -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -lean_dec_ref(x_16); -lean_dec(x_2); -x_28 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg___closed__8; -x_29 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_26); -x_30 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg___closed__10; -x_31 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_30); -x_32 = l_Lean_MessageData_note(x_31); -x_33 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_33, 0, x_1); -lean_ctor_set(x_33, 1, x_32); -x_7 = x_33; -x_8 = x_3; -x_9 = x_4; -x_10 = lean_box(0); -goto block_14; -} -else -{ -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; -x_34 = lean_ctor_get(x_27, 0); -lean_inc(x_34); -lean_dec_ref(x_27); -x_35 = lean_box(0); -x_36 = l_Lean_Environment_header(x_16); -lean_dec_ref(x_16); -x_37 = l_Lean_EnvironmentHeader_moduleNames(x_36); -x_38 = lean_array_get(x_35, x_37, x_34); -lean_dec(x_34); -lean_dec_ref(x_37); -x_39 = l_Lean_isPrivateName(x_2); -lean_dec(x_2); -if (x_39 == 0) -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_40 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg___closed__12; -x_41 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_26); -x_42 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg___closed__14; -x_43 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_43, 0, x_41); -lean_ctor_set(x_43, 1, x_42); -x_44 = l_Lean_MessageData_ofName(x_38); -x_45 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_45, 0, x_43); -lean_ctor_set(x_45, 1, x_44); -x_46 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg___closed__16; -x_47 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_46); -x_48 = l_Lean_MessageData_note(x_47); -x_49 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_49, 0, x_1); -lean_ctor_set(x_49, 1, x_48); -x_7 = x_49; -x_8 = x_3; -x_9 = x_4; -x_10 = lean_box(0); -goto block_14; +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; +x_34 = lean_ctor_get(x_27, 0); +lean_inc(x_34); +lean_dec_ref(x_27); +x_35 = lean_box(0); +x_36 = l_Lean_Environment_header(x_16); +lean_dec_ref(x_16); +x_37 = l_Lean_EnvironmentHeader_moduleNames(x_36); +x_38 = lean_array_get(x_35, x_37, x_34); +lean_dec(x_34); +lean_dec_ref(x_37); +x_39 = l_Lean_isPrivateName(x_2); +lean_dec(x_2); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_40 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg___closed__12; +x_41 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_26); +x_42 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg___closed__14; +x_43 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +x_44 = l_Lean_MessageData_ofName(x_38); +x_45 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +x_46 = l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_Meta_matchMatcherApp_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__1_spec__2_spec__8_spec__33_spec__42_spec__46_spec__48___redArg___closed__16; +x_47 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +x_48 = l_Lean_MessageData_note(x_47); +x_49 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_49, 0, x_1); +lean_ctor_set(x_49, 1, x_48); +x_7 = x_49; +x_8 = x_3; +x_9 = x_4; +x_10 = lean_box(0); +goto block_14; } else { @@ -66887,6 +59393,23 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__1___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(":", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__1___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__1___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { @@ -66917,7 +59440,7 @@ x_25 = l_Lean_MessageData_ofExpr(x_3); x_26 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_26, 0, x_24); lean_ctor_set(x_26, 1, x_25); -x_27 = l_Lean_Tactic_FunInd_rwMatcher___closed__3; +x_27 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__1___closed__5; x_28 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_28, 0, x_26); lean_ctor_set(x_28, 1, x_27); @@ -67113,7 +59636,7 @@ lean_closure_set(x_17, 2, x_3); lean_closure_set(x_17, 3, x_4); lean_closure_set(x_17, 4, x_5); lean_closure_set(x_17, 5, x_9); -x_18 = lean_alloc_closure((void*)(l_Lean_Tactic_FunInd_rwMatcher___boxed), 7, 1); +x_18 = lean_alloc_closure((void*)(l_Lean_Meta_rwMatcher___boxed), 7, 1); lean_closure_set(x_18, 0, x_6); x_19 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_withRewrittenMotiveArg___boxed), 10, 3); lean_closure_set(x_19, 0, x_7); @@ -67997,7 +60520,7 @@ return x_9; LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__16_spec__20_spec__27(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_57; lean_object* x_58; uint8_t x_59; uint8_t x_60; lean_object* x_61; uint8_t x_62; lean_object* x_63; lean_object* x_64; lean_object* x_82; lean_object* x_83; uint8_t x_84; uint8_t x_85; uint8_t x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_93; lean_object* x_94; lean_object* x_95; uint8_t x_96; uint8_t x_97; lean_object* x_98; uint8_t x_99; uint8_t x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; uint8_t x_109; lean_object* x_110; uint8_t x_111; uint8_t x_112; uint8_t x_114; uint8_t x_132; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_57; lean_object* x_58; uint8_t x_59; uint8_t x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; lean_object* x_64; lean_object* x_82; lean_object* x_83; uint8_t x_84; lean_object* x_85; uint8_t x_86; lean_object* x_87; uint8_t x_88; lean_object* x_89; lean_object* x_93; uint8_t x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; uint8_t x_98; uint8_t x_99; uint8_t x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; uint8_t x_110; uint8_t x_111; uint8_t x_112; uint8_t x_114; uint8_t x_132; x_105 = 2; x_132 = l_Lean_instBEqMessageSeverity_beq(x_3, x_105); if (x_132 == 0) @@ -68045,15 +60568,15 @@ lean_ctor_set(x_34, 0, x_30); lean_ctor_set(x_34, 1, x_31); x_35 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_19); +lean_ctor_set(x_35, 1, x_25); x_36 = lean_alloc_ctor(0, 5, 3); -lean_ctor_set(x_36, 0, x_20); -lean_ctor_set(x_36, 1, x_23); -lean_ctor_set(x_36, 2, x_25); -lean_ctor_set(x_36, 3, x_22); +lean_ctor_set(x_36, 0, x_24); +lean_ctor_set(x_36, 1, x_20); +lean_ctor_set(x_36, 2, x_21); +lean_ctor_set(x_36, 3, x_19); lean_ctor_set(x_36, 4, x_35); -lean_ctor_set_uint8(x_36, sizeof(void*)*5, x_21); -lean_ctor_set_uint8(x_36, sizeof(void*)*5 + 1, x_24); +lean_ctor_set_uint8(x_36, sizeof(void*)*5, x_22); +lean_ctor_set_uint8(x_36, sizeof(void*)*5 + 1, x_23); lean_ctor_set_uint8(x_36, sizeof(void*)*5 + 2, x_4); x_37 = l_Lean_MessageLog_add(x_36, x_33); lean_ctor_set(x_29, 6, x_37); @@ -68091,15 +60614,15 @@ lean_ctor_set(x_49, 0, x_30); lean_ctor_set(x_49, 1, x_31); x_50 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_50, 0, x_49); -lean_ctor_set(x_50, 1, x_19); +lean_ctor_set(x_50, 1, x_25); x_51 = lean_alloc_ctor(0, 5, 3); -lean_ctor_set(x_51, 0, x_20); -lean_ctor_set(x_51, 1, x_23); -lean_ctor_set(x_51, 2, x_25); -lean_ctor_set(x_51, 3, x_22); +lean_ctor_set(x_51, 0, x_24); +lean_ctor_set(x_51, 1, x_20); +lean_ctor_set(x_51, 2, x_21); +lean_ctor_set(x_51, 3, x_19); lean_ctor_set(x_51, 4, x_50); -lean_ctor_set_uint8(x_51, sizeof(void*)*5, x_21); -lean_ctor_set_uint8(x_51, sizeof(void*)*5 + 1, x_24); +lean_ctor_set_uint8(x_51, sizeof(void*)*5, x_22); +lean_ctor_set_uint8(x_51, sizeof(void*)*5 + 1, x_23); lean_ctor_set_uint8(x_51, sizeof(void*)*5 + 2, x_4); x_52 = l_Lean_MessageLog_add(x_51, x_46); x_53 = lean_alloc_ctor(0, 9, 0); @@ -68130,24 +60653,24 @@ if (x_67 == 0) { lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; x_68 = lean_ctor_get(x_66, 0); -lean_inc_ref(x_63); -x_69 = l_Lean_FileMap_toPosition(x_63, x_61); -lean_dec(x_61); -x_70 = l_Lean_FileMap_toPosition(x_63, x_64); +lean_inc_ref(x_61); +x_69 = l_Lean_FileMap_toPosition(x_61, x_58); +lean_dec(x_58); +x_70 = l_Lean_FileMap_toPosition(x_61, x_64); lean_dec(x_64); lean_ctor_set_tag(x_66, 1); lean_ctor_set(x_66, 0, x_70); x_71 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__21___closed__3; -if (x_62 == 0) +if (x_63 == 0) { lean_dec_ref(x_57); -x_19 = x_68; -x_20 = x_58; -x_21 = x_59; -x_22 = x_71; -x_23 = x_69; -x_24 = x_60; -x_25 = x_66; +x_19 = x_71; +x_20 = x_69; +x_21 = x_66; +x_22 = x_59; +x_23 = x_60; +x_24 = x_62; +x_25 = x_68; x_26 = x_9; x_27 = x_10; x_28 = lean_box(0); @@ -68164,7 +60687,7 @@ lean_object* x_73; lean_dec_ref(x_66); lean_dec_ref(x_69); lean_dec(x_68); -lean_dec_ref(x_58); +lean_dec_ref(x_62); lean_dec_ref(x_9); x_73 = lean_box(0); x_12 = x_73; @@ -68174,13 +60697,13 @@ goto block_18; } else { -x_19 = x_68; -x_20 = x_58; -x_21 = x_59; -x_22 = x_71; -x_23 = x_69; -x_24 = x_60; -x_25 = x_66; +x_19 = x_71; +x_20 = x_69; +x_21 = x_66; +x_22 = x_59; +x_23 = x_60; +x_24 = x_62; +x_25 = x_68; x_26 = x_9; x_27 = x_10; x_28 = lean_box(0); @@ -68194,24 +60717,24 @@ lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean x_74 = lean_ctor_get(x_66, 0); lean_inc(x_74); lean_dec(x_66); -lean_inc_ref(x_63); -x_75 = l_Lean_FileMap_toPosition(x_63, x_61); -lean_dec(x_61); -x_76 = l_Lean_FileMap_toPosition(x_63, x_64); +lean_inc_ref(x_61); +x_75 = l_Lean_FileMap_toPosition(x_61, x_58); +lean_dec(x_58); +x_76 = l_Lean_FileMap_toPosition(x_61, x_64); lean_dec(x_64); x_77 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_77, 0, x_76); x_78 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__21___closed__3; -if (x_62 == 0) +if (x_63 == 0) { lean_dec_ref(x_57); -x_19 = x_74; -x_20 = x_58; -x_21 = x_59; -x_22 = x_78; -x_23 = x_75; -x_24 = x_60; -x_25 = x_77; +x_19 = x_78; +x_20 = x_75; +x_21 = x_77; +x_22 = x_59; +x_23 = x_60; +x_24 = x_62; +x_25 = x_74; x_26 = x_9; x_27 = x_10; x_28 = lean_box(0); @@ -68228,7 +60751,7 @@ lean_object* x_80; lean_dec_ref(x_77); lean_dec_ref(x_75); lean_dec(x_74); -lean_dec_ref(x_58); +lean_dec_ref(x_62); lean_dec_ref(x_9); x_80 = lean_box(0); x_12 = x_80; @@ -68238,13 +60761,13 @@ goto block_18; } else { -x_19 = x_74; -x_20 = x_58; -x_21 = x_59; -x_22 = x_78; -x_23 = x_75; -x_24 = x_60; -x_25 = x_77; +x_19 = x_78; +x_20 = x_75; +x_21 = x_77; +x_22 = x_59; +x_23 = x_60; +x_24 = x_62; +x_25 = x_74; x_26 = x_9; x_27 = x_10; x_28 = lean_box(0); @@ -68256,17 +60779,17 @@ goto block_56; block_92: { lean_object* x_90; -x_90 = l_Lean_Syntax_getTailPos_x3f(x_87, x_84); -lean_dec(x_87); +x_90 = l_Lean_Syntax_getTailPos_x3f(x_83, x_84); +lean_dec(x_83); if (lean_obj_tag(x_90) == 0) { lean_inc(x_89); x_57 = x_82; -x_58 = x_83; +x_58 = x_89; x_59 = x_84; -x_60 = x_85; -x_61 = x_89; -x_62 = x_86; +x_60 = x_86; +x_61 = x_85; +x_62 = x_87; x_63 = x_88; x_64 = x_89; goto block_81; @@ -68278,11 +60801,11 @@ x_91 = lean_ctor_get(x_90, 0); lean_inc(x_91); lean_dec_ref(x_90); x_57 = x_82; -x_58 = x_83; +x_58 = x_89; x_59 = x_84; -x_60 = x_85; -x_61 = x_89; -x_62 = x_86; +x_60 = x_86; +x_61 = x_85; +x_62 = x_87; x_63 = x_88; x_64 = x_91; goto block_81; @@ -68291,19 +60814,19 @@ goto block_81; block_104: { lean_object* x_100; lean_object* x_101; -x_100 = l_Lean_replaceRef(x_1, x_94); -lean_dec(x_94); -x_101 = l_Lean_Syntax_getPos_x3f(x_100, x_96); +x_100 = l_Lean_replaceRef(x_1, x_96); +lean_dec(x_96); +x_101 = l_Lean_Syntax_getPos_x3f(x_100, x_94); if (lean_obj_tag(x_101) == 0) { lean_object* x_102; x_102 = lean_unsigned_to_nat(0u); x_82 = x_93; -x_83 = x_95; -x_84 = x_96; -x_85 = x_99; -x_86 = x_97; -x_87 = x_100; +x_83 = x_100; +x_84 = x_94; +x_85 = x_95; +x_86 = x_99; +x_87 = x_97; x_88 = x_98; x_89 = x_102; goto block_92; @@ -68315,11 +60838,11 @@ x_103 = lean_ctor_get(x_101, 0); lean_inc(x_103); lean_dec_ref(x_101); x_82 = x_93; -x_83 = x_95; -x_84 = x_96; -x_85 = x_99; -x_86 = x_97; -x_87 = x_100; +x_83 = x_100; +x_84 = x_94; +x_85 = x_95; +x_86 = x_99; +x_87 = x_97; x_88 = x_98; x_89 = x_103; goto block_92; @@ -68329,10 +60852,10 @@ goto block_92; { if (x_112 == 0) { -x_93 = x_108; -x_94 = x_106; -x_95 = x_107; -x_96 = x_111; +x_93 = x_106; +x_94 = x_111; +x_95 = x_108; +x_96 = x_107; x_97 = x_109; x_98 = x_110; x_99 = x_3; @@ -68340,10 +60863,10 @@ goto block_104; } else { -x_93 = x_108; -x_94 = x_106; -x_95 = x_107; -x_96 = x_111; +x_93 = x_106; +x_94 = x_111; +x_95 = x_108; +x_96 = x_107; x_97 = x_109; x_98 = x_110; x_99 = x_105; @@ -68369,14 +60892,14 @@ x_123 = 1; x_124 = l_Lean_instBEqMessageSeverity_beq(x_3, x_123); if (x_124 == 0) { -lean_inc_ref(x_116); lean_inc_ref(x_115); +lean_inc_ref(x_116); lean_inc(x_118); -x_106 = x_118; -x_107 = x_115; -x_108 = x_122; -x_109 = x_119; -x_110 = x_116; +x_106 = x_122; +x_107 = x_118; +x_108 = x_116; +x_109 = x_115; +x_110 = x_119; x_111 = x_114; x_112 = x_124; goto block_113; @@ -68386,14 +60909,14 @@ else lean_object* x_125; uint8_t x_126; x_125 = l_Lean_logAt___at___00Lean_log___at___00Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__25_spec__27_spec__41___closed__0; x_126 = l_Lean_Option_get___at___00Lean_withTraceNode___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__13_spec__34(x_117, x_125); -lean_inc_ref(x_116); lean_inc_ref(x_115); +lean_inc_ref(x_116); lean_inc(x_118); -x_106 = x_118; -x_107 = x_115; -x_108 = x_122; -x_109 = x_119; -x_110 = x_116; +x_106 = x_122; +x_107 = x_118; +x_108 = x_116; +x_109 = x_115; +x_110 = x_119; x_111 = x_114; x_112 = x_126; goto block_113; @@ -71294,7 +63817,7 @@ return x_18; LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5(lean_object* x_1, uint8_t x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; uint8_t x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_238; lean_object* x_239; uint8_t x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; uint8_t x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; uint8_t x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; uint8_t x_400; lean_object* x_401; lean_object* x_402; size_t x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; uint8_t x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; uint8_t x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_238; lean_object* x_239; lean_object* x_240; uint8_t x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; uint8_t x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; uint8_t x_369; lean_object* x_370; lean_object* x_371; uint8_t x_400; uint8_t x_401; size_t x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; x_15 = lean_st_ref_get(x_13); x_16 = lean_ctor_get(x_15, 0); lean_inc_ref(x_16); @@ -71510,8 +64033,8 @@ lean_inc(x_46); lean_inc_ref(x_45); lean_inc(x_44); lean_inc_ref(x_43); -lean_inc(x_38); -x_48 = l_Lean_Meta_inferArgumentTypesN(x_38, x_29, x_43, x_44, x_45, x_46); +lean_inc(x_25); +x_48 = l_Lean_Meta_inferArgumentTypesN(x_25, x_28, x_43, x_44, x_45, x_46); if (lean_obj_tag(x_48) == 0) { lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; @@ -71529,40 +64052,40 @@ lean_inc(x_53); x_54 = lean_ctor_get(x_17, 5); lean_inc_ref(x_54); lean_dec_ref(x_17); -x_55 = !lean_is_exclusive(x_34); +x_55 = !lean_is_exclusive(x_38); if (x_55 == 0) { lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_56 = lean_ctor_get(x_34, 2); -x_57 = lean_ctor_get(x_34, 5); +x_56 = lean_ctor_get(x_38, 2); +x_57 = lean_ctor_get(x_38, 5); lean_dec(x_57); -x_58 = lean_ctor_get(x_34, 4); +x_58 = lean_ctor_get(x_38, 4); lean_dec(x_58); -x_59 = lean_ctor_get(x_34, 3); +x_59 = lean_ctor_get(x_38, 3); lean_dec(x_59); -x_60 = lean_ctor_get(x_34, 1); +x_60 = lean_ctor_get(x_38, 1); lean_dec(x_60); -x_61 = lean_ctor_get(x_34, 0); +x_61 = lean_ctor_get(x_38, 0); lean_dec(x_61); -lean_inc(x_38); -lean_inc(x_39); -x_62 = l_Array_toSubarray___redArg(x_23, x_39, x_38); +lean_inc(x_25); +lean_inc(x_26); +x_62 = l_Array_toSubarray___redArg(x_23, x_26, x_25); x_63 = lean_array_get_size(x_52); -lean_inc(x_39); +lean_inc(x_26); lean_inc_ref(x_52); -x_64 = l_Array_toSubarray___redArg(x_52, x_39, x_63); +x_64 = l_Array_toSubarray___redArg(x_52, x_26, x_63); x_65 = lean_array_get_size(x_56); -lean_inc(x_39); -x_66 = l_Array_toSubarray___redArg(x_56, x_39, x_65); -x_67 = lean_array_get_size(x_28); -lean_inc(x_39); -x_68 = l_Array_toSubarray___redArg(x_28, x_39, x_67); +lean_inc(x_26); +x_66 = l_Array_toSubarray___redArg(x_56, x_26, x_65); +x_67 = lean_array_get_size(x_35); +lean_inc(x_26); +x_68 = l_Array_toSubarray___redArg(x_35, x_26, x_67); x_69 = lean_array_get_size(x_49); -lean_inc(x_39); -x_70 = l_Array_toSubarray___redArg(x_49, x_39, x_69); +lean_inc(x_26); +x_70 = l_Array_toSubarray___redArg(x_49, x_26, x_69); x_71 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_33); +lean_ctor_set(x_71, 1, x_27); x_72 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_72, 0, x_68); lean_ctor_set(x_72, 1, x_71); @@ -71579,10 +64102,10 @@ lean_inc(x_46); lean_inc_ref(x_45); lean_inc(x_44); lean_inc_ref(x_43); -lean_inc(x_39); -x_76 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__19___redArg(x_38, x_39, x_6, x_30, x_2, x_25, x_26, x_39, x_75, x_41, x_42, x_43, x_44, x_45, x_46); -lean_dec(x_39); -lean_dec(x_38); +lean_inc(x_26); +x_76 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__19___redArg(x_25, x_26, x_6, x_36, x_2, x_33, x_31, x_26, x_75, x_41, x_42, x_43, x_44, x_45, x_46); +lean_dec(x_26); +lean_dec(x_25); if (lean_obj_tag(x_76) == 0) { lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; @@ -71627,19 +64150,19 @@ lean_object* x_93; lean_object* x_94; lean_object* x_95; x_93 = lean_ctor_get(x_91, 0); x_94 = l_Array_append___redArg(x_40, x_93); lean_dec(x_93); -lean_ctor_set(x_34, 5, x_54); -lean_ctor_set(x_34, 4, x_35); -lean_ctor_set(x_34, 3, x_53); -lean_ctor_set(x_34, 2, x_52); -lean_ctor_set(x_34, 1, x_51); -lean_ctor_set(x_34, 0, x_50); +lean_ctor_set(x_38, 5, x_54); +lean_ctor_set(x_38, 4, x_39); +lean_ctor_set(x_38, 3, x_53); +lean_ctor_set(x_38, 2, x_52); +lean_ctor_set(x_38, 1, x_51); +lean_ctor_set(x_38, 0, x_50); x_95 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_95, 0, x_34); -lean_ctor_set(x_95, 1, x_27); -lean_ctor_set(x_95, 2, x_36); -lean_ctor_set(x_95, 3, x_32); -lean_ctor_set(x_95, 4, x_31); -lean_ctor_set(x_95, 5, x_37); +lean_ctor_set(x_95, 0, x_38); +lean_ctor_set(x_95, 1, x_37); +lean_ctor_set(x_95, 2, x_29); +lean_ctor_set(x_95, 3, x_34); +lean_ctor_set(x_95, 4, x_32); +lean_ctor_set(x_95, 5, x_30); lean_ctor_set(x_95, 6, x_86); lean_ctor_set(x_95, 7, x_94); lean_ctor_set(x_91, 0, x_95); @@ -71655,19 +64178,19 @@ lean_inc(x_96); lean_dec(x_91); x_98 = l_Array_append___redArg(x_40, x_96); lean_dec(x_96); -lean_ctor_set(x_34, 5, x_54); -lean_ctor_set(x_34, 4, x_35); -lean_ctor_set(x_34, 3, x_53); -lean_ctor_set(x_34, 2, x_52); -lean_ctor_set(x_34, 1, x_51); -lean_ctor_set(x_34, 0, x_50); +lean_ctor_set(x_38, 5, x_54); +lean_ctor_set(x_38, 4, x_39); +lean_ctor_set(x_38, 3, x_53); +lean_ctor_set(x_38, 2, x_52); +lean_ctor_set(x_38, 1, x_51); +lean_ctor_set(x_38, 0, x_50); x_99 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_99, 0, x_34); -lean_ctor_set(x_99, 1, x_27); -lean_ctor_set(x_99, 2, x_36); -lean_ctor_set(x_99, 3, x_32); -lean_ctor_set(x_99, 4, x_31); -lean_ctor_set(x_99, 5, x_37); +lean_ctor_set(x_99, 0, x_38); +lean_ctor_set(x_99, 1, x_37); +lean_ctor_set(x_99, 2, x_29); +lean_ctor_set(x_99, 3, x_34); +lean_ctor_set(x_99, 4, x_32); +lean_ctor_set(x_99, 5, x_30); lean_ctor_set(x_99, 6, x_86); lean_ctor_set(x_99, 7, x_98); x_100 = lean_alloc_ctor(0, 2, 0); @@ -71699,19 +64222,19 @@ if (lean_is_exclusive(x_101)) { } x_106 = l_Array_append___redArg(x_40, x_103); lean_dec(x_103); -lean_ctor_set(x_34, 5, x_54); -lean_ctor_set(x_34, 4, x_35); -lean_ctor_set(x_34, 3, x_53); -lean_ctor_set(x_34, 2, x_52); -lean_ctor_set(x_34, 1, x_51); -lean_ctor_set(x_34, 0, x_50); +lean_ctor_set(x_38, 5, x_54); +lean_ctor_set(x_38, 4, x_39); +lean_ctor_set(x_38, 3, x_53); +lean_ctor_set(x_38, 2, x_52); +lean_ctor_set(x_38, 1, x_51); +lean_ctor_set(x_38, 0, x_50); x_107 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_107, 0, x_34); -lean_ctor_set(x_107, 1, x_27); -lean_ctor_set(x_107, 2, x_36); -lean_ctor_set(x_107, 3, x_32); -lean_ctor_set(x_107, 4, x_31); -lean_ctor_set(x_107, 5, x_37); +lean_ctor_set(x_107, 0, x_38); +lean_ctor_set(x_107, 1, x_37); +lean_ctor_set(x_107, 2, x_29); +lean_ctor_set(x_107, 3, x_34); +lean_ctor_set(x_107, 4, x_32); +lean_ctor_set(x_107, 5, x_30); lean_ctor_set(x_107, 6, x_86); lean_ctor_set(x_107, 7, x_106); if (lean_is_scalar(x_105)) { @@ -71760,19 +64283,19 @@ if (lean_is_exclusive(x_111)) { } x_117 = l_Array_append___redArg(x_40, x_114); lean_dec(x_114); -lean_ctor_set(x_34, 5, x_54); -lean_ctor_set(x_34, 4, x_35); -lean_ctor_set(x_34, 3, x_53); -lean_ctor_set(x_34, 2, x_52); -lean_ctor_set(x_34, 1, x_51); -lean_ctor_set(x_34, 0, x_50); +lean_ctor_set(x_38, 5, x_54); +lean_ctor_set(x_38, 4, x_39); +lean_ctor_set(x_38, 3, x_53); +lean_ctor_set(x_38, 2, x_52); +lean_ctor_set(x_38, 1, x_51); +lean_ctor_set(x_38, 0, x_50); x_118 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_118, 0, x_34); -lean_ctor_set(x_118, 1, x_27); -lean_ctor_set(x_118, 2, x_36); -lean_ctor_set(x_118, 3, x_32); -lean_ctor_set(x_118, 4, x_31); -lean_ctor_set(x_118, 5, x_37); +lean_ctor_set(x_118, 0, x_38); +lean_ctor_set(x_118, 1, x_37); +lean_ctor_set(x_118, 2, x_29); +lean_ctor_set(x_118, 3, x_34); +lean_ctor_set(x_118, 4, x_32); +lean_ctor_set(x_118, 5, x_30); lean_ctor_set(x_118, 6, x_86); lean_ctor_set(x_118, 7, x_117); if (lean_is_scalar(x_116)) { @@ -71798,19 +64321,19 @@ else { uint8_t x_122; lean_dec(x_86); -lean_free_object(x_34); +lean_free_object(x_38); lean_dec_ref(x_54); lean_dec(x_53); lean_dec_ref(x_52); lean_dec(x_51); lean_dec(x_50); lean_dec_ref(x_40); -lean_dec_ref(x_37); -lean_dec_ref(x_36); -lean_dec_ref(x_35); +lean_dec_ref(x_39); +lean_dec(x_37); +lean_dec_ref(x_34); lean_dec_ref(x_32); -lean_dec_ref(x_31); -lean_dec(x_27); +lean_dec_ref(x_30); +lean_dec_ref(x_29); x_122 = !lean_is_exclusive(x_87); if (x_122 == 0) { @@ -71831,7 +64354,7 @@ return x_124; else { uint8_t x_125; -lean_free_object(x_34); +lean_free_object(x_38); lean_dec_ref(x_54); lean_dec(x_53); lean_dec_ref(x_52); @@ -71842,12 +64365,12 @@ lean_dec_ref(x_45); lean_dec(x_44); lean_dec_ref(x_43); lean_dec_ref(x_40); -lean_dec_ref(x_37); -lean_dec_ref(x_36); -lean_dec_ref(x_35); +lean_dec_ref(x_39); +lean_dec(x_37); +lean_dec_ref(x_34); lean_dec_ref(x_32); -lean_dec_ref(x_31); -lean_dec(x_27); +lean_dec_ref(x_30); +lean_dec_ref(x_29); lean_dec_ref(x_24); lean_dec_ref(x_7); x_125 = !lean_is_exclusive(x_76); @@ -71870,28 +64393,28 @@ return x_127; else { lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; -x_128 = lean_ctor_get(x_34, 2); +x_128 = lean_ctor_get(x_38, 2); lean_inc(x_128); -lean_dec(x_34); -lean_inc(x_38); -lean_inc(x_39); -x_129 = l_Array_toSubarray___redArg(x_23, x_39, x_38); +lean_dec(x_38); +lean_inc(x_25); +lean_inc(x_26); +x_129 = l_Array_toSubarray___redArg(x_23, x_26, x_25); x_130 = lean_array_get_size(x_52); -lean_inc(x_39); +lean_inc(x_26); lean_inc_ref(x_52); -x_131 = l_Array_toSubarray___redArg(x_52, x_39, x_130); +x_131 = l_Array_toSubarray___redArg(x_52, x_26, x_130); x_132 = lean_array_get_size(x_128); -lean_inc(x_39); -x_133 = l_Array_toSubarray___redArg(x_128, x_39, x_132); -x_134 = lean_array_get_size(x_28); -lean_inc(x_39); -x_135 = l_Array_toSubarray___redArg(x_28, x_39, x_134); +lean_inc(x_26); +x_133 = l_Array_toSubarray___redArg(x_128, x_26, x_132); +x_134 = lean_array_get_size(x_35); +lean_inc(x_26); +x_135 = l_Array_toSubarray___redArg(x_35, x_26, x_134); x_136 = lean_array_get_size(x_49); -lean_inc(x_39); -x_137 = l_Array_toSubarray___redArg(x_49, x_39, x_136); +lean_inc(x_26); +x_137 = l_Array_toSubarray___redArg(x_49, x_26, x_136); x_138 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_138, 0, x_137); -lean_ctor_set(x_138, 1, x_33); +lean_ctor_set(x_138, 1, x_27); x_139 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_139, 0, x_135); lean_ctor_set(x_139, 1, x_138); @@ -71908,10 +64431,10 @@ lean_inc(x_46); lean_inc_ref(x_45); lean_inc(x_44); lean_inc_ref(x_43); -lean_inc(x_39); -x_143 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__19___redArg(x_38, x_39, x_6, x_30, x_2, x_25, x_26, x_39, x_142, x_41, x_42, x_43, x_44, x_45, x_46); -lean_dec(x_39); -lean_dec(x_38); +lean_inc(x_26); +x_143 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__19___redArg(x_25, x_26, x_6, x_36, x_2, x_33, x_31, x_26, x_142, x_41, x_42, x_43, x_44, x_45, x_46); +lean_dec(x_26); +lean_dec(x_25); if (lean_obj_tag(x_143) == 0) { lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; @@ -71979,15 +64502,15 @@ lean_ctor_set(x_164, 0, x_50); lean_ctor_set(x_164, 1, x_51); lean_ctor_set(x_164, 2, x_52); lean_ctor_set(x_164, 3, x_53); -lean_ctor_set(x_164, 4, x_35); +lean_ctor_set(x_164, 4, x_39); lean_ctor_set(x_164, 5, x_54); x_165 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_165, 0, x_164); -lean_ctor_set(x_165, 1, x_27); -lean_ctor_set(x_165, 2, x_36); -lean_ctor_set(x_165, 3, x_32); -lean_ctor_set(x_165, 4, x_31); -lean_ctor_set(x_165, 5, x_37); +lean_ctor_set(x_165, 1, x_37); +lean_ctor_set(x_165, 2, x_29); +lean_ctor_set(x_165, 3, x_34); +lean_ctor_set(x_165, 4, x_32); +lean_ctor_set(x_165, 5, x_30); lean_ctor_set(x_165, 6, x_153); lean_ctor_set(x_165, 7, x_163); if (lean_is_scalar(x_162)) { @@ -72022,12 +64545,12 @@ lean_dec_ref(x_52); lean_dec(x_51); lean_dec(x_50); lean_dec_ref(x_40); -lean_dec_ref(x_37); -lean_dec_ref(x_36); -lean_dec_ref(x_35); +lean_dec_ref(x_39); +lean_dec(x_37); +lean_dec_ref(x_34); lean_dec_ref(x_32); -lean_dec_ref(x_31); -lean_dec(x_27); +lean_dec_ref(x_30); +lean_dec_ref(x_29); x_169 = lean_ctor_get(x_154, 0); lean_inc(x_169); if (lean_is_exclusive(x_154)) { @@ -72059,12 +64582,12 @@ lean_dec_ref(x_45); lean_dec(x_44); lean_dec_ref(x_43); lean_dec_ref(x_40); -lean_dec_ref(x_37); -lean_dec_ref(x_36); -lean_dec_ref(x_35); +lean_dec_ref(x_39); +lean_dec(x_37); +lean_dec_ref(x_34); lean_dec_ref(x_32); -lean_dec_ref(x_31); -lean_dec(x_27); +lean_dec_ref(x_30); +lean_dec_ref(x_29); lean_dec_ref(x_24); lean_dec_ref(x_7); x_172 = lean_ctor_get(x_143, 0); @@ -72096,17 +64619,17 @@ lean_dec_ref(x_43); lean_dec_ref(x_42); lean_dec_ref(x_41); lean_dec_ref(x_40); -lean_dec(x_39); -lean_dec(x_38); -lean_dec_ref(x_37); -lean_dec_ref(x_36); +lean_dec_ref(x_39); +lean_dec_ref(x_38); +lean_dec(x_37); lean_dec_ref(x_35); lean_dec_ref(x_34); -lean_dec_ref(x_33); +lean_dec(x_33); lean_dec_ref(x_32); -lean_dec_ref(x_31); -lean_dec_ref(x_28); -lean_dec(x_27); +lean_dec(x_31); +lean_dec_ref(x_30); +lean_dec_ref(x_29); +lean_dec_ref(x_27); lean_dec(x_26); lean_dec(x_25); lean_dec_ref(x_24); @@ -72139,7 +64662,7 @@ lean_inc(x_197); lean_inc_ref(x_196); lean_inc(x_195); lean_inc_ref(x_194); -x_200 = l_Lean_Meta_inferArgumentTypesN(x_199, x_187, x_194, x_195, x_196, x_197); +x_200 = l_Lean_Meta_inferArgumentTypesN(x_199, x_179, x_194, x_195, x_196, x_197); if (lean_obj_tag(x_200) == 0) { lean_object* x_201; lean_object* x_202; @@ -72164,10 +64687,10 @@ lean_inc_ref(x_205); lean_dec(x_203); lean_inc(x_204); x_206 = l_Lean_mkConst(x_204, x_182); -x_207 = l_Lean_mkAppN(x_206, x_183); -lean_inc_ref(x_181); -x_208 = l_Lean_Expr_app___override(x_207, x_181); -x_209 = l_Lean_mkAppN(x_208, x_189); +x_207 = l_Lean_mkAppN(x_206, x_190); +lean_inc_ref(x_189); +x_208 = l_Lean_Expr_app___override(x_207, x_189); +x_209 = l_Lean_mkAppN(x_208, x_186); lean_inc(x_197); lean_inc_ref(x_196); lean_inc(x_195); @@ -72221,21 +64744,21 @@ lean_dec(x_221); x_224 = lean_ctor_get(x_222, 1); lean_inc(x_224); lean_dec(x_222); -x_25 = x_179; -x_26 = x_180; -x_27 = x_204; -x_28 = x_201; -x_29 = x_209; -x_30 = x_188; -x_31 = x_181; -x_32 = x_183; -x_33 = x_184; -x_34 = x_205; -x_35 = x_186; -x_36 = x_185; -x_37 = x_189; -x_38 = x_199; -x_39 = x_190; +x_25 = x_199; +x_26 = x_181; +x_27 = x_180; +x_28 = x_209; +x_29 = x_184; +x_30 = x_186; +x_31 = x_185; +x_32 = x_189; +x_33 = x_188; +x_34 = x_190; +x_35 = x_201; +x_36 = x_183; +x_37 = x_204; +x_38 = x_205; +x_39 = x_187; x_40 = x_191; x_41 = x_224; x_42 = x_223; @@ -72258,15 +64781,15 @@ lean_dec_ref(x_196); lean_dec(x_195); lean_dec_ref(x_194); lean_dec_ref(x_191); -lean_dec(x_190); +lean_dec_ref(x_190); lean_dec_ref(x_189); +lean_dec(x_188); +lean_dec_ref(x_187); lean_dec_ref(x_186); -lean_dec_ref(x_185); +lean_dec(x_185); lean_dec_ref(x_184); -lean_dec_ref(x_183); -lean_dec_ref(x_181); -lean_dec(x_180); -lean_dec(x_179); +lean_dec(x_181); +lean_dec_ref(x_180); lean_dec_ref(x_24); lean_dec_ref(x_23); lean_dec_ref(x_17); @@ -72291,21 +64814,21 @@ return x_227; } else { -x_25 = x_179; -x_26 = x_180; -x_27 = x_204; -x_28 = x_201; -x_29 = x_209; -x_30 = x_188; -x_31 = x_181; -x_32 = x_183; -x_33 = x_184; -x_34 = x_205; -x_35 = x_186; -x_36 = x_185; -x_37 = x_189; -x_38 = x_199; -x_39 = x_190; +x_25 = x_199; +x_26 = x_181; +x_27 = x_180; +x_28 = x_209; +x_29 = x_184; +x_30 = x_186; +x_31 = x_185; +x_32 = x_189; +x_33 = x_188; +x_34 = x_190; +x_35 = x_201; +x_36 = x_183; +x_37 = x_204; +x_38 = x_205; +x_39 = x_187; x_40 = x_191; x_41 = x_192; x_42 = x_193; @@ -72331,15 +64854,15 @@ lean_dec_ref(x_194); lean_dec_ref(x_193); lean_dec_ref(x_192); lean_dec_ref(x_191); -lean_dec(x_190); +lean_dec_ref(x_190); lean_dec_ref(x_189); +lean_dec(x_188); +lean_dec_ref(x_187); lean_dec_ref(x_186); -lean_dec_ref(x_185); +lean_dec(x_185); lean_dec_ref(x_184); -lean_dec_ref(x_183); -lean_dec_ref(x_181); -lean_dec(x_180); -lean_dec(x_179); +lean_dec(x_181); +lean_dec_ref(x_180); lean_dec_ref(x_24); lean_dec_ref(x_23); lean_dec_ref(x_17); @@ -72373,16 +64896,16 @@ lean_dec_ref(x_194); lean_dec_ref(x_193); lean_dec_ref(x_192); lean_dec_ref(x_191); -lean_dec(x_190); +lean_dec_ref(x_190); lean_dec_ref(x_189); +lean_dec(x_188); +lean_dec_ref(x_187); lean_dec_ref(x_186); -lean_dec_ref(x_185); +lean_dec(x_185); lean_dec_ref(x_184); -lean_dec_ref(x_183); lean_dec(x_182); -lean_dec_ref(x_181); -lean_dec(x_180); -lean_dec(x_179); +lean_dec(x_181); +lean_dec_ref(x_180); lean_dec_ref(x_24); lean_dec_ref(x_23); lean_dec_ref(x_17); @@ -72415,16 +64938,16 @@ lean_dec_ref(x_194); lean_dec_ref(x_193); lean_dec_ref(x_192); lean_dec_ref(x_191); -lean_dec(x_190); +lean_dec_ref(x_190); lean_dec_ref(x_189); +lean_dec(x_188); +lean_dec_ref(x_187); lean_dec_ref(x_186); -lean_dec_ref(x_185); +lean_dec(x_185); lean_dec_ref(x_184); -lean_dec_ref(x_183); lean_dec(x_182); -lean_dec_ref(x_181); -lean_dec(x_180); -lean_dec(x_179); +lean_dec(x_181); +lean_dec_ref(x_180); lean_dec_ref(x_24); lean_dec_ref(x_23); lean_dec(x_18); @@ -72456,26 +64979,26 @@ lean_inc(x_255); lean_inc_ref(x_254); lean_inc(x_253); lean_inc_ref(x_252); -x_258 = l_Lean_Meta_inferArgumentTypesN(x_257, x_241, x_252, x_253, x_254, x_255); +x_258 = l_Lean_Meta_inferArgumentTypesN(x_257, x_240, x_252, x_253, x_254, x_255); if (lean_obj_tag(x_258) == 0) { lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; x_259 = lean_ctor_get(x_258, 0); lean_inc(x_259); lean_dec_ref(x_258); -lean_inc(x_248); +lean_inc(x_239); lean_inc_ref(x_23); -x_260 = l_Array_toSubarray___redArg(x_23, x_248, x_257); +x_260 = l_Array_toSubarray___redArg(x_23, x_239, x_257); x_261 = l_Lean_Meta_MatcherApp_altNumParams(x_1); x_262 = lean_array_get_size(x_261); -lean_inc(x_248); -x_263 = l_Array_toSubarray___redArg(x_261, x_248, x_262); +lean_inc(x_239); +x_263 = l_Array_toSubarray___redArg(x_261, x_239, x_262); x_264 = lean_array_get_size(x_259); -lean_inc(x_248); -x_265 = l_Array_toSubarray___redArg(x_259, x_248, x_264); +lean_inc(x_239); +x_265 = l_Array_toSubarray___redArg(x_259, x_239, x_264); x_266 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_266, 0, x_265); -lean_ctor_set(x_266, 1, x_243); +lean_ctor_set(x_266, 1, x_238); x_267 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_267, 0, x_263); lean_ctor_set(x_267, 1, x_266); @@ -72486,7 +65009,7 @@ lean_inc(x_255); lean_inc_ref(x_254); lean_inc(x_253); lean_inc_ref(x_252); -x_269 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__15___redArg(x_257, x_247, x_6, x_240, x_238, x_248, x_268, x_250, x_251, x_252, x_253, x_254, x_255); +x_269 = l_WellFounded_opaqueFix_u2083___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__15___redArg(x_257, x_241, x_6, x_247, x_246, x_239, x_268, x_250, x_251, x_252, x_253, x_254, x_255); if (lean_obj_tag(x_269) == 0) { lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; @@ -72539,10 +65062,10 @@ lean_ctor_set(x_17, 4, x_245); x_288 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_288, 0, x_17); lean_ctor_set(x_288, 1, x_18); -lean_ctor_set(x_288, 2, x_244); -lean_ctor_set(x_288, 3, x_242); -lean_ctor_set(x_288, 4, x_239); -lean_ctor_set(x_288, 5, x_246); +lean_ctor_set(x_288, 2, x_242); +lean_ctor_set(x_288, 3, x_248); +lean_ctor_set(x_288, 4, x_244); +lean_ctor_set(x_288, 5, x_243); lean_ctor_set(x_288, 6, x_277); lean_ctor_set(x_288, 7, x_287); lean_ctor_set(x_282, 0, x_288); @@ -72575,10 +65098,10 @@ lean_ctor_set(x_296, 5, x_294); x_297 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_297, 0, x_296); lean_ctor_set(x_297, 1, x_18); -lean_ctor_set(x_297, 2, x_244); -lean_ctor_set(x_297, 3, x_242); -lean_ctor_set(x_297, 4, x_239); -lean_ctor_set(x_297, 5, x_246); +lean_ctor_set(x_297, 2, x_242); +lean_ctor_set(x_297, 3, x_248); +lean_ctor_set(x_297, 4, x_244); +lean_ctor_set(x_297, 5, x_243); lean_ctor_set(x_297, 6, x_277); lean_ctor_set(x_297, 7, x_295); lean_ctor_set(x_282, 0, x_297); @@ -72631,10 +65154,10 @@ lean_ctor_set(x_307, 5, x_304); x_308 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_308, 0, x_307); lean_ctor_set(x_308, 1, x_18); -lean_ctor_set(x_308, 2, x_244); -lean_ctor_set(x_308, 3, x_242); -lean_ctor_set(x_308, 4, x_239); -lean_ctor_set(x_308, 5, x_246); +lean_ctor_set(x_308, 2, x_242); +lean_ctor_set(x_308, 3, x_248); +lean_ctor_set(x_308, 4, x_244); +lean_ctor_set(x_308, 5, x_243); lean_ctor_set(x_308, 6, x_277); lean_ctor_set(x_308, 7, x_306); x_309 = lean_alloc_ctor(0, 2, 0); @@ -72702,10 +65225,10 @@ lean_ctor_set(x_322, 5, x_319); x_323 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_323, 0, x_322); lean_ctor_set(x_323, 1, x_18); -lean_ctor_set(x_323, 2, x_244); -lean_ctor_set(x_323, 3, x_242); -lean_ctor_set(x_323, 4, x_239); -lean_ctor_set(x_323, 5, x_246); +lean_ctor_set(x_323, 2, x_242); +lean_ctor_set(x_323, 3, x_248); +lean_ctor_set(x_323, 4, x_244); +lean_ctor_set(x_323, 5, x_243); lean_ctor_set(x_323, 6, x_277); lean_ctor_set(x_323, 7, x_321); if (lean_is_scalar(x_314)) { @@ -72790,10 +65313,10 @@ lean_ctor_set(x_340, 5, x_337); x_341 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_341, 0, x_340); lean_ctor_set(x_341, 1, x_18); -lean_ctor_set(x_341, 2, x_244); -lean_ctor_set(x_341, 3, x_242); -lean_ctor_set(x_341, 4, x_239); -lean_ctor_set(x_341, 5, x_246); +lean_ctor_set(x_341, 2, x_242); +lean_ctor_set(x_341, 3, x_248); +lean_ctor_set(x_341, 4, x_244); +lean_ctor_set(x_341, 5, x_243); lean_ctor_set(x_341, 6, x_277); lean_ctor_set(x_341, 7, x_339); if (lean_is_scalar(x_332)) { @@ -72820,11 +65343,11 @@ else uint8_t x_345; lean_dec(x_277); lean_dec_ref(x_249); -lean_dec_ref(x_246); +lean_dec_ref(x_248); lean_dec_ref(x_245); lean_dec_ref(x_244); +lean_dec_ref(x_243); lean_dec_ref(x_242); -lean_dec_ref(x_239); lean_dec(x_18); lean_dec_ref(x_17); x_345 = !lean_is_exclusive(x_278); @@ -72852,11 +65375,11 @@ lean_dec_ref(x_254); lean_dec(x_253); lean_dec_ref(x_252); lean_dec_ref(x_249); -lean_dec_ref(x_246); +lean_dec_ref(x_248); lean_dec_ref(x_245); lean_dec_ref(x_244); +lean_dec_ref(x_243); lean_dec_ref(x_242); -lean_dec_ref(x_239); lean_dec_ref(x_24); lean_dec(x_18); lean_dec_ref(x_17); @@ -72888,14 +65411,14 @@ lean_dec_ref(x_252); lean_dec_ref(x_251); lean_dec_ref(x_250); lean_dec_ref(x_249); -lean_dec(x_248); -lean_dec_ref(x_246); +lean_dec_ref(x_248); +lean_dec(x_246); lean_dec_ref(x_245); lean_dec_ref(x_244); lean_dec_ref(x_243); lean_dec_ref(x_242); -lean_dec_ref(x_239); -lean_dec(x_238); +lean_dec(x_239); +lean_dec_ref(x_238); lean_dec_ref(x_24); lean_dec(x_18); lean_dec_ref(x_17); @@ -72922,20 +65445,20 @@ return x_353; block_399: { lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; -lean_inc_ref(x_367); -x_372 = lean_array_to_list(x_367); +lean_inc_ref(x_359); +x_372 = lean_array_to_list(x_359); lean_inc(x_18); x_373 = l_Lean_mkConst(x_18, x_372); -x_374 = l_Lean_mkAppN(x_373, x_364); -lean_inc_ref(x_362); -x_375 = l_Lean_Expr_app___override(x_374, x_362); -x_376 = l_Lean_mkAppN(x_375, x_370); -lean_inc(x_356); -lean_inc_ref(x_360); -lean_inc(x_357); +x_374 = l_Lean_mkAppN(x_373, x_363); +lean_inc_ref(x_361); +x_375 = l_Lean_Expr_app___override(x_374, x_361); +x_376 = l_Lean_mkAppN(x_375, x_360); +lean_inc(x_370); +lean_inc_ref(x_368); +lean_inc(x_365); lean_inc_ref(x_358); lean_inc_ref(x_376); -x_377 = l_Lean_Meta_isTypeCorrect(x_376, x_358, x_357, x_360, x_356); +x_377 = l_Lean_Meta_isTypeCorrect(x_376, x_358, x_365, x_368, x_370); if (lean_obj_tag(x_377) == 0) { lean_object* x_378; uint8_t x_379; uint8_t x_380; @@ -72954,8 +65477,8 @@ x_382 = l_Lean_indentExpr(x_376); x_383 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_383, 0, x_381); lean_ctor_set(x_383, 1, x_382); -lean_inc_ref(x_360); -x_384 = l_Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__16(x_383, x_369, x_366, x_358, x_357, x_360, x_356); +lean_inc_ref(x_368); +x_384 = l_Lean_logError___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__16(x_383, x_355, x_367, x_358, x_365, x_368, x_370); if (lean_obj_tag(x_384) == 0) { lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; @@ -72970,33 +65493,33 @@ lean_dec(x_385); x_388 = lean_ctor_get(x_386, 1); lean_inc(x_388); lean_dec(x_386); -lean_inc(x_356); -lean_inc_ref(x_360); -lean_inc(x_357); +lean_inc(x_370); +lean_inc_ref(x_368); +lean_inc(x_365); lean_inc_ref(x_358); lean_inc_ref(x_376); -x_389 = l_Lean_Meta_check(x_376, x_358, x_357, x_360, x_356); +x_389 = l_Lean_Meta_check(x_376, x_358, x_365, x_368, x_370); if (lean_obj_tag(x_389) == 0) { lean_dec_ref(x_389); -x_238 = x_355; -x_239 = x_362; -x_240 = x_379; -x_241 = x_376; -x_242 = x_364; -x_243 = x_365; -x_244 = x_367; -x_245 = x_368; -x_246 = x_370; -x_247 = x_361; -x_248 = x_371; -x_249 = x_363; +x_238 = x_356; +x_239 = x_357; +x_240 = x_376; +x_241 = x_369; +x_242 = x_359; +x_243 = x_360; +x_244 = x_361; +x_245 = x_371; +x_246 = x_362; +x_247 = x_379; +x_248 = x_363; +x_249 = x_366; x_250 = x_388; x_251 = x_387; x_252 = x_358; -x_253 = x_357; -x_254 = x_360; -x_255 = x_356; +x_253 = x_365; +x_254 = x_368; +x_255 = x_370; x_256 = lean_box(0); goto block_354; } @@ -73006,19 +65529,19 @@ uint8_t x_390; lean_dec(x_388); lean_dec(x_387); lean_dec_ref(x_376); -lean_dec(x_371); -lean_dec_ref(x_370); +lean_dec_ref(x_371); +lean_dec(x_370); lean_dec_ref(x_368); -lean_dec_ref(x_367); -lean_dec_ref(x_365); -lean_dec_ref(x_364); +lean_dec_ref(x_366); +lean_dec(x_365); lean_dec_ref(x_363); -lean_dec_ref(x_362); +lean_dec(x_362); +lean_dec_ref(x_361); lean_dec_ref(x_360); +lean_dec_ref(x_359); lean_dec_ref(x_358); lean_dec(x_357); -lean_dec(x_356); -lean_dec(x_355); +lean_dec_ref(x_356); lean_dec_ref(x_24); lean_dec(x_18); lean_dec_ref(x_17); @@ -73046,19 +65569,19 @@ else { uint8_t x_393; lean_dec_ref(x_376); -lean_dec(x_371); -lean_dec_ref(x_370); +lean_dec_ref(x_371); +lean_dec(x_370); lean_dec_ref(x_368); -lean_dec_ref(x_367); -lean_dec_ref(x_365); -lean_dec_ref(x_364); +lean_dec_ref(x_366); +lean_dec(x_365); lean_dec_ref(x_363); -lean_dec_ref(x_362); +lean_dec(x_362); +lean_dec_ref(x_361); lean_dec_ref(x_360); +lean_dec_ref(x_359); lean_dec_ref(x_358); lean_dec(x_357); -lean_dec(x_356); -lean_dec(x_355); +lean_dec_ref(x_356); lean_dec_ref(x_24); lean_dec(x_18); lean_dec_ref(x_17); @@ -73084,24 +65607,24 @@ return x_395; } else { -x_238 = x_355; -x_239 = x_362; -x_240 = x_379; -x_241 = x_376; -x_242 = x_364; -x_243 = x_365; -x_244 = x_367; -x_245 = x_368; -x_246 = x_370; -x_247 = x_361; -x_248 = x_371; -x_249 = x_363; -x_250 = x_369; -x_251 = x_366; +x_238 = x_356; +x_239 = x_357; +x_240 = x_376; +x_241 = x_369; +x_242 = x_359; +x_243 = x_360; +x_244 = x_361; +x_245 = x_371; +x_246 = x_362; +x_247 = x_379; +x_248 = x_363; +x_249 = x_366; +x_250 = x_355; +x_251 = x_367; x_252 = x_358; -x_253 = x_357; -x_254 = x_360; -x_255 = x_356; +x_253 = x_365; +x_254 = x_368; +x_255 = x_370; x_256 = lean_box(0); goto block_354; } @@ -73110,21 +65633,21 @@ else { uint8_t x_396; lean_dec_ref(x_376); -lean_dec(x_371); -lean_dec_ref(x_370); -lean_dec_ref(x_369); +lean_dec_ref(x_371); +lean_dec(x_370); lean_dec_ref(x_368); lean_dec_ref(x_367); lean_dec_ref(x_366); -lean_dec_ref(x_365); -lean_dec_ref(x_364); +lean_dec(x_365); lean_dec_ref(x_363); -lean_dec_ref(x_362); +lean_dec(x_362); +lean_dec_ref(x_361); lean_dec_ref(x_360); +lean_dec_ref(x_359); lean_dec_ref(x_358); lean_dec(x_357); -lean_dec(x_356); -lean_dec(x_355); +lean_dec_ref(x_356); +lean_dec_ref(x_355); lean_dec_ref(x_24); lean_dec(x_18); lean_dec_ref(x_17); @@ -73153,11 +65676,11 @@ return x_398; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; size_t x_425; lean_object* x_426; x_417 = lean_unsigned_to_nat(0u); x_418 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M_run___redArg___closed__0; -x_419 = l_Array_reverse___redArg(x_406); +x_419 = l_Array_reverse___redArg(x_403); x_420 = lean_array_get_size(x_419); x_421 = l_Array_toSubarray___redArg(x_419, x_417, x_420); -lean_inc_ref(x_408); -x_422 = l_Array_reverse___redArg(x_408); +lean_inc_ref(x_405); +x_422 = l_Array_reverse___redArg(x_405); x_423 = l_Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12___closed__6; x_424 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_424, 0, x_421); @@ -73167,7 +65690,7 @@ lean_inc(x_415); lean_inc_ref(x_414); lean_inc(x_413); lean_inc_ref(x_412); -x_426 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__13(x_422, x_425, x_403, x_424, x_410, x_411, x_412, x_413, x_414, x_415); +x_426 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5_spec__13(x_422, x_425, x_402, x_424, x_410, x_411, x_412, x_413, x_414, x_415); lean_dec_ref(x_422); if (lean_obj_tag(x_426) == 0) { @@ -73190,7 +65713,7 @@ if (x_2 == 0) { lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_free_object(x_429); -lean_dec(x_401); +lean_dec(x_404); x_433 = lean_ctor_get(x_427, 1); lean_inc(x_433); lean_dec(x_427); @@ -73202,23 +65725,23 @@ lean_inc(x_435); x_436 = lean_ctor_get(x_431, 1); lean_inc(x_436); lean_dec(x_431); -x_355 = x_435; -x_356 = x_415; -x_357 = x_413; +x_355 = x_434; +x_356 = x_418; +x_357 = x_417; x_358 = x_412; -x_359 = lean_box(0); -x_360 = x_414; +x_359 = x_409; +x_360 = x_405; x_361 = x_407; -x_362 = x_402; -x_363 = x_436; -x_364 = x_404; -x_365 = x_418; -x_366 = x_433; -x_367 = x_409; -x_368 = x_405; -x_369 = x_434; -x_370 = x_408; -x_371 = x_417; +x_362 = x_435; +x_363 = x_408; +x_364 = lean_box(0); +x_365 = x_413; +x_366 = x_436; +x_367 = x_433; +x_368 = x_414; +x_369 = x_401; +x_370 = x_415; +x_371 = x_406; goto block_399; } else @@ -73245,10 +65768,10 @@ x_442 = lean_array_to_list(x_409); lean_inc(x_442); lean_inc(x_18); x_443 = l_Lean_mkConst(x_18, x_442); -x_444 = l_Lean_mkAppN(x_443, x_404); -lean_inc_ref(x_402); -x_445 = l_Lean_Expr_app___override(x_444, x_402); -x_446 = l_Lean_mkAppN(x_445, x_408); +x_444 = l_Lean_mkAppN(x_443, x_408); +lean_inc_ref(x_407); +x_445 = l_Lean_Expr_app___override(x_444, x_407); +x_446 = l_Lean_mkAppN(x_445, x_405); lean_inc(x_415); lean_inc_ref(x_414); lean_inc(x_413); @@ -73302,18 +65825,18 @@ lean_dec(x_456); x_459 = lean_ctor_get(x_457, 1); lean_inc(x_459); lean_dec(x_457); -x_179 = x_440; -x_180 = x_401; -x_181 = x_402; +x_179 = x_446; +x_180 = x_418; +x_181 = x_417; x_182 = x_442; -x_183 = x_404; -x_184 = x_418; -x_185 = x_409; +x_183 = x_401; +x_184 = x_409; +x_185 = x_404; x_186 = x_405; -x_187 = x_446; -x_188 = x_407; -x_189 = x_408; -x_190 = x_417; +x_187 = x_406; +x_188 = x_440; +x_189 = x_407; +x_190 = x_408; x_191 = x_441; x_192 = x_459; x_193 = x_458; @@ -73337,10 +65860,10 @@ lean_dec(x_413); lean_dec_ref(x_412); lean_dec_ref(x_409); lean_dec_ref(x_408); +lean_dec_ref(x_407); +lean_dec_ref(x_406); lean_dec_ref(x_405); -lean_dec_ref(x_404); -lean_dec_ref(x_402); -lean_dec(x_401); +lean_dec(x_404); lean_dec_ref(x_24); lean_dec_ref(x_23); lean_dec(x_18); @@ -73368,18 +65891,18 @@ else { lean_free_object(x_431); lean_free_object(x_429); -x_179 = x_440; -x_180 = x_401; -x_181 = x_402; +x_179 = x_446; +x_180 = x_418; +x_181 = x_417; x_182 = x_442; -x_183 = x_404; -x_184 = x_418; -x_185 = x_409; +x_183 = x_401; +x_184 = x_409; +x_185 = x_404; x_186 = x_405; -x_187 = x_446; -x_188 = x_407; -x_189 = x_408; -x_190 = x_417; +x_187 = x_406; +x_188 = x_440; +x_189 = x_407; +x_190 = x_408; x_191 = x_441; x_192 = x_438; x_193 = x_437; @@ -73408,10 +65931,10 @@ lean_dec(x_413); lean_dec_ref(x_412); lean_dec_ref(x_409); lean_dec_ref(x_408); +lean_dec_ref(x_407); +lean_dec_ref(x_406); lean_dec_ref(x_405); -lean_dec_ref(x_404); -lean_dec_ref(x_402); -lean_dec(x_401); +lean_dec(x_404); lean_dec_ref(x_24); lean_dec_ref(x_23); lean_dec(x_18); @@ -73448,10 +65971,10 @@ x_468 = lean_array_to_list(x_409); lean_inc(x_468); lean_inc(x_18); x_469 = l_Lean_mkConst(x_18, x_468); -x_470 = l_Lean_mkAppN(x_469, x_404); -lean_inc_ref(x_402); -x_471 = l_Lean_Expr_app___override(x_470, x_402); -x_472 = l_Lean_mkAppN(x_471, x_408); +x_470 = l_Lean_mkAppN(x_469, x_408); +lean_inc_ref(x_407); +x_471 = l_Lean_Expr_app___override(x_470, x_407); +x_472 = l_Lean_mkAppN(x_471, x_405); lean_inc(x_415); lean_inc_ref(x_414); lean_inc(x_413); @@ -73505,18 +66028,18 @@ lean_dec(x_483); x_486 = lean_ctor_get(x_484, 1); lean_inc(x_486); lean_dec(x_484); -x_179 = x_466; -x_180 = x_401; -x_181 = x_402; +x_179 = x_472; +x_180 = x_418; +x_181 = x_417; x_182 = x_468; -x_183 = x_404; -x_184 = x_418; -x_185 = x_409; +x_183 = x_401; +x_184 = x_409; +x_185 = x_404; x_186 = x_405; -x_187 = x_472; -x_188 = x_407; -x_189 = x_408; -x_190 = x_417; +x_187 = x_406; +x_188 = x_466; +x_189 = x_407; +x_190 = x_408; x_191 = x_467; x_192 = x_486; x_193 = x_485; @@ -73540,10 +66063,10 @@ lean_dec(x_413); lean_dec_ref(x_412); lean_dec_ref(x_409); lean_dec_ref(x_408); +lean_dec_ref(x_407); +lean_dec_ref(x_406); lean_dec_ref(x_405); -lean_dec_ref(x_404); -lean_dec_ref(x_402); -lean_dec(x_401); +lean_dec(x_404); lean_dec_ref(x_24); lean_dec_ref(x_23); lean_dec(x_18); @@ -73571,18 +66094,18 @@ return x_489; else { lean_free_object(x_429); -x_179 = x_466; -x_180 = x_401; -x_181 = x_402; +x_179 = x_472; +x_180 = x_418; +x_181 = x_417; x_182 = x_468; -x_183 = x_404; -x_184 = x_418; -x_185 = x_409; +x_183 = x_401; +x_184 = x_409; +x_185 = x_404; x_186 = x_405; -x_187 = x_472; -x_188 = x_407; -x_189 = x_408; -x_190 = x_417; +x_187 = x_406; +x_188 = x_466; +x_189 = x_407; +x_190 = x_408; x_191 = x_467; x_192 = x_438; x_193 = x_437; @@ -73610,10 +66133,10 @@ lean_dec(x_413); lean_dec_ref(x_412); lean_dec_ref(x_409); lean_dec_ref(x_408); +lean_dec_ref(x_407); +lean_dec_ref(x_406); lean_dec_ref(x_405); -lean_dec_ref(x_404); -lean_dec_ref(x_402); -lean_dec(x_401); +lean_dec(x_404); lean_dec_ref(x_24); lean_dec_ref(x_23); lean_dec(x_18); @@ -73643,7 +66166,7 @@ else { lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_free_object(x_429); -lean_dec(x_401); +lean_dec(x_404); x_493 = lean_ctor_get(x_427, 1); lean_inc(x_493); lean_dec(x_427); @@ -73655,23 +66178,23 @@ lean_inc(x_495); x_496 = lean_ctor_get(x_431, 1); lean_inc(x_496); lean_dec(x_431); -x_355 = x_495; -x_356 = x_415; -x_357 = x_413; +x_355 = x_494; +x_356 = x_418; +x_357 = x_417; x_358 = x_412; -x_359 = lean_box(0); -x_360 = x_414; +x_359 = x_409; +x_360 = x_405; x_361 = x_407; -x_362 = x_402; -x_363 = x_496; -x_364 = x_404; -x_365 = x_418; -x_366 = x_493; -x_367 = x_409; -x_368 = x_405; -x_369 = x_494; -x_370 = x_408; -x_371 = x_417; +x_362 = x_495; +x_363 = x_408; +x_364 = lean_box(0); +x_365 = x_413; +x_366 = x_496; +x_367 = x_493; +x_368 = x_414; +x_369 = x_401; +x_370 = x_415; +x_371 = x_406; goto block_399; } } @@ -73685,7 +66208,7 @@ lean_dec(x_429); if (x_2 == 0) { lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; -lean_dec(x_401); +lean_dec(x_404); x_498 = lean_ctor_get(x_427, 1); lean_inc(x_498); lean_dec(x_427); @@ -73697,23 +66220,23 @@ lean_inc(x_500); x_501 = lean_ctor_get(x_497, 1); lean_inc(x_501); lean_dec(x_497); -x_355 = x_500; -x_356 = x_415; -x_357 = x_413; +x_355 = x_499; +x_356 = x_418; +x_357 = x_417; x_358 = x_412; -x_359 = lean_box(0); -x_360 = x_414; +x_359 = x_409; +x_360 = x_405; x_361 = x_407; -x_362 = x_402; -x_363 = x_501; -x_364 = x_404; -x_365 = x_418; -x_366 = x_498; -x_367 = x_409; -x_368 = x_405; -x_369 = x_499; -x_370 = x_408; -x_371 = x_417; +x_362 = x_500; +x_363 = x_408; +x_364 = lean_box(0); +x_365 = x_413; +x_366 = x_501; +x_367 = x_498; +x_368 = x_414; +x_369 = x_401; +x_370 = x_415; +x_371 = x_406; goto block_399; } else @@ -73746,10 +66269,10 @@ x_507 = lean_array_to_list(x_409); lean_inc(x_507); lean_inc(x_18); x_508 = l_Lean_mkConst(x_18, x_507); -x_509 = l_Lean_mkAppN(x_508, x_404); -lean_inc_ref(x_402); -x_510 = l_Lean_Expr_app___override(x_509, x_402); -x_511 = l_Lean_mkAppN(x_510, x_408); +x_509 = l_Lean_mkAppN(x_508, x_408); +lean_inc_ref(x_407); +x_510 = l_Lean_Expr_app___override(x_509, x_407); +x_511 = l_Lean_mkAppN(x_510, x_405); lean_inc(x_415); lean_inc_ref(x_414); lean_inc(x_413); @@ -73808,18 +66331,18 @@ lean_dec(x_523); x_526 = lean_ctor_get(x_524, 1); lean_inc(x_526); lean_dec(x_524); -x_179 = x_504; -x_180 = x_401; -x_181 = x_402; +x_179 = x_511; +x_180 = x_418; +x_181 = x_417; x_182 = x_507; -x_183 = x_404; -x_184 = x_418; -x_185 = x_409; +x_183 = x_401; +x_184 = x_409; +x_185 = x_404; x_186 = x_405; -x_187 = x_511; -x_188 = x_407; -x_189 = x_408; -x_190 = x_417; +x_187 = x_406; +x_188 = x_504; +x_189 = x_407; +x_190 = x_408; x_191 = x_505; x_192 = x_526; x_193 = x_525; @@ -73843,10 +66366,10 @@ lean_dec(x_413); lean_dec_ref(x_412); lean_dec_ref(x_409); lean_dec_ref(x_408); +lean_dec_ref(x_407); +lean_dec_ref(x_406); lean_dec_ref(x_405); -lean_dec_ref(x_404); -lean_dec_ref(x_402); -lean_dec(x_401); +lean_dec(x_404); lean_dec_ref(x_24); lean_dec_ref(x_23); lean_dec(x_18); @@ -73874,18 +66397,18 @@ return x_529; else { lean_dec(x_506); -x_179 = x_504; -x_180 = x_401; -x_181 = x_402; +x_179 = x_511; +x_180 = x_418; +x_181 = x_417; x_182 = x_507; -x_183 = x_404; -x_184 = x_418; -x_185 = x_409; +x_183 = x_401; +x_184 = x_409; +x_185 = x_404; x_186 = x_405; -x_187 = x_511; -x_188 = x_407; -x_189 = x_408; -x_190 = x_417; +x_187 = x_406; +x_188 = x_504; +x_189 = x_407; +x_190 = x_408; x_191 = x_505; x_192 = x_503; x_193 = x_502; @@ -73913,10 +66436,10 @@ lean_dec(x_413); lean_dec_ref(x_412); lean_dec_ref(x_409); lean_dec_ref(x_408); +lean_dec_ref(x_407); +lean_dec_ref(x_406); lean_dec_ref(x_405); -lean_dec_ref(x_404); -lean_dec_ref(x_402); -lean_dec(x_401); +lean_dec(x_404); lean_dec_ref(x_24); lean_dec_ref(x_23); lean_dec(x_18); @@ -73944,7 +66467,7 @@ return x_532; else { lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; -lean_dec(x_401); +lean_dec(x_404); x_533 = lean_ctor_get(x_427, 1); lean_inc(x_533); lean_dec(x_427); @@ -73956,23 +66479,23 @@ lean_inc(x_535); x_536 = lean_ctor_get(x_497, 1); lean_inc(x_536); lean_dec(x_497); -x_355 = x_535; -x_356 = x_415; -x_357 = x_413; +x_355 = x_534; +x_356 = x_418; +x_357 = x_417; x_358 = x_412; -x_359 = lean_box(0); -x_360 = x_414; +x_359 = x_409; +x_360 = x_405; x_361 = x_407; -x_362 = x_402; -x_363 = x_536; -x_364 = x_404; -x_365 = x_418; -x_366 = x_533; -x_367 = x_409; -x_368 = x_405; -x_369 = x_534; -x_370 = x_408; -x_371 = x_417; +x_362 = x_535; +x_363 = x_408; +x_364 = lean_box(0); +x_365 = x_413; +x_366 = x_536; +x_367 = x_533; +x_368 = x_414; +x_369 = x_401; +x_370 = x_415; +x_371 = x_406; goto block_399; } } @@ -73987,10 +66510,10 @@ lean_dec(x_413); lean_dec_ref(x_412); lean_dec_ref(x_409); lean_dec_ref(x_408); +lean_dec_ref(x_407); +lean_dec_ref(x_406); lean_dec_ref(x_405); -lean_dec_ref(x_404); -lean_dec_ref(x_402); -lean_dec(x_401); +lean_dec(x_404); lean_dec_ref(x_24); lean_dec(x_18); lean_dec_ref(x_17); @@ -74116,14 +66639,14 @@ x_579 = lean_ctor_get(x_573, 1); lean_inc(x_579); lean_dec(x_573); lean_inc_ref(x_19); -x_401 = x_541; -x_402 = x_577; -x_403 = x_550; -x_404 = x_555; -x_405 = x_579; -x_406 = x_578; -x_407 = x_567; -x_408 = x_562; +x_401 = x_567; +x_402 = x_550; +x_403 = x_578; +x_404 = x_541; +x_405 = x_562; +x_406 = x_579; +x_407 = x_577; +x_408 = x_555; x_409 = x_19; x_410 = x_576; x_411 = x_575; @@ -74158,14 +66681,14 @@ lean_dec(x_573); x_586 = lean_ctor_get(x_574, 0); lean_inc_ref(x_19); x_587 = lean_array_set(x_19, x_586, x_583); -x_401 = x_541; -x_402 = x_582; -x_403 = x_550; -x_404 = x_555; -x_405 = x_585; -x_406 = x_584; -x_407 = x_567; -x_408 = x_562; +x_401 = x_567; +x_402 = x_550; +x_403 = x_584; +x_404 = x_541; +x_405 = x_562; +x_406 = x_585; +x_407 = x_582; +x_408 = x_555; x_409 = x_587; x_410 = x_581; x_411 = x_580; @@ -78714,7 +71237,7 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_Fun lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__17___closed__2; x_2 = lean_unsigned_to_nat(40u); -x_3 = lean_unsigned_to_nat(818u); +x_3 = lean_unsigned_to_nat(699u); x_4 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__13___closed__0; x_5 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__17___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -78867,7 +71390,7 @@ lean_closure_set(x_34, 2, x_2); lean_closure_set(x_34, 3, x_26); lean_closure_set(x_34, 4, x_6); lean_closure_set(x_34, 5, x_7); -x_35 = lean_alloc_closure((void*)(l_Lean_Tactic_FunInd_rwMatcher___boxed), 7, 1); +x_35 = lean_alloc_closure((void*)(l_Lean_Meta_rwMatcher___boxed), 7, 1); lean_closure_set(x_35, 0, x_8); lean_inc(x_18); lean_inc_ref(x_17); @@ -79465,6 +71988,57 @@ x_3 = l_Lean_Name_mkStr2(x_2, x_1); return x_3; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__19() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("cond", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__19; +x_2 = l_Lean_Name_mkStr1(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__21() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("dite", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__21; +x_2 = l_Lean_Name_mkStr1(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__23() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ite", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__23; +x_2 = l_Lean_Name_mkStr1(x_1); +return x_2; +} +} LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__18(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { @@ -79742,7 +72316,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunI { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_inc_ref(x_5); -x_13 = lean_alloc_closure((void*)(l_Lean_Tactic_FunInd_rwIfWith___boxed), 7, 1); +x_13 = lean_alloc_closure((void*)(l_Lean_Meta_rwIfWith___boxed), 7, 1); lean_closure_set(x_13, 0, x_5); x_14 = lean_box(x_3); x_15 = lean_box(x_4); @@ -79767,7 +72341,7 @@ x_15 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionB return x_15; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__19() { +static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__25() { _start: { lean_object* x_1; @@ -79775,11 +72349,11 @@ x_1 = lean_mk_string_unchecked("h", 1, 1); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__20() { +static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__26() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__19; +x_1 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__25; x_2 = l_Lean_Name_mkStr1(x_1); return x_2; } @@ -79822,7 +72396,7 @@ lean_closure_set(x_21, 2, x_5); lean_closure_set(x_21, 3, x_6); lean_closure_set(x_21, 4, x_7); lean_closure_set(x_21, 5, x_20); -x_22 = lean_alloc_closure((void*)(l_Lean_Tactic_FunInd_rwIfWith___boxed), 7, 1); +x_22 = lean_alloc_closure((void*)(l_Lean_Meta_rwIfWith___boxed), 7, 1); lean_closure_set(x_22, 0, x_8); lean_inc(x_17); lean_inc_ref(x_16); @@ -80170,6 +72744,42 @@ x_20 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionB return x_20; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__28() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("true", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__27() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Bool", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__29() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__28; +x_2 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__27; +x_3 = l_Lean_Name_mkStr2(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__30() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__29; +x_3 = l_Lean_mkConst(x_2, x_1); +return x_3; +} +} LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__27(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { @@ -80433,7 +73043,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunI { lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_inc_ref(x_4); -x_12 = lean_alloc_closure((void*)(l_Lean_Tactic_FunInd_rwIfWith___boxed), 7, 1); +x_12 = lean_alloc_closure((void*)(l_Lean_Meta_rwIfWith___boxed), 7, 1); lean_closure_set(x_12, 0, x_4); x_13 = lean_box(x_3); x_14 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__27___boxed), 12, 5); @@ -80455,7 +73065,35 @@ x_13 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionB return x_13; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__21() { +static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__31() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("false", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__32() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__31; +x_2 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__27; +x_3 = l_Lean_Name_mkStr2(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__33() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__32; +x_3 = l_Lean_mkConst(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__34() { _start: { lean_object* x_1; @@ -80463,12 +73101,12 @@ x_1 = lean_mk_string_unchecked("dcond", 5, 5); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__22() { +static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__35() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__21; -x_2 = l_Lean_Tactic_FunInd_rwIfWith___closed__15; +x_1 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__34; +x_2 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__27; x_3 = l_Lean_Name_mkStr2(x_2, x_1); return x_3; } @@ -80476,7 +73114,7 @@ return x_3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20) { _start: { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_110; uint8_t x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; uint8_t x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; uint8_t x_123; uint8_t x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; uint8_t x_260; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; uint8_t x_424; lean_object* x_450; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_110; lean_object* x_111; uint8_t x_112; lean_object* x_113; uint8_t x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; uint8_t x_123; uint8_t x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; uint8_t x_260; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; uint8_t x_424; lean_object* x_450; lean_inc_ref(x_1); x_450 = l_Lean_Meta_instantiateMVarsIfMVarApp___redArg(x_1, x_18); if (lean_obj_tag(x_450) == 0) @@ -80569,7 +73207,7 @@ lean_inc_ref(x_474); lean_dec_ref(x_468); lean_inc_ref(x_470); x_475 = l_Lean_Expr_appFnCleanup___redArg(x_470); -x_476 = l_Lean_Tactic_FunInd_rwIfWith___closed__1; +x_476 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__20; x_477 = l_Lean_Expr_isConstOf(x_475, x_476); if (x_477 == 0) { @@ -80597,12 +73235,12 @@ x_479 = lean_ctor_get(x_470, 1); lean_inc_ref(x_479); lean_dec_ref(x_470); x_480 = l_Lean_Expr_appFnCleanup___redArg(x_475); -x_481 = l_Lean_Tactic_FunInd_rwIfWith___closed__3; +x_481 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__22; x_482 = l_Lean_Expr_isConstOf(x_480, x_481); if (x_482 == 0) { lean_object* x_483; uint8_t x_484; -x_483 = l_Lean_Tactic_FunInd_rwIfWith___closed__5; +x_483 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__24; x_484 = l_Lean_Expr_isConstOf(x_480, x_483); lean_dec_ref(x_480); if (x_484 == 0) @@ -80688,7 +73326,7 @@ lean_closure_set(x_496, 0, x_2); lean_closure_set(x_496, 1, x_493); lean_closure_set(x_496, 2, x_494); lean_closure_set(x_496, 3, x_495); -x_497 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__20; +x_497 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__26; x_498 = 0; x_499 = 0; lean_inc(x_20); @@ -81298,7 +73936,7 @@ lean_closure_set(x_609, 5, x_5); lean_closure_set(x_609, 6, x_2); lean_closure_set(x_609, 7, x_607); lean_closure_set(x_609, 8, x_608); -x_610 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__20; +x_610 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__26; x_611 = 0; x_612 = 0; lean_inc(x_20); @@ -81870,7 +74508,7 @@ x_713 = lean_ctor_get(x_711, 1); lean_inc(x_713); lean_dec(x_711); x_714 = lean_box(0); -x_715 = l_Lean_Tactic_FunInd_rwIfWith___closed__18; +x_715 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__30; lean_inc(x_20); lean_inc_ref(x_19); lean_inc(x_18); @@ -81901,7 +74539,7 @@ x_720 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_FunInd_0__Lean_T lean_closure_set(x_720, 0, x_2); lean_closure_set(x_720, 1, x_718); lean_closure_set(x_720, 2, x_719); -x_721 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__20; +x_721 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__26; x_722 = 0; x_723 = 0; lean_inc(x_20); @@ -81926,7 +74564,7 @@ if (x_728 == 0) lean_object* x_729; lean_object* x_730; lean_object* x_731; lean_object* x_732; x_729 = lean_ctor_get(x_726, 0); x_730 = lean_ctor_get(x_726, 1); -x_731 = l_Lean_Tactic_FunInd_rwIfWith___closed__21; +x_731 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__33; lean_inc(x_20); lean_inc_ref(x_19); lean_inc(x_18); @@ -81985,7 +74623,7 @@ if (x_746 == 0) { lean_object* x_747; lean_object* x_748; lean_object* x_749; lean_object* x_750; x_747 = lean_ctor_get(x_745, 0); -x_748 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__22; +x_748 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__35; lean_ctor_set_tag(x_726, 1); lean_ctor_set(x_726, 1, x_714); lean_ctor_set(x_726, 0, x_747); @@ -82001,7 +74639,7 @@ lean_object* x_751; lean_object* x_752; lean_object* x_753; lean_object* x_754; x_751 = lean_ctor_get(x_745, 0); lean_inc(x_751); lean_dec(x_745); -x_752 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__22; +x_752 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__35; lean_ctor_set_tag(x_726, 1); lean_ctor_set(x_726, 1, x_714); lean_ctor_set(x_726, 0, x_751); @@ -82065,7 +74703,7 @@ if (lean_is_exclusive(x_762)) { lean_dec_ref(x_762); x_764 = lean_box(0); } -x_765 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__22; +x_765 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__35; lean_ctor_set_tag(x_726, 1); lean_ctor_set(x_726, 1, x_714); lean_ctor_set(x_726, 0, x_763); @@ -82147,7 +74785,7 @@ if (lean_is_exclusive(x_778)) { lean_dec_ref(x_778); x_780 = lean_box(0); } -x_781 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__22; +x_781 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__35; lean_ctor_set_tag(x_726, 1); lean_ctor_set(x_726, 1, x_714); lean_ctor_set(x_726, 0, x_779); @@ -82258,7 +74896,7 @@ x_794 = lean_ctor_get(x_726, 1); lean_inc(x_794); lean_inc(x_793); lean_dec(x_726); -x_795 = l_Lean_Tactic_FunInd_rwIfWith___closed__21; +x_795 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__33; lean_inc(x_20); lean_inc_ref(x_19); lean_inc(x_18); @@ -82333,7 +74971,7 @@ if (lean_is_exclusive(x_809)) { lean_dec_ref(x_809); x_811 = lean_box(0); } -x_812 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__22; +x_812 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__35; x_813 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_813, 0, x_810); lean_ctor_set(x_813, 1, x_714); @@ -82535,13 +75173,13 @@ x_458 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInduction x_459 = l_Lean_Expr_isAppOf(x_1, x_458); if (x_459 == 0) { -x_417 = x_453; -x_418 = x_454; -x_419 = x_456; -x_420 = x_457; -x_421 = x_455; -x_422 = lean_box(0); -x_423 = x_452; +x_417 = x_454; +x_418 = x_452; +x_419 = x_453; +x_420 = lean_box(0); +x_421 = x_457; +x_422 = x_455; +x_423 = x_456; x_424 = x_459; goto block_449; } @@ -82552,13 +75190,13 @@ x_460 = lean_unsigned_to_nat(1u); x_461 = l_Lean_Expr_getAppNumArgs(x_1); x_462 = lean_nat_dec_lt(x_460, x_461); lean_dec(x_461); -x_417 = x_453; -x_418 = x_454; -x_419 = x_456; -x_420 = x_457; -x_421 = x_455; -x_422 = lean_box(0); -x_423 = x_452; +x_417 = x_454; +x_418 = x_452; +x_419 = x_453; +x_420 = lean_box(0); +x_421 = x_457; +x_422 = x_455; +x_423 = x_456; x_424 = x_462; goto block_449; } @@ -82625,7 +75263,7 @@ return x_32; lean_object* x_38; lean_object* x_39; lean_object* x_40; x_38 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_38, 1, x_36); +lean_ctor_set(x_38, 1, x_34); x_39 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_39, 0, x_38); lean_ctor_set(x_39, 1, x_35); @@ -82912,18 +75550,18 @@ lean_object* x_104; lean_inc(x_99); lean_dec_ref(x_1); x_104 = l_Lean_Expr_mdata___override(x_99, x_97); -x_34 = lean_box(0); +x_34 = x_98; x_35 = x_96; -x_36 = x_98; +x_36 = lean_box(0); x_37 = x_104; goto block_41; } else { lean_dec(x_97); -x_34 = lean_box(0); +x_34 = x_98; x_35 = x_96; -x_36 = x_98; +x_36 = lean_box(0); x_37 = x_1; goto block_41; } @@ -82940,9 +75578,9 @@ lean_inc(x_106); lean_dec(x_95); x_107 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__6; x_108 = l_panic___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__3(x_107); -x_34 = lean_box(0); +x_34 = x_106; x_35 = x_105; -x_36 = x_106; +x_36 = lean_box(0); x_37 = x_108; goto block_41; } @@ -82962,22 +75600,22 @@ uint8_t x_124; lean_dec_ref(x_110); lean_dec_ref(x_13); lean_dec_ref(x_12); -x_124 = l_Array_isEmpty___redArg(x_113); -lean_dec_ref(x_113); +x_124 = l_Array_isEmpty___redArg(x_111); +lean_dec_ref(x_111); if (x_124 == 0) { lean_dec_ref(x_117); -lean_dec_ref(x_112); +lean_dec_ref(x_113); lean_dec_ref(x_11); lean_dec_ref(x_10); lean_dec_ref(x_9); -x_42 = x_111; -x_43 = x_120; -x_44 = x_119; -x_45 = x_114; +x_42 = x_112; +x_43 = x_119; +x_44 = x_118; +x_45 = x_116; x_46 = x_122; -x_47 = x_121; -x_48 = x_116; +x_47 = x_115; +x_48 = x_120; x_49 = lean_box(0); goto block_109; } @@ -82991,8 +75629,8 @@ lean_dec_ref(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec_ref(x_1); -lean_inc_ref(x_114); -x_125 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_mkLambdaFVarsMasked(x_112, x_2, x_114, x_122, x_121, x_116); +lean_inc_ref(x_116); +x_125 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_mkLambdaFVarsMasked(x_113, x_2, x_116, x_122, x_115, x_120); if (lean_obj_tag(x_125) == 0) { lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; @@ -83007,7 +75645,7 @@ lean_dec(x_126); x_129 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__16___boxed), 11, 2); lean_closure_set(x_129, 0, x_127); lean_closure_set(x_129, 1, x_128); -x_130 = l_Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5(x_117, x_115, x_115, x_9, x_129, x_10, x_11, x_120, x_119, x_114, x_122, x_121, x_116); +x_130 = l_Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5(x_117, x_114, x_114, x_9, x_129, x_10, x_11, x_119, x_118, x_116, x_122, x_115, x_120); if (lean_obj_tag(x_130) == 0) { uint8_t x_131; @@ -83155,12 +75793,12 @@ else { uint8_t x_164; lean_dec(x_122); -lean_dec_ref(x_121); -lean_dec_ref(x_120); +lean_dec(x_120); lean_dec_ref(x_119); +lean_dec_ref(x_118); lean_dec_ref(x_117); -lean_dec(x_116); -lean_dec_ref(x_114); +lean_dec_ref(x_116); +lean_dec_ref(x_115); lean_dec_ref(x_11); lean_dec_ref(x_10); lean_dec_ref(x_9); @@ -83185,7 +75823,7 @@ return x_166; else { lean_object* x_167; -lean_dec_ref(x_113); +lean_dec_ref(x_111); lean_dec_ref(x_11); lean_dec_ref(x_10); lean_dec_ref(x_9); @@ -83195,25 +75833,25 @@ lean_dec_ref(x_6); lean_dec_ref(x_5); lean_dec(x_3); lean_dec_ref(x_1); -lean_inc_ref(x_114); -x_167 = l_Lean_FVarId_getType___redArg(x_4, x_114, x_121, x_116); +lean_inc_ref(x_116); +x_167 = l_Lean_FVarId_getType___redArg(x_4, x_116, x_115, x_120); if (lean_obj_tag(x_167) == 0) { lean_object* x_168; lean_object* x_169; x_168 = lean_ctor_get(x_167, 0); lean_inc(x_168); lean_dec_ref(x_167); -lean_inc(x_116); -lean_inc_ref(x_121); -x_169 = l_Lean_mkArrow(x_168, x_2, x_121, x_116); +lean_inc(x_120); +lean_inc_ref(x_115); +x_169 = l_Lean_mkArrow(x_168, x_2, x_115, x_120); if (lean_obj_tag(x_169) == 0) { lean_object* x_170; lean_object* x_171; x_170 = lean_ctor_get(x_169, 0); lean_inc(x_170); lean_dec_ref(x_169); -lean_inc_ref(x_114); -x_171 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_mkLambdaFVarsMasked(x_112, x_170, x_114, x_122, x_121, x_116); +lean_inc_ref(x_116); +x_171 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_mkLambdaFVarsMasked(x_113, x_170, x_116, x_122, x_115, x_120); if (lean_obj_tag(x_171) == 0) { lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; @@ -83228,7 +75866,7 @@ lean_dec(x_172); x_175 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__16___boxed), 11, 2); lean_closure_set(x_175, 0, x_173); lean_closure_set(x_175, 1, x_174); -x_176 = l_Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5(x_117, x_115, x_115, x_12, x_175, x_110, x_13, x_120, x_119, x_114, x_122, x_121, x_116); +x_176 = l_Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__5(x_117, x_114, x_114, x_12, x_175, x_110, x_13, x_119, x_118, x_116, x_122, x_115, x_120); if (lean_obj_tag(x_176) == 0) { uint8_t x_177; @@ -83376,12 +76014,12 @@ else { uint8_t x_210; lean_dec(x_122); -lean_dec_ref(x_121); -lean_dec_ref(x_120); +lean_dec(x_120); lean_dec_ref(x_119); +lean_dec_ref(x_118); lean_dec_ref(x_117); -lean_dec(x_116); -lean_dec_ref(x_114); +lean_dec_ref(x_116); +lean_dec_ref(x_115); lean_dec_ref(x_110); lean_dec_ref(x_13); lean_dec_ref(x_12); @@ -83406,13 +76044,13 @@ else { uint8_t x_213; lean_dec(x_122); -lean_dec_ref(x_121); -lean_dec_ref(x_120); +lean_dec(x_120); lean_dec_ref(x_119); +lean_dec_ref(x_118); lean_dec_ref(x_117); -lean_dec(x_116); -lean_dec_ref(x_114); -lean_dec_ref(x_112); +lean_dec_ref(x_116); +lean_dec_ref(x_115); +lean_dec_ref(x_113); lean_dec_ref(x_110); lean_dec_ref(x_13); lean_dec_ref(x_12); @@ -83437,13 +76075,13 @@ else { uint8_t x_216; lean_dec(x_122); -lean_dec_ref(x_121); -lean_dec_ref(x_120); +lean_dec(x_120); lean_dec_ref(x_119); +lean_dec_ref(x_118); lean_dec_ref(x_117); -lean_dec(x_116); -lean_dec_ref(x_114); -lean_dec_ref(x_112); +lean_dec_ref(x_116); +lean_dec_ref(x_115); +lean_dec_ref(x_113); lean_dec_ref(x_110); lean_dec_ref(x_13); lean_dec_ref(x_12); @@ -83521,17 +76159,17 @@ if (x_243 == 0) { lean_dec_ref(x_14); x_110 = x_242; -x_111 = x_228; -x_112 = x_236; -x_113 = x_237; -x_114 = x_223; -x_115 = x_228; -x_116 = x_226; +x_111 = x_237; +x_112 = x_228; +x_113 = x_236; +x_114 = x_228; +x_115 = x_225; +x_116 = x_223; x_117 = x_233; -x_118 = lean_box(0); -x_119 = x_234; -x_120 = x_235; -x_121 = x_225; +x_118 = x_234; +x_119 = x_235; +x_120 = x_226; +x_121 = lean_box(0); x_122 = x_224; x_123 = x_243; goto block_219; @@ -83543,17 +76181,17 @@ x_244 = lean_unsigned_to_nat(0u); x_245 = lean_array_get_borrowed(x_14, x_237, x_244); x_246 = l_Lean_Expr_isFVarOf(x_245, x_3); x_110 = x_242; -x_111 = x_228; -x_112 = x_236; -x_113 = x_237; -x_114 = x_223; -x_115 = x_228; -x_116 = x_226; +x_111 = x_237; +x_112 = x_228; +x_113 = x_236; +x_114 = x_228; +x_115 = x_225; +x_116 = x_223; x_117 = x_233; -x_118 = lean_box(0); -x_119 = x_234; -x_120 = x_235; -x_121 = x_225; +x_118 = x_234; +x_119 = x_235; +x_120 = x_226; +x_121 = lean_box(0); x_122 = x_224; x_123 = x_246; goto block_219; @@ -83629,7 +76267,7 @@ if (x_260 == 0) { lean_object* x_261; lean_inc_ref(x_2); -x_261 = l_Lean_Meta_instantiateMVarsIfMVarApp___redArg(x_2, x_257); +x_261 = l_Lean_Meta_instantiateMVarsIfMVarApp___redArg(x_2, x_258); if (lean_obj_tag(x_261) == 0) { uint8_t x_262; @@ -83650,12 +76288,12 @@ if (x_267 == 0) { lean_dec_ref(x_264); x_220 = x_260; -x_221 = x_258; -x_222 = x_256; -x_223 = x_255; -x_224 = x_257; -x_225 = x_254; -x_226 = x_253; +x_221 = x_255; +x_222 = x_254; +x_223 = x_253; +x_224 = x_258; +x_225 = x_259; +x_226 = x_257; x_227 = lean_box(0); goto block_252; } @@ -83670,12 +76308,12 @@ if (x_269 == 0) lean_dec_ref(x_268); lean_dec_ref(x_264); x_220 = x_260; -x_221 = x_258; -x_222 = x_256; -x_223 = x_255; -x_224 = x_257; -x_225 = x_254; -x_226 = x_253; +x_221 = x_255; +x_222 = x_254; +x_223 = x_253; +x_224 = x_258; +x_225 = x_259; +x_226 = x_257; x_227 = lean_box(0); goto block_252; } @@ -83692,12 +76330,12 @@ if (x_272 == 0) lean_dec_ref(x_268); lean_dec_ref(x_264); x_220 = x_260; -x_221 = x_258; -x_222 = x_256; -x_223 = x_255; -x_224 = x_257; -x_225 = x_254; -x_226 = x_253; +x_221 = x_255; +x_222 = x_254; +x_223 = x_253; +x_224 = x_258; +x_225 = x_259; +x_226 = x_257; x_227 = lean_box(0); goto block_252; } @@ -83713,7 +76351,7 @@ lean_dec_ref(x_9); lean_dec_ref(x_8); lean_dec_ref(x_2); lean_inc_ref(x_1); -x_273 = l_Lean_Meta_instantiateMVarsIfMVarApp___redArg(x_1, x_257); +x_273 = l_Lean_Meta_instantiateMVarsIfMVarApp___redArg(x_1, x_258); if (lean_obj_tag(x_273) == 0) { lean_object* x_274; lean_object* x_275; uint8_t x_276; @@ -83732,12 +76370,12 @@ lean_dec_ref(x_6); lean_dec_ref(x_5); lean_dec(x_4); lean_dec(x_3); -x_22 = x_258; -x_23 = x_256; -x_24 = x_255; -x_25 = x_257; -x_26 = x_254; -x_27 = x_253; +x_22 = x_255; +x_23 = x_254; +x_24 = x_253; +x_25 = x_258; +x_26 = x_259; +x_27 = x_257; x_28 = lean_box(0); goto block_33; } @@ -83758,12 +76396,12 @@ lean_dec_ref(x_6); lean_dec_ref(x_5); lean_dec(x_4); lean_dec(x_3); -x_22 = x_258; -x_23 = x_256; -x_24 = x_255; -x_25 = x_257; -x_26 = x_254; -x_27 = x_253; +x_22 = x_255; +x_23 = x_254; +x_24 = x_253; +x_25 = x_258; +x_26 = x_259; +x_27 = x_257; x_28 = lean_box(0); goto block_33; } @@ -83785,12 +76423,12 @@ lean_dec_ref(x_6); lean_dec_ref(x_5); lean_dec(x_4); lean_dec(x_3); -x_22 = x_258; -x_23 = x_256; -x_24 = x_255; -x_25 = x_257; -x_26 = x_254; -x_27 = x_253; +x_22 = x_255; +x_23 = x_254; +x_24 = x_253; +x_25 = x_258; +x_26 = x_259; +x_27 = x_257; x_28 = lean_box(0); goto block_33; } @@ -83811,12 +76449,12 @@ lean_dec_ref(x_6); lean_dec_ref(x_5); lean_dec(x_4); lean_dec(x_3); -x_22 = x_258; -x_23 = x_256; -x_24 = x_255; -x_25 = x_257; -x_26 = x_254; -x_27 = x_253; +x_22 = x_255; +x_23 = x_254; +x_24 = x_253; +x_25 = x_258; +x_26 = x_259; +x_27 = x_257; x_28 = lean_box(0); goto block_33; } @@ -83838,12 +76476,12 @@ lean_dec_ref(x_6); lean_dec_ref(x_5); lean_dec(x_4); lean_dec(x_3); -x_22 = x_258; -x_23 = x_256; -x_24 = x_255; -x_25 = x_257; -x_26 = x_254; -x_27 = x_253; +x_22 = x_255; +x_23 = x_254; +x_24 = x_253; +x_25 = x_258; +x_26 = x_259; +x_27 = x_257; x_28 = lean_box(0); goto block_33; } @@ -83857,17 +76495,17 @@ lean_dec_ref(x_268); x_287 = lean_ctor_get(x_277, 1); lean_inc_ref(x_287); lean_dec_ref(x_277); -lean_inc(x_253); -lean_inc_ref(x_254); lean_inc(x_257); -lean_inc_ref(x_255); +lean_inc_ref(x_259); +lean_inc(x_258); +lean_inc_ref(x_253); lean_inc_ref(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc_ref(x_286); lean_inc_ref(x_7); lean_inc_ref(x_6); -x_288 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody(x_6, x_7, x_286, x_3, x_4, x_5, x_287, x_258, x_256, x_255, x_257, x_254, x_253); +x_288 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody(x_6, x_7, x_286, x_3, x_4, x_5, x_287, x_255, x_254, x_253, x_258, x_259, x_257); if (lean_obj_tag(x_288) == 0) { lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; @@ -83891,7 +76529,7 @@ x_295 = lean_ctor_get(x_275, 1); lean_inc_ref(x_295); lean_dec_ref(x_275); lean_inc_ref(x_294); -x_296 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody(x_6, x_7, x_294, x_3, x_4, x_5, x_295, x_293, x_291, x_255, x_257, x_254, x_253); +x_296 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody(x_6, x_7, x_294, x_3, x_4, x_5, x_295, x_293, x_291, x_253, x_258, x_259, x_257); if (lean_obj_tag(x_296) == 0) { uint8_t x_297; @@ -84032,10 +76670,10 @@ else lean_dec_ref(x_286); lean_dec_ref(x_275); lean_dec_ref(x_264); +lean_dec_ref(x_259); +lean_dec(x_258); lean_dec(x_257); -lean_dec_ref(x_255); -lean_dec_ref(x_254); -lean_dec(x_253); +lean_dec_ref(x_253); lean_dec_ref(x_7); lean_dec_ref(x_6); lean_dec_ref(x_5); @@ -84054,12 +76692,12 @@ else uint8_t x_331; lean_dec_ref(x_268); lean_dec_ref(x_264); -lean_dec_ref(x_258); +lean_dec_ref(x_259); +lean_dec(x_258); lean_dec(x_257); -lean_dec_ref(x_256); lean_dec_ref(x_255); lean_dec_ref(x_254); -lean_dec(x_253); +lean_dec_ref(x_253); lean_dec_ref(x_7); lean_dec_ref(x_6); lean_dec_ref(x_5); @@ -84090,10 +76728,10 @@ else { lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_dec_ref(x_264); +lean_dec_ref(x_259); +lean_dec(x_258); lean_dec(x_257); -lean_dec_ref(x_255); -lean_dec_ref(x_254); -lean_dec(x_253); +lean_dec_ref(x_253); lean_dec_ref(x_14); lean_dec_ref(x_13); lean_dec_ref(x_12); @@ -84111,10 +76749,10 @@ lean_dec_ref(x_1); x_334 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__13; x_335 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_335, 0, x_334); -lean_ctor_set(x_335, 1, x_258); +lean_ctor_set(x_335, 1, x_255); x_336 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_336, 0, x_335); -lean_ctor_set(x_336, 1, x_256); +lean_ctor_set(x_336, 1, x_254); lean_ctor_set(x_261, 0, x_336); return x_261; } @@ -84136,12 +76774,12 @@ if (x_341 == 0) { lean_dec_ref(x_338); x_220 = x_260; -x_221 = x_258; -x_222 = x_256; -x_223 = x_255; -x_224 = x_257; -x_225 = x_254; -x_226 = x_253; +x_221 = x_255; +x_222 = x_254; +x_223 = x_253; +x_224 = x_258; +x_225 = x_259; +x_226 = x_257; x_227 = lean_box(0); goto block_252; } @@ -84156,12 +76794,12 @@ if (x_343 == 0) lean_dec_ref(x_342); lean_dec_ref(x_338); x_220 = x_260; -x_221 = x_258; -x_222 = x_256; -x_223 = x_255; -x_224 = x_257; -x_225 = x_254; -x_226 = x_253; +x_221 = x_255; +x_222 = x_254; +x_223 = x_253; +x_224 = x_258; +x_225 = x_259; +x_226 = x_257; x_227 = lean_box(0); goto block_252; } @@ -84178,12 +76816,12 @@ if (x_346 == 0) lean_dec_ref(x_342); lean_dec_ref(x_338); x_220 = x_260; -x_221 = x_258; -x_222 = x_256; -x_223 = x_255; -x_224 = x_257; -x_225 = x_254; -x_226 = x_253; +x_221 = x_255; +x_222 = x_254; +x_223 = x_253; +x_224 = x_258; +x_225 = x_259; +x_226 = x_257; x_227 = lean_box(0); goto block_252; } @@ -84199,7 +76837,7 @@ lean_dec_ref(x_9); lean_dec_ref(x_8); lean_dec_ref(x_2); lean_inc_ref(x_1); -x_347 = l_Lean_Meta_instantiateMVarsIfMVarApp___redArg(x_1, x_257); +x_347 = l_Lean_Meta_instantiateMVarsIfMVarApp___redArg(x_1, x_258); if (lean_obj_tag(x_347) == 0) { lean_object* x_348; lean_object* x_349; uint8_t x_350; @@ -84218,12 +76856,12 @@ lean_dec_ref(x_6); lean_dec_ref(x_5); lean_dec(x_4); lean_dec(x_3); -x_22 = x_258; -x_23 = x_256; -x_24 = x_255; -x_25 = x_257; -x_26 = x_254; -x_27 = x_253; +x_22 = x_255; +x_23 = x_254; +x_24 = x_253; +x_25 = x_258; +x_26 = x_259; +x_27 = x_257; x_28 = lean_box(0); goto block_33; } @@ -84244,12 +76882,12 @@ lean_dec_ref(x_6); lean_dec_ref(x_5); lean_dec(x_4); lean_dec(x_3); -x_22 = x_258; -x_23 = x_256; -x_24 = x_255; -x_25 = x_257; -x_26 = x_254; -x_27 = x_253; +x_22 = x_255; +x_23 = x_254; +x_24 = x_253; +x_25 = x_258; +x_26 = x_259; +x_27 = x_257; x_28 = lean_box(0); goto block_33; } @@ -84271,12 +76909,12 @@ lean_dec_ref(x_6); lean_dec_ref(x_5); lean_dec(x_4); lean_dec(x_3); -x_22 = x_258; -x_23 = x_256; -x_24 = x_255; -x_25 = x_257; -x_26 = x_254; -x_27 = x_253; +x_22 = x_255; +x_23 = x_254; +x_24 = x_253; +x_25 = x_258; +x_26 = x_259; +x_27 = x_257; x_28 = lean_box(0); goto block_33; } @@ -84297,12 +76935,12 @@ lean_dec_ref(x_6); lean_dec_ref(x_5); lean_dec(x_4); lean_dec(x_3); -x_22 = x_258; -x_23 = x_256; -x_24 = x_255; -x_25 = x_257; -x_26 = x_254; -x_27 = x_253; +x_22 = x_255; +x_23 = x_254; +x_24 = x_253; +x_25 = x_258; +x_26 = x_259; +x_27 = x_257; x_28 = lean_box(0); goto block_33; } @@ -84324,12 +76962,12 @@ lean_dec_ref(x_6); lean_dec_ref(x_5); lean_dec(x_4); lean_dec(x_3); -x_22 = x_258; -x_23 = x_256; -x_24 = x_255; -x_25 = x_257; -x_26 = x_254; -x_27 = x_253; +x_22 = x_255; +x_23 = x_254; +x_24 = x_253; +x_25 = x_258; +x_26 = x_259; +x_27 = x_257; x_28 = lean_box(0); goto block_33; } @@ -84343,17 +76981,17 @@ lean_dec_ref(x_342); x_361 = lean_ctor_get(x_351, 1); lean_inc_ref(x_361); lean_dec_ref(x_351); -lean_inc(x_253); -lean_inc_ref(x_254); lean_inc(x_257); -lean_inc_ref(x_255); +lean_inc_ref(x_259); +lean_inc(x_258); +lean_inc_ref(x_253); lean_inc_ref(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc_ref(x_360); lean_inc_ref(x_7); lean_inc_ref(x_6); -x_362 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody(x_6, x_7, x_360, x_3, x_4, x_5, x_361, x_258, x_256, x_255, x_257, x_254, x_253); +x_362 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody(x_6, x_7, x_360, x_3, x_4, x_5, x_361, x_255, x_254, x_253, x_258, x_259, x_257); if (lean_obj_tag(x_362) == 0) { lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; @@ -84377,7 +77015,7 @@ x_369 = lean_ctor_get(x_349, 1); lean_inc_ref(x_369); lean_dec_ref(x_349); lean_inc_ref(x_368); -x_370 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody(x_6, x_7, x_368, x_3, x_4, x_5, x_369, x_367, x_365, x_255, x_257, x_254, x_253); +x_370 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody(x_6, x_7, x_368, x_3, x_4, x_5, x_369, x_367, x_365, x_253, x_258, x_259, x_257); if (lean_obj_tag(x_370) == 0) { lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; @@ -84451,10 +77089,10 @@ else lean_dec_ref(x_360); lean_dec_ref(x_349); lean_dec_ref(x_338); +lean_dec_ref(x_259); +lean_dec(x_258); lean_dec(x_257); -lean_dec_ref(x_255); -lean_dec_ref(x_254); -lean_dec(x_253); +lean_dec_ref(x_253); lean_dec_ref(x_7); lean_dec_ref(x_6); lean_dec_ref(x_5); @@ -84473,12 +77111,12 @@ else lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_dec_ref(x_342); lean_dec_ref(x_338); -lean_dec_ref(x_258); +lean_dec_ref(x_259); +lean_dec(x_258); lean_dec(x_257); -lean_dec_ref(x_256); lean_dec_ref(x_255); lean_dec_ref(x_254); -lean_dec(x_253); +lean_dec_ref(x_253); lean_dec_ref(x_7); lean_dec_ref(x_6); lean_dec_ref(x_5); @@ -84510,10 +77148,10 @@ else { lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_dec_ref(x_338); +lean_dec_ref(x_259); +lean_dec(x_258); lean_dec(x_257); -lean_dec_ref(x_255); -lean_dec_ref(x_254); -lean_dec(x_253); +lean_dec_ref(x_253); lean_dec_ref(x_14); lean_dec_ref(x_13); lean_dec_ref(x_12); @@ -84531,10 +77169,10 @@ lean_dec_ref(x_1); x_387 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__13; x_388 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_388, 0, x_387); -lean_ctor_set(x_388, 1, x_258); +lean_ctor_set(x_388, 1, x_255); x_389 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_389, 0, x_388); -lean_ctor_set(x_389, 1, x_256); +lean_ctor_set(x_389, 1, x_254); x_390 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_390, 0, x_389); return x_390; @@ -84544,12 +77182,12 @@ return x_390; else { uint8_t x_391; -lean_dec_ref(x_258); +lean_dec_ref(x_259); +lean_dec(x_258); lean_dec(x_257); -lean_dec_ref(x_256); lean_dec_ref(x_255); lean_dec_ref(x_254); -lean_dec(x_253); +lean_dec_ref(x_253); lean_dec_ref(x_14); lean_dec_ref(x_13); lean_dec_ref(x_12); @@ -84609,7 +77247,7 @@ x_401 = lean_array_get(x_14, x_399, x_400); x_402 = lean_unsigned_to_nat(3u); x_403 = lean_array_get(x_14, x_399, x_402); lean_dec_ref(x_399); -x_404 = l_Lean_Meta_mkAbsurd(x_2, x_401, x_403, x_255, x_257, x_254, x_253); +x_404 = l_Lean_Meta_mkAbsurd(x_2, x_401, x_403, x_253, x_258, x_259, x_257); if (lean_obj_tag(x_404) == 0) { uint8_t x_405; @@ -84620,10 +77258,10 @@ lean_object* x_406; lean_object* x_407; lean_object* x_408; x_406 = lean_ctor_get(x_404, 0); x_407 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_407, 0, x_406); -lean_ctor_set(x_407, 1, x_258); +lean_ctor_set(x_407, 1, x_255); x_408 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_408, 0, x_407); -lean_ctor_set(x_408, 1, x_256); +lean_ctor_set(x_408, 1, x_254); lean_ctor_set(x_404, 0, x_408); return x_404; } @@ -84635,10 +77273,10 @@ lean_inc(x_409); lean_dec(x_404); x_410 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_410, 0, x_409); -lean_ctor_set(x_410, 1, x_258); +lean_ctor_set(x_410, 1, x_255); x_411 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_411, 0, x_410); -lean_ctor_set(x_411, 1, x_256); +lean_ctor_set(x_411, 1, x_254); x_412 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_412, 0, x_411); return x_412; @@ -84647,8 +77285,8 @@ return x_412; else { uint8_t x_413; -lean_dec_ref(x_258); -lean_dec_ref(x_256); +lean_dec_ref(x_255); +lean_dec_ref(x_254); x_413 = !lean_is_exclusive(x_404); if (x_413 == 0) { @@ -84676,13 +77314,13 @@ x_425 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInduction x_426 = l_Lean_Expr_isAppOf(x_1, x_425); if (x_426 == 0) { -x_253 = x_420; +x_253 = x_417; x_254 = x_419; x_255 = x_418; -x_256 = x_417; +x_256 = lean_box(0); x_257 = x_421; -x_258 = x_423; -x_259 = lean_box(0); +x_258 = x_422; +x_259 = x_423; x_260 = x_426; goto block_416; } @@ -84693,13 +77331,13 @@ x_427 = lean_unsigned_to_nat(3u); x_428 = l_Lean_Expr_getAppNumArgs(x_1); x_429 = lean_nat_dec_lt(x_427, x_428); lean_dec(x_428); -x_253 = x_420; +x_253 = x_417; x_254 = x_419; x_255 = x_418; -x_256 = x_417; +x_256 = lean_box(0); x_257 = x_421; -x_258 = x_423; -x_259 = lean_box(0); +x_258 = x_422; +x_259 = x_423; x_260 = x_429; goto block_416; } @@ -84728,7 +77366,7 @@ lean_dec(x_431); x_435 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_432, x_434); x_436 = lean_array_get(x_14, x_435, x_433); lean_dec_ref(x_435); -x_437 = l_Lean_Meta_mkFalseElim(x_2, x_436, x_418, x_421, x_419, x_420); +x_437 = l_Lean_Meta_mkFalseElim(x_2, x_436, x_417, x_422, x_423, x_421); if (lean_obj_tag(x_437) == 0) { uint8_t x_438; @@ -84739,10 +77377,10 @@ lean_object* x_439; lean_object* x_440; lean_object* x_441; x_439 = lean_ctor_get(x_437, 0); x_440 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_440, 0, x_439); -lean_ctor_set(x_440, 1, x_423); +lean_ctor_set(x_440, 1, x_418); x_441 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_441, 0, x_440); -lean_ctor_set(x_441, 1, x_417); +lean_ctor_set(x_441, 1, x_419); lean_ctor_set(x_437, 0, x_441); return x_437; } @@ -84754,10 +77392,10 @@ lean_inc(x_442); lean_dec(x_437); x_443 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_443, 0, x_442); -lean_ctor_set(x_443, 1, x_423); +lean_ctor_set(x_443, 1, x_418); x_444 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_444, 0, x_443); -lean_ctor_set(x_444, 1, x_417); +lean_ctor_set(x_444, 1, x_419); x_445 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_445, 0, x_444); return x_445; @@ -84766,8 +77404,8 @@ return x_445; else { uint8_t x_446; -lean_dec_ref(x_423); -lean_dec_ref(x_417); +lean_dec_ref(x_419); +lean_dec_ref(x_418); x_446 = !lean_is_exclusive(x_437); if (x_446 == 0) { @@ -85538,7 +78176,27 @@ x_17 = l_Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunI return x_17; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___00Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6_spec__12_spec__15___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static size_t _init_l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7___redArg___closed__0() { +_start: +{ +size_t x_1; size_t x_2; size_t x_3; +x_1 = 5; +x_2 = 1; +x_3 = lean_usize_shift_left(x_2, x_1); +return x_3; +} +} +static size_t _init_l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7___redArg___closed__1() { +_start: +{ +size_t x_1; size_t x_2; size_t x_3; +x_1 = 1; +x_2 = l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7___redArg___closed__0; +x_3 = lean_usize_sub(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___00Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7_spec__14_spec__17___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -85640,16 +78298,16 @@ return x_34; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6_spec__12___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7_spec__14___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; x_4 = lean_unsigned_to_nat(0u); -x_5 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___00Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6_spec__12_spec__15___redArg(x_1, x_4, x_2, x_3); +x_5 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___00Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7_spec__14_spec__17___redArg(x_1, x_4, x_2, x_3); return x_5; } } -static lean_object* _init_l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6___redArg___closed__0() { +static lean_object* _init_l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7___redArg___closed__2() { _start: { lean_object* x_1; @@ -85657,7 +78315,7 @@ x_1 = l_Lean_PersistentHashMap_mkEmptyEntries(lean_box(0), lean_box(0)); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6_spec__13___redArg(size_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7_spec__15___redArg(size_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; uint8_t x_7; @@ -85685,14 +78343,14 @@ x_18 = lean_nat_add(x_4, x_13); lean_dec(x_4); lean_inc(x_9); lean_inc(x_8); -x_19 = l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6___redArg(x_5, x_17, x_1, x_8, x_9); +x_19 = l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7___redArg(x_5, x_17, x_1, x_8, x_9); x_4 = x_18; x_5 = x_19; goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6___redArg(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7___redArg(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { _start: { if (lean_obj_tag(x_1) == 0) @@ -85701,7 +78359,7 @@ lean_object* x_6; size_t x_7; size_t x_8; size_t x_9; size_t x_10; lean_object* x_6 = lean_ctor_get(x_1, 0); x_7 = 5; x_8 = 1; -x_9 = l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3___redArg___closed__1; +x_9 = l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7___redArg___closed__1; x_10 = lean_usize_land(x_2, x_9); x_11 = lean_usize_to_nat(x_10); x_12 = lean_array_get_size(x_6); @@ -85799,7 +78457,7 @@ lean_object* x_35; size_t x_36; size_t x_37; lean_object* x_38; x_35 = lean_ctor_get(x_15, 0); x_36 = lean_usize_shift_right(x_2, x_7); x_37 = lean_usize_add(x_3, x_8); -x_38 = l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6___redArg(x_35, x_36, x_37, x_4, x_5); +x_38 = l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7___redArg(x_35, x_36, x_37, x_4, x_5); lean_ctor_set(x_15, 0, x_38); x_18 = x_15; goto block_21; @@ -85812,7 +78470,7 @@ lean_inc(x_39); lean_dec(x_15); x_40 = lean_usize_shift_right(x_2, x_7); x_41 = lean_usize_add(x_3, x_8); -x_42 = l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6___redArg(x_39, x_40, x_41, x_4, x_5); +x_42 = l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7___redArg(x_39, x_40, x_41, x_4, x_5); x_43 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_43, 0, x_42); x_18 = x_43; @@ -85851,7 +78509,7 @@ x_45 = !lean_is_exclusive(x_1); if (x_45 == 0) { lean_object* x_46; uint8_t x_47; size_t x_54; uint8_t x_55; -x_46 = l_Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6_spec__12___redArg(x_1, x_4, x_5); +x_46 = l_Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7_spec__14___redArg(x_1, x_4, x_5); x_54 = 7; x_55 = lean_usize_dec_le(x_54, x_3); if (x_55 == 0) @@ -85880,8 +78538,8 @@ x_49 = lean_ctor_get(x_46, 1); lean_inc_ref(x_49); lean_dec_ref(x_46); x_50 = lean_unsigned_to_nat(0u); -x_51 = l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6___redArg___closed__0; -x_52 = l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6_spec__13___redArg(x_3, x_48, x_49, x_50, x_51); +x_51 = l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7___redArg___closed__2; +x_52 = l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7_spec__15___redArg(x_3, x_48, x_49, x_50, x_51); lean_dec_ref(x_49); lean_dec_ref(x_48); return x_52; @@ -85903,7 +78561,7 @@ lean_dec(x_1); x_61 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_61, 0, x_59); lean_ctor_set(x_61, 1, x_60); -x_62 = l_Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6_spec__12___redArg(x_61, x_4, x_5); +x_62 = l_Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7_spec__14___redArg(x_61, x_4, x_5); x_70 = 7; x_71 = lean_usize_dec_le(x_70, x_3); if (x_71 == 0) @@ -85932,8 +78590,8 @@ x_65 = lean_ctor_get(x_62, 1); lean_inc_ref(x_65); lean_dec_ref(x_62); x_66 = lean_unsigned_to_nat(0u); -x_67 = l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6___redArg___closed__0; -x_68 = l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6_spec__13___redArg(x_3, x_64, x_65, x_66, x_67); +x_67 = l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7___redArg___closed__2; +x_68 = l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7_spec__15___redArg(x_3, x_64, x_65, x_66, x_67); lean_dec_ref(x_65); lean_dec_ref(x_64); return x_68; @@ -85954,7 +78612,7 @@ uint64_t x_4; size_t x_5; size_t x_6; lean_object* x_7; x_4 = l_Lean_instHashableMVarId_hash(x_2); x_5 = lean_uint64_to_usize(x_4); x_6 = 1; -x_7 = l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6___redArg(x_1, x_5, x_6, x_2, x_3); +x_7 = l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7___redArg(x_1, x_5, x_6, x_2, x_3); return x_7; } } @@ -86106,7 +78764,7 @@ x_8 = l_Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Ta return x_8; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18___redArg___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20___redArg___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_8; lean_object* x_9; lean_object* x_10; @@ -86116,16 +78774,16 @@ x_10 = lean_apply_5(x_9, x_3, x_4, x_5, x_6, lean_box(0)); return x_10; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18___redArg___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20___redArg___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { lean_object* x_8; -x_8 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18___redArg___lam__0(x_1, x_2, x_3, x_4, x_5, x_6); +x_8 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20___redArg___lam__0(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec_ref(x_2); return x_8; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18___redArg(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20___redArg(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_11; lean_object* x_12; uint8_t x_13; @@ -86225,7 +78883,7 @@ return x_51; else { lean_object* x_52; lean_object* x_53; uint8_t x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_52 = lean_alloc_closure((void*)(l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18___redArg___lam__0___boxed), 7, 1); +x_52 = lean_alloc_closure((void*)(l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20___redArg___lam__0___boxed), 7, 1); lean_closure_set(x_52, 0, x_30); x_53 = lean_box(0); x_54 = 0; @@ -86257,7 +78915,7 @@ x_65 = lean_ctor_get(x_64, 0); lean_inc(x_65); lean_dec_ref(x_64); x_66 = lean_unbox(x_62); -x_67 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18_spec__21___redArg(x_5, x_1, x_2, x_3, x_4, x_61, x_66, x_65, x_3, x_6, x_7, x_8, x_9); +x_67 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20_spec__23___redArg(x_5, x_1, x_2, x_3, x_4, x_61, x_66, x_65, x_3, x_6, x_7, x_8, x_9); return x_67; } else @@ -86341,7 +78999,7 @@ return x_87; else { lean_object* x_88; lean_object* x_89; uint8_t x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_88 = lean_alloc_closure((void*)(l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18___redArg___lam__0___boxed), 7, 1); +x_88 = lean_alloc_closure((void*)(l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20___redArg___lam__0___boxed), 7, 1); lean_closure_set(x_88, 0, x_30); x_89 = lean_box(0); x_90 = 0; @@ -86373,7 +79031,7 @@ x_101 = lean_ctor_get(x_100, 0); lean_inc(x_101); lean_dec_ref(x_100); x_102 = lean_unbox(x_98); -x_103 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18_spec__21___redArg(x_5, x_1, x_2, x_3, x_4, x_97, x_102, x_101, x_3, x_6, x_7, x_8, x_9); +x_103 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20_spec__23___redArg(x_5, x_1, x_2, x_3, x_4, x_97, x_102, x_101, x_3, x_6, x_7, x_8, x_9); return x_103; } else @@ -86477,7 +79135,7 @@ return x_126; else { lean_object* x_127; lean_object* x_128; uint8_t x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; -x_127 = lean_alloc_closure((void*)(l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18___redArg___lam__0___boxed), 7, 1); +x_127 = lean_alloc_closure((void*)(l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20___redArg___lam__0___boxed), 7, 1); lean_closure_set(x_127, 0, x_122); x_128 = lean_box(0); x_129 = 0; @@ -86509,7 +79167,7 @@ x_140 = lean_ctor_get(x_139, 0); lean_inc(x_140); lean_dec_ref(x_139); x_141 = lean_unbox(x_137); -x_142 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18_spec__21___redArg(x_5, x_1, x_2, x_3, x_4, x_136, x_141, x_140, x_3, x_6, x_7, x_8, x_9); +x_142 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20_spec__23___redArg(x_5, x_1, x_2, x_3, x_4, x_136, x_141, x_140, x_3, x_6, x_7, x_8, x_9); return x_142; } else @@ -86658,7 +79316,7 @@ return x_180; else { lean_object* x_181; lean_object* x_182; uint8_t x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; -x_181 = lean_alloc_closure((void*)(l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18___redArg___lam__0___boxed), 7, 1); +x_181 = lean_alloc_closure((void*)(l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20___redArg___lam__0___boxed), 7, 1); lean_closure_set(x_181, 0, x_176); x_182 = lean_box(0); x_183 = 0; @@ -86690,7 +79348,7 @@ x_194 = lean_ctor_get(x_193, 0); lean_inc(x_194); lean_dec_ref(x_193); x_195 = lean_unbox(x_191); -x_196 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18_spec__21___redArg(x_5, x_1, x_2, x_3, x_4, x_190, x_195, x_194, x_3, x_6, x_7, x_8, x_9); +x_196 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20_spec__23___redArg(x_5, x_1, x_2, x_3, x_4, x_190, x_195, x_194, x_3, x_6, x_7, x_8, x_9); return x_196; } else @@ -86857,7 +79515,7 @@ return x_237; else { lean_object* x_238; lean_object* x_239; uint8_t x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; -x_238 = lean_alloc_closure((void*)(l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18___redArg___lam__0___boxed), 7, 1); +x_238 = lean_alloc_closure((void*)(l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20___redArg___lam__0___boxed), 7, 1); lean_closure_set(x_238, 0, x_233); x_239 = lean_box(0); x_240 = 0; @@ -86889,7 +79547,7 @@ x_251 = lean_ctor_get(x_250, 0); lean_inc(x_251); lean_dec_ref(x_250); x_252 = lean_unbox(x_248); -x_253 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18_spec__21___redArg(x_5, x_1, x_2, x_3, x_4, x_247, x_252, x_251, x_3, x_6, x_7, x_8, x_9); +x_253 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20_spec__23___redArg(x_5, x_1, x_2, x_3, x_4, x_247, x_252, x_251, x_3, x_6, x_7, x_8, x_9); return x_253; } else @@ -86925,30 +79583,30 @@ return x_256; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18_spec__21___redArg___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20_spec__23___redArg___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_12; lean_object* x_13; x_12 = lean_array_push(x_1, x_6); -x_13 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18___redArg(x_2, x_3, x_4, x_5, x_12, x_7, x_8, x_9, x_10); +x_13 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20___redArg(x_2, x_3, x_4, x_5, x_12, x_7, x_8, x_9, x_10); return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18_spec__21___redArg___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20_spec__23___redArg___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { uint8_t x_12; lean_object* x_13; x_12 = lean_unbox(x_4); -x_13 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18_spec__21___redArg___lam__0(x_1, x_2, x_3, x_12, x_5, x_6, x_7, x_8, x_9, x_10); +x_13 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20_spec__23___redArg___lam__0(x_1, x_2, x_3, x_12, x_5, x_6, x_7, x_8, x_9, x_10); return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18_spec__21___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, uint8_t x_7, lean_object* x_8, uint8_t x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20_spec__23___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, uint8_t x_7, lean_object* x_8, uint8_t x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { lean_object* x_15; lean_object* x_16; lean_object* x_17; x_15 = lean_box(x_4); -x_16 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18_spec__21___redArg___lam__0___boxed), 11, 5); +x_16 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20_spec__23___redArg___lam__0___boxed), 11, 5); lean_closure_set(x_16, 0, x_1); lean_closure_set(x_16, 1, x_2); lean_closure_set(x_16, 2, x_3); @@ -86995,32 +79653,32 @@ return x_23; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18_spec__21(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, uint8_t x_8, lean_object* x_9, uint8_t x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20_spec__23(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, uint8_t x_8, lean_object* x_9, uint8_t x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { lean_object* x_16; -x_16 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18_spec__21___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +x_16 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20_spec__23___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); return x_16; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_10; lean_object* x_11; x_10 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_M_run___redArg___closed__0; -x_11 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18___redArg(x_2, x_3, x_4, x_1, x_10, x_5, x_6, x_7, x_8); +x_11 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20___redArg(x_2, x_3, x_4, x_1, x_10, x_5, x_6, x_7, x_8); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_11; -x_11 = l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__13(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__14(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -87081,26 +79739,26 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; x_10 = lean_array_size(x_2); x_11 = 0; -x_12 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__13(x_10, x_11, x_2); -x_13 = l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14___redArg(x_1, x_12, x_3, x_4, x_5, x_6, x_7, x_8); +x_12 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__14(x_10, x_11, x_2); +x_13 = l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15___redArg(x_1, x_12, x_3, x_4, x_5, x_6, x_7, x_8); return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_11; -x_11 = l_Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__7___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__8___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_8; @@ -87109,11 +79767,11 @@ lean_ctor_set(x_8, 0, x_1); return x_8; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__7___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__8___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { lean_object* x_8; -x_8 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__7___lam__0(x_1, x_2, x_3, x_4, x_5, x_6); +x_8 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__8___lam__0(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_6); lean_dec_ref(x_5); lean_dec(x_4); @@ -87122,7 +79780,7 @@ lean_dec_ref(x_2); return x_8; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__7(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__8(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -87142,7 +79800,7 @@ lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; size_t x_7 = lean_ctor_get(x_5, 1); x_8 = lean_unsigned_to_nat(0u); x_9 = lean_array_uset(x_3, x_2, x_8); -x_10 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__7___lam__0___boxed), 7, 1); +x_10 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__8___lam__0___boxed), 7, 1); lean_closure_set(x_10, 0, x_7); lean_ctor_set(x_5, 1, x_10); x_11 = 1; @@ -87162,7 +79820,7 @@ lean_inc(x_15); lean_dec(x_5); x_17 = lean_unsigned_to_nat(0u); x_18 = lean_array_uset(x_3, x_2, x_17); -x_19 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__7___lam__0___boxed), 7, 1); +x_19 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__8___lam__0___boxed), 7, 1); lean_closure_set(x_19, 0, x_16); x_20 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_20, 0, x_15); @@ -87177,22 +79835,22 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; x_10 = lean_array_size(x_2); x_11 = 0; -x_12 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__7(x_10, x_11, x_2); -x_13 = l_Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8___redArg(x_1, x_12, x_3, x_4, x_5, x_6, x_7, x_8); +x_12 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__8(x_10, x_11, x_2); +x_13 = l_Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9___redArg(x_1, x_12, x_3, x_4, x_5, x_6, x_7, x_8); return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_11; -x_11 = l_Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } @@ -87204,47 +79862,119 @@ x_5 = l_Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__pri return x_5; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6_spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7_spec__14(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_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6_spec__12___redArg(x_2, x_3, x_4); +x_5 = l_Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7_spec__14___redArg(x_2, x_3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___00Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6_spec__12_spec__15(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___00Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7_spec__14_spec__17(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; -x_6 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___00Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6_spec__12_spec__15___redArg(x_2, x_3, x_4, x_5); +x_6 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___00Lean_PersistentHashMap_insertAtCollisionNode___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7_spec__14_spec__17___redArg(x_2, x_3, x_4, x_5); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; -x_7 = l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6___redArg(x_2, x_3, x_4, x_5, x_6); +x_7 = l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7___redArg(x_2, x_3, x_4, x_5, x_6); return x_7; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6_spec__13(lean_object* x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__4___redArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_4; +x_4 = l_Lean_Expr_hasMVar(x_1); +if (x_4 == 0) +{ +lean_object* x_5; +x_5 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_5, 0, x_1); +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_6 = lean_st_ref_get(x_2); +x_7 = lean_ctor_get(x_6, 0); +lean_inc_ref(x_7); +lean_dec(x_6); +x_8 = l_Lean_instantiateMVarsCore(x_7, x_1); +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_ref(x_8); +x_11 = lean_st_ref_take(x_2); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_11, 0); +lean_dec(x_13); +lean_ctor_set(x_11, 0, x_10); +x_14 = lean_st_ref_set(x_2, x_11); +x_15 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_15, 0, x_9); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_16 = lean_ctor_get(x_11, 1); +x_17 = lean_ctor_get(x_11, 2); +x_18 = lean_ctor_get(x_11, 3); +x_19 = lean_ctor_get(x_11, 4); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_11); +x_20 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_20, 0, x_10); +lean_ctor_set(x_20, 1, x_16); +lean_ctor_set(x_20, 2, x_17); +lean_ctor_set(x_20, 3, x_18); +lean_ctor_set(x_20, 4, x_19); +x_21 = lean_st_ref_set(x_2, x_20); +x_22 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_22, 0, x_9); +return x_22; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_instantiateMVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__4___redArg(x_1, x_3); +return x_7; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7_spec__15(lean_object* x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { lean_object* x_8; -x_8 = l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6_spec__13___redArg(x_2, x_3, x_4, x_6, x_7); +x_8 = l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7_spec__15___redArg(x_2, x_3, x_4, x_6, x_7); return x_8; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_12; -x_12 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_12 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_12; } } -static lean_object* _init_l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__1___closed__0() { +static lean_object* _init_l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2___closed__0() { _start: { lean_object* x_1; @@ -87252,7 +79982,7 @@ x_1 = l_Lean_instInhabitedPersistentArrayNode_default(lean_box(0)); return x_1; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__3(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { lean_object* x_5; uint8_t x_10; @@ -87294,7 +80024,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__3(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__4(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -87323,7 +80053,7 @@ else size_t x_8; size_t x_9; lean_object* x_10; x_8 = 0; x_9 = lean_usize_of_nat(x_5); -x_10 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__1_spec__7(x_3, x_8, x_9, x_2); +x_10 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2_spec__9(x_3, x_8, x_9, x_2); return x_10; } } @@ -87332,7 +80062,7 @@ else size_t x_11; size_t x_12; lean_object* x_13; x_11 = 0; x_12 = lean_usize_of_nat(x_5); -x_13 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__1_spec__7(x_3, x_11, x_12, x_2); +x_13 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2_spec__9(x_3, x_11, x_12, x_2); return x_13; } } @@ -87363,7 +80093,7 @@ else size_t x_19; size_t x_20; lean_object* x_21; x_19 = 0; x_20 = lean_usize_of_nat(x_16); -x_21 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2(x_14, x_19, x_20, x_2); +x_21 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__3(x_14, x_19, x_20, x_2); return x_21; } } @@ -87372,14 +80102,14 @@ else size_t x_22; size_t x_23; lean_object* x_24; x_22 = 0; x_23 = lean_usize_of_nat(x_16); -x_24 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2(x_14, x_22, x_23, x_2); +x_24 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__3(x_14, x_22, x_23, x_2); return x_24; } } } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__1_spec__7(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2_spec__9(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -87388,7 +80118,7 @@ if (x_5 == 0) { lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; x_6 = lean_array_uget(x_1, x_2); -x_7 = l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__3(x_6, x_4); +x_7 = l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__4(x_6, x_4); lean_dec(x_6); x_8 = 1; x_9 = lean_usize_add(x_2, x_8); @@ -87402,14 +80132,14 @@ return x_4; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__1(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { if (lean_obj_tag(x_1) == 0) { lean_object* x_5; lean_object* x_6; size_t x_7; lean_object* x_8; lean_object* x_9; size_t x_10; size_t x_11; size_t x_12; size_t x_13; size_t x_14; size_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; x_5 = lean_ctor_get(x_1, 0); -x_6 = l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__1___closed__0; +x_6 = l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2___closed__0; x_7 = lean_usize_shift_right(x_2, x_3); x_8 = lean_usize_to_nat(x_7); x_9 = lean_array_get_borrowed(x_6, x_5, x_8); @@ -87419,7 +80149,7 @@ x_12 = lean_usize_sub(x_11, x_10); x_13 = lean_usize_land(x_2, x_12); x_14 = 5; x_15 = lean_usize_sub(x_3, x_14); -x_16 = l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__1(x_9, x_13, x_15, x_4); +x_16 = l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2(x_9, x_13, x_15, x_4); x_17 = lean_unsigned_to_nat(1u); x_18 = lean_nat_add(x_8, x_17); lean_dec(x_8); @@ -87447,7 +80177,7 @@ size_t x_22; size_t x_23; lean_object* x_24; x_22 = lean_usize_of_nat(x_18); lean_dec(x_18); x_23 = lean_usize_of_nat(x_19); -x_24 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__1_spec__7(x_5, x_22, x_23, x_16); +x_24 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2_spec__9(x_5, x_22, x_23, x_16); return x_24; } } @@ -87457,7 +80187,7 @@ size_t x_25; size_t x_26; lean_object* x_27; x_25 = lean_usize_of_nat(x_18); lean_dec(x_18); x_26 = lean_usize_of_nat(x_19); -x_27 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__1_spec__7(x_5, x_25, x_26, x_16); +x_27 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2_spec__9(x_5, x_25, x_26, x_16); return x_27; } } @@ -87491,7 +80221,7 @@ size_t x_33; size_t x_34; lean_object* x_35; x_33 = lean_usize_of_nat(x_29); lean_dec(x_29); x_34 = lean_usize_of_nat(x_30); -x_35 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2(x_28, x_33, x_34, x_4); +x_35 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__3(x_28, x_33, x_34, x_4); return x_35; } } @@ -87501,7 +80231,7 @@ size_t x_36; size_t x_37; lean_object* x_38; x_36 = lean_usize_of_nat(x_29); lean_dec(x_29); x_37 = lean_usize_of_nat(x_30); -x_38 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2(x_28, x_36, x_37, x_4); +x_38 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__3(x_28, x_36, x_37, x_4); return x_38; } } @@ -87526,7 +80256,7 @@ if (x_10 == 0) { size_t x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; x_11 = lean_usize_of_nat(x_3); -x_12 = l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__1(x_6, x_11, x_8, x_2); +x_12 = l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2(x_6, x_11, x_8, x_2); x_13 = lean_array_get_size(x_7); x_14 = lean_nat_dec_lt(x_4, x_13); if (x_14 == 0) @@ -87548,7 +80278,7 @@ else size_t x_16; size_t x_17; lean_object* x_18; x_16 = 0; x_17 = lean_usize_of_nat(x_13); -x_18 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2(x_7, x_16, x_17, x_12); +x_18 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__3(x_7, x_16, x_17, x_12); return x_18; } } @@ -87557,7 +80287,7 @@ else size_t x_19; size_t x_20; lean_object* x_21; x_19 = 0; x_20 = lean_usize_of_nat(x_13); -x_21 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2(x_7, x_19, x_20, x_12); +x_21 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__3(x_7, x_19, x_20, x_12); return x_21; } } @@ -87590,7 +80320,7 @@ size_t x_26; size_t x_27; lean_object* x_28; x_26 = lean_usize_of_nat(x_22); lean_dec(x_22); x_27 = lean_usize_of_nat(x_23); -x_28 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2(x_7, x_26, x_27, x_2); +x_28 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__3(x_7, x_26, x_27, x_2); return x_28; } } @@ -87600,7 +80330,7 @@ size_t x_29; size_t x_30; lean_object* x_31; x_29 = lean_usize_of_nat(x_22); lean_dec(x_22); x_30 = lean_usize_of_nat(x_23); -x_31 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2(x_7, x_29, x_30, x_2); +x_31 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__3(x_7, x_29, x_30, x_2); return x_31; } } @@ -87611,7 +80341,7 @@ else lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; x_32 = lean_ctor_get(x_1, 0); x_33 = lean_ctor_get(x_1, 1); -x_34 = l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__3(x_32, x_2); +x_34 = l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__4(x_32, x_2); x_35 = lean_array_get_size(x_33); x_36 = lean_nat_dec_lt(x_4, x_35); if (x_36 == 0) @@ -87633,7 +80363,7 @@ else size_t x_38; size_t x_39; lean_object* x_40; x_38 = 0; x_39 = lean_usize_of_nat(x_35); -x_40 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2(x_33, x_38, x_39, x_34); +x_40 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__3(x_33, x_38, x_39, x_34); return x_40; } } @@ -87642,7 +80372,7 @@ else size_t x_41; size_t x_42; lean_object* x_43; x_41 = 0; x_42 = lean_usize_of_nat(x_35); -x_43 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2(x_33, x_41, x_42, x_34); +x_43 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__3(x_33, x_41, x_42, x_34); return x_43; } } @@ -87658,6 +80388,56 @@ x_5 = l_Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00_ return x_5; } } +LEAN_EXPORT lean_object* l_List_mapTR_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__7(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; +x_3 = l_List_reverse___redArg(x_2); +return x_3; +} +else +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_1, 0); +x_6 = lean_ctor_get(x_1, 1); +x_7 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_7, 0, x_5); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set(x_1, 0, x_7); +{ +lean_object* _tmp_0 = x_6; +lean_object* _tmp_1 = x_1; +x_1 = _tmp_0; +x_2 = _tmp_1; +} +goto _start; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_1, 0); +x_10 = lean_ctor_get(x_1, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_1); +x_11 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, x_9); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_2); +x_1 = x_10; +x_2 = x_12; +goto _start; +} +} +} +} static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars___lam__0___closed__0() { _start: { @@ -87742,7 +80522,7 @@ return x_23; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__4(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_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5(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) { _start: { uint8_t x_9; @@ -87947,7 +80727,7 @@ if (lean_obj_tag(x_14) == 0) { lean_object* x_15; lean_object* x_16; uint8_t x_17; uint8_t x_18; uint8_t x_19; lean_object* x_20; lean_dec_ref(x_14); -x_15 = l_Lean_instantiateMVars___at___00Lean_Tactic_FunInd_rwMatcher_spec__3___redArg(x_3, x_6); +x_15 = l_Lean_instantiateMVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__4___redArg(x_3, x_6); x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); lean_dec_ref(x_15); @@ -88180,7 +80960,7 @@ x_10 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta_Tactic_FunInd_0__ x_11 = !lean_is_exclusive(x_10); if (x_11 == 0) { -lean_object* x_12; lean_object* x_13; lean_object* x_14; size_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_63; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; size_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_63; x_12 = lean_ctor_get(x_10, 0); x_13 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars___lam__0___boxed), 1, 0); x_14 = l_Lean_instInhabitedExpr; @@ -88215,7 +80995,7 @@ lean_ctor_set(x_69, 1, x_68); lean_inc_ref(x_1); x_70 = lean_array_to_list(x_1); x_71 = lean_box(0); -x_72 = l_List_mapTR_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__4(x_70, x_71); +x_72 = l_List_mapTR_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__7(x_70, x_71); x_73 = l_Lean_MessageData_ofList(x_72); x_74 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_74, 0, x_69); @@ -88264,7 +81044,7 @@ return x_78; size_t x_23; lean_object* x_24; x_23 = lean_array_size(x_17); lean_inc_ref(x_17); -x_24 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__4(x_23, x_15, x_17, x_18, x_19, x_20, x_21); +x_24 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5(x_23, x_16, x_17, x_18, x_19, x_20, x_21); if (lean_obj_tag(x_24) == 0) { lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; @@ -88278,7 +81058,7 @@ x_28 = l_Array_zip___redArg(x_27, x_25); lean_dec(x_25); lean_dec_ref(x_27); x_29 = 0; -x_30 = l_Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5___redArg(x_14, x_28, x_16, x_29, x_18, x_19, x_20, x_21); +x_30 = l_Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6___redArg(x_14, x_28, x_15, x_29, x_18, x_19, x_20, x_21); return x_30; } else @@ -88289,7 +81069,7 @@ lean_dec_ref(x_20); lean_dec(x_19); lean_dec_ref(x_18); lean_dec_ref(x_17); -lean_dec_ref(x_16); +lean_dec_ref(x_15); lean_dec_ref(x_13); x_31 = !lean_is_exclusive(x_24); if (x_31 == 0) @@ -88338,8 +81118,8 @@ x_48 = lean_unbox(x_45); lean_dec(x_45); if (x_48 == 0) { -x_15 = x_41; -x_16 = x_47; +x_15 = x_47; +x_16 = x_41; x_17 = x_43; x_18 = x_35; x_19 = x_36; @@ -88355,7 +81135,7 @@ x_49 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndepen lean_inc(x_43); x_50 = lean_array_to_list(x_43); x_51 = lean_box(0); -x_52 = l_List_mapTR_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__4(x_50, x_51); +x_52 = l_List_mapTR_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__7(x_50, x_51); x_53 = l_Lean_MessageData_ofList(x_52); x_54 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_54, 0, x_49); @@ -88364,8 +81144,8 @@ x_55 = l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic if (lean_obj_tag(x_55) == 0) { lean_dec_ref(x_55); -x_15 = x_41; -x_16 = x_47; +x_15 = x_47; +x_16 = x_41; x_17 = x_43; x_18 = x_35; x_19 = x_36; @@ -88431,7 +81211,7 @@ return x_61; } else { -lean_object* x_79; lean_object* x_80; lean_object* x_81; size_t x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; uint8_t x_130; +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; size_t x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; uint8_t x_130; x_79 = lean_ctor_get(x_10, 0); lean_inc(x_79); lean_dec(x_10); @@ -88467,7 +81247,7 @@ lean_ctor_set(x_137, 1, x_136); lean_inc_ref(x_1); x_138 = lean_array_to_list(x_1); x_139 = lean_box(0); -x_140 = l_List_mapTR_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__4(x_138, x_139); +x_140 = l_List_mapTR_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__7(x_138, x_139); x_141 = l_Lean_MessageData_ofList(x_140); x_142 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_142, 0, x_137); @@ -88517,7 +81297,7 @@ return x_146; size_t x_90; lean_object* x_91; x_90 = lean_array_size(x_84); lean_inc_ref(x_84); -x_91 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__4(x_90, x_82, x_84, x_85, x_86, x_87, x_88); +x_91 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5(x_90, x_83, x_84, x_85, x_86, x_87, x_88); if (lean_obj_tag(x_91) == 0) { lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; uint8_t x_96; lean_object* x_97; @@ -88531,7 +81311,7 @@ x_95 = l_Array_zip___redArg(x_94, x_92); lean_dec(x_92); lean_dec_ref(x_94); x_96 = 0; -x_97 = l_Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5___redArg(x_81, x_95, x_83, x_96, x_85, x_86, x_87, x_88); +x_97 = l_Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6___redArg(x_81, x_95, x_82, x_96, x_85, x_86, x_87, x_88); return x_97; } else @@ -88542,7 +81322,7 @@ lean_dec_ref(x_87); lean_dec(x_86); lean_dec_ref(x_85); lean_dec_ref(x_84); -lean_dec_ref(x_83); +lean_dec_ref(x_82); lean_dec_ref(x_80); x_98 = lean_ctor_get(x_91, 0); lean_inc(x_98); @@ -88592,8 +81372,8 @@ x_115 = lean_unbox(x_112); lean_dec(x_112); if (x_115 == 0) { -x_82 = x_108; -x_83 = x_114; +x_82 = x_114; +x_83 = x_108; x_84 = x_110; x_85 = x_102; x_86 = x_103; @@ -88609,7 +81389,7 @@ x_116 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndepe lean_inc(x_110); x_117 = lean_array_to_list(x_110); x_118 = lean_box(0); -x_119 = l_List_mapTR_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__4(x_117, x_118); +x_119 = l_List_mapTR_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__7(x_117, x_118); x_120 = l_Lean_MessageData_ofList(x_119); x_121 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_121, 0, x_116); @@ -88618,8 +81398,8 @@ x_122 = l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tacti if (lean_obj_tag(x_122) == 0) { lean_dec_ref(x_122); -x_82 = x_108; -x_83 = x_114; +x_82 = x_114; +x_83 = x_108; x_84 = x_110; x_85 = x_102; x_86 = x_103; @@ -88699,45 +81479,45 @@ lean_dec_ref(x_3); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18_spec__21___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20_spec__23___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { uint8_t x_16; uint8_t x_17; uint8_t x_18; lean_object* x_19; x_16 = lean_unbox(x_5); x_17 = lean_unbox(x_8); x_18 = lean_unbox(x_10); -x_19 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18_spec__21(x_1, x_2, x_3, x_4, x_16, x_6, x_7, x_17, x_9, x_18, x_11, x_12, x_13, x_14); +x_19 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20_spec__23(x_1, x_2, x_3, x_4, x_16, x_6, x_7, x_17, x_9, x_18, x_11, x_12, x_13, x_14); return x_19; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { uint8_t x_11; lean_object* x_12; x_11 = lean_unbox(x_5); -x_12 = l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14(x_1, x_2, x_3, x_4, x_11, x_6, x_7, x_8, x_9); +x_12 = l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15(x_1, x_2, x_3, x_4, x_11, x_6, x_7, x_8, x_9); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { uint8_t x_11; lean_object* x_12; x_11 = lean_unbox(x_5); -x_12 = l_Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8(x_1, x_2, x_3, x_4, x_11, x_6, x_7, x_8, x_9); +x_12 = l_Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9(x_1, x_2, x_3, x_4, x_11, x_6, x_7, x_8, x_9); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { uint8_t x_11; lean_object* x_12; x_11 = lean_unbox(x_5); -x_12 = l_Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5(x_1, x_2, x_3, x_4, x_11, x_6, x_7, x_8, x_9); +x_12 = l_Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6(x_1, x_2, x_3, x_4, x_11, x_6, x_7, x_8, x_9); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { size_t x_7; size_t x_8; lean_object* x_9; @@ -88745,28 +81525,40 @@ x_7 = lean_unbox_usize(x_3); lean_dec(x_3); x_8 = lean_unbox_usize(x_4); lean_dec(x_4); -x_9 = l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6(x_1, x_2, x_7, x_8, x_5, x_6); +x_9 = l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7(x_1, x_2, x_7, x_8, x_5, x_6); return x_9; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6_spec__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_instantiateMVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__4(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec_ref(x_4); +lean_dec(x_3); +lean_dec_ref(x_2); +return x_7; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7_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) { _start: { size_t x_8; lean_object* x_9; x_8 = lean_unbox_usize(x_2); lean_dec(x_2); -x_9 = l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6_spec__13(x_1, x_8, x_3, x_4, x_5, x_6, x_7); +x_9 = l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7_spec__15(x_1, x_8, x_3, x_4, x_5, x_6, x_7); lean_dec_ref(x_4); lean_dec_ref(x_3); return x_9; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { uint8_t x_12; lean_object* x_13; x_12 = lean_unbox(x_4); -x_13 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18(x_1, x_2, x_3, x_12, x_5, x_6, x_7, x_8, x_9, x_10); +x_13 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20(x_1, x_2, x_3, x_12, x_5, x_6, x_7, x_8, x_9, x_10); return x_13; } } @@ -88780,45 +81572,45 @@ lean_dec_ref(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { uint8_t x_10; lean_object* x_11; x_10 = lean_unbox(x_4); -x_11 = l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14___redArg(x_1, x_2, x_3, x_10, x_5, x_6, x_7, x_8); +x_11 = l_Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15___redArg(x_1, x_2, x_3, x_10, x_5, x_6, x_7, x_8); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { uint8_t x_10; lean_object* x_11; x_10 = lean_unbox(x_4); -x_11 = l_Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8___redArg(x_1, x_2, x_3, x_10, x_5, x_6, x_7, x_8); +x_11 = l_Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9___redArg(x_1, x_2, x_3, x_10, x_5, x_6, x_7, x_8); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { uint8_t x_10; lean_object* x_11; x_10 = lean_unbox(x_4); -x_11 = l_Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5___redArg(x_1, x_2, x_3, x_10, x_5, x_6, x_7, x_8); +x_11 = l_Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6___redArg(x_1, x_2, x_3, x_10, x_5, x_6, x_7, x_8); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18_spec__21___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20_spec__23___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { uint8_t x_15; uint8_t x_16; uint8_t x_17; lean_object* x_18; x_15 = lean_unbox(x_4); x_16 = lean_unbox(x_7); x_17 = lean_unbox(x_9); -x_18 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18_spec__21___redArg(x_1, x_2, x_3, x_15, x_5, x_6, x_16, x_8, x_17, x_10, x_11, x_12, x_13); +x_18 = l_Lean_Meta_withLocalDecl___at___00Lean_Meta_withLocalDeclD___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__0_spec__0___at___00__private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20_spec__23___redArg(x_1, x_2, x_3, x_15, x_5, x_6, x_16, x_8, x_17, x_10, x_11, x_12, x_13); return x_18; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__1_spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2_spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { size_t x_5; size_t x_6; lean_object* x_7; @@ -88826,7 +81618,7 @@ x_5 = lean_unbox_usize(x_2); lean_dec(x_2); x_6 = lean_unbox_usize(x_3); lean_dec(x_3); -x_7 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__1_spec__7(x_1, x_5, x_6, x_4); +x_7 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2_spec__9(x_1, x_5, x_6, x_4); lean_dec_ref(x_1); return x_7; } @@ -88840,7 +81632,7 @@ lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { size_t x_5; size_t x_6; lean_object* x_7; @@ -88848,12 +81640,12 @@ x_5 = lean_unbox_usize(x_2); lean_dec(x_2); x_6 = lean_unbox_usize(x_3); lean_dec(x_3); -x_7 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2(x_1, x_5, x_6, x_4); +x_7 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__3(x_1, x_5, x_6, x_4); lean_dec_ref(x_1); return x_7; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -88861,11 +81653,11 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__7(x_4, x_5, x_3); +x_6 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__8(x_4, x_5, x_3); return x_6; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__14___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; @@ -88873,11 +81665,11 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__13(x_4, x_5, x_3); +x_6 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__14(x_4, x_5, x_3); return x_6; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { size_t x_9; size_t x_10; lean_object* x_11; @@ -88885,7 +81677,7 @@ x_9 = lean_unbox_usize(x_1); lean_dec(x_1); x_10 = lean_unbox_usize(x_2); lean_dec(x_2); -x_11 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__4(x_9, x_10, x_3, x_4, x_5, x_6, x_7); +x_11 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5(x_9, x_10, x_3, x_4, x_5, x_6, x_7); lean_dec(x_7); lean_dec_ref(x_6); lean_dec(x_5); @@ -88893,18 +81685,27 @@ lean_dec_ref(x_4); return x_11; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6_spec__13___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7_spec__15___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { size_t x_6; lean_object* x_7; x_6 = lean_unbox_usize(x_1); lean_dec(x_1); -x_7 = l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6_spec__13___redArg(x_6, x_2, x_3, x_4, x_5); +x_7 = l___private_Lean_Data_PersistentHashMap_0__Lean_PersistentHashMap_insertAux_traverse___at___00Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7_spec__15___redArg(x_6, x_2, x_3, x_4, x_5); lean_dec_ref(x_3); lean_dec_ref(x_2); return x_7; } } +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__4___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_instantiateMVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__4___redArg(x_1, x_2); +lean_dec(x_2); +return x_4; +} +} LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_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: { @@ -88934,16 +81735,16 @@ lean_dec_ref(x_1); return x_12; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__3___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__4___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__3(x_1, x_2); +x_3 = l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__4(x_1, x_2); lean_dec_ref(x_1); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { size_t x_5; size_t x_6; lean_object* x_7; @@ -88951,21 +81752,21 @@ 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___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__1(x_1, x_5, x_6, x_4); +x_7 = l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2(x_1, x_5, x_6, x_4); lean_dec_ref(x_1); return x_7; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { uint8_t x_11; lean_object* x_12; x_11 = lean_unbox(x_3); -x_12 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5_spec__8_spec__14_spec__18___redArg(x_1, x_2, x_11, x_4, x_5, x_6, x_7, x_8, x_9); +x_12 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDecls_loop___at___00Lean_Meta_withLocalDecls___at___00Lean_Meta_withLocalDeclsD___at___00Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6_spec__9_spec__15_spec__20___redArg(x_1, x_2, x_11, x_4, x_5, x_6, x_7, x_8, x_9); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7___redArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { size_t x_6; size_t x_7; lean_object* x_8; @@ -88973,7 +81774,7 @@ x_6 = lean_unbox_usize(x_2); lean_dec(x_2); x_7 = lean_unbox_usize(x_3); lean_dec(x_3); -x_8 = l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6___redArg(x_1, x_6, x_7, x_4, x_5); +x_8 = l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7___redArg(x_1, x_6, x_7, x_4, x_5); return x_8; } } @@ -90997,7 +83798,7 @@ if (x_54 == 0) lean_object* x_55; lean_object* x_56; lean_object* x_57; uint8_t x_58; x_55 = lean_ctor_get(x_53, 0); x_56 = lean_ctor_get(x_53, 1); -x_57 = l_Lean_instantiateMVars___at___00Lean_Tactic_FunInd_rwMatcher_spec__3___redArg(x_56, x_25); +x_57 = l_Lean_instantiateMVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__4___redArg(x_56, x_25); lean_dec(x_25); x_58 = !lean_is_exclusive(x_57); if (x_58 == 0) @@ -91030,7 +83831,7 @@ x_63 = lean_ctor_get(x_53, 1); lean_inc(x_63); lean_inc(x_62); lean_dec(x_53); -x_64 = l_Lean_instantiateMVars___at___00Lean_Tactic_FunInd_rwMatcher_spec__3___redArg(x_63, x_25); +x_64 = l_Lean_instantiateMVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__4___redArg(x_63, x_25); lean_dec(x_25); x_65 = lean_ctor_get(x_64, 0); lean_inc(x_65); @@ -91255,11 +84056,11 @@ x_102 = l_Lean_indentExpr(x_7); x_103 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_103, 0, x_101); lean_ctor_set(x_103, 1, x_102); -x_104 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___redArg(x_103, x_94, x_92, x_93, x_95); -lean_dec(x_95); -lean_dec_ref(x_93); +x_104 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___redArg(x_103, x_95, x_94, x_93, x_92); lean_dec(x_92); -lean_dec_ref(x_94); +lean_dec_ref(x_93); +lean_dec(x_94); +lean_dec_ref(x_95); x_105 = !lean_is_exclusive(x_104); if (x_105 == 0) { @@ -91296,7 +84097,7 @@ x_116 = l_Lean_MessageData_ofFormat(x_115); x_117 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_117, 0, x_112); lean_ctor_set(x_117, 1, x_116); -x_118 = l_Lean_Tactic_FunInd_rwMatcher___closed__3; +x_118 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__1___closed__5; x_119 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_119, 0, x_117); lean_ctor_set(x_119, 1, x_118); @@ -91304,11 +84105,11 @@ x_120 = l_Lean_indentExpr(x_7); x_121 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_121, 0, x_119); lean_ctor_set(x_121, 1, x_120); -x_122 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___redArg(x_121, x_94, x_92, x_93, x_95); -lean_dec(x_95); -lean_dec_ref(x_93); +x_122 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___redArg(x_121, x_95, x_94, x_93, x_92); lean_dec(x_92); -lean_dec_ref(x_94); +lean_dec_ref(x_93); +lean_dec(x_94); +lean_dec_ref(x_95); x_123 = !lean_is_exclusive(x_122); if (x_123 == 0) { @@ -91337,20 +84138,20 @@ lean_object* x_134; uint8_t x_135; x_134 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__8___closed__7; x_135 = l_Lean_Expr_isAppOf(x_7, x_134); x_91 = lean_box(0); -x_92 = x_128; +x_92 = x_130; x_93 = x_129; -x_94 = x_127; -x_95 = x_130; +x_94 = x_128; +x_95 = x_127; x_96 = x_135; goto block_126; } else { x_91 = lean_box(0); -x_92 = x_128; +x_92 = x_130; x_93 = x_129; -x_94 = x_127; -x_95 = x_130; +x_94 = x_128; +x_95 = x_127; x_96 = x_133; goto block_126; } @@ -94812,7 +87613,7 @@ static lean_object* _init_l_WellFounded_opaqueFix_u2083___at___00__private_Lean_ 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_WellFounded_opaqueFix_u2083___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_setNaryFunIndInfo_spec__2___redArg___closed__1; x_2 = lean_unsigned_to_nat(8u); -x_3 = lean_unsigned_to_nat(1058u); +x_3 = lean_unsigned_to_nat(939u); x_4 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_setNaryFunIndInfo_spec__2___redArg___closed__0; x_5 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__17___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -95169,7 +87970,7 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_Fun lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_setNaryFunIndInfo___closed__0; x_2 = lean_unsigned_to_nat(2u); -x_3 = lean_unsigned_to_nat(1048u); +x_3 = lean_unsigned_to_nat(929u); x_4 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_setNaryFunIndInfo_spec__2___redArg___closed__0; x_5 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__17___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -95211,7 +88012,7 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_Fun lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_setNaryFunIndInfo___closed__4; x_2 = lean_unsigned_to_nat(4u); -x_3 = lean_unsigned_to_nat(1063u); +x_3 = lean_unsigned_to_nat(944u); x_4 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_setNaryFunIndInfo_spec__2___redArg___closed__0; x_5 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__17___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -95556,25 +88357,51 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } +static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("casesOn", 7, 7); +return x_1; +} +} static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__2() { _start: { lean_object* x_1; +x_1 = lean_mk_string_unchecked("PSum", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__3; +x_2 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__2; +x_3 = l_Lean_Name_mkStr2(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__5() { +_start: +{ +lean_object* x_1; x_1 = lean_mk_string_unchecked("inl", 3, 3); return x_1; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__3() { +static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__2; -x_2 = l_Lean_Tactic_FunInd_rwMatcher___closed__12; +x_1 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__5; +x_2 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__2; x_3 = l_Lean_Name_mkStr2(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__4() { +static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__7() { _start: { lean_object* x_1; @@ -95582,27 +88409,35 @@ x_1 = lean_mk_string_unchecked("inr", 3, 3); return x_1; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__5() { +static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__4; -x_2 = l_Lean_Tactic_FunInd_rwMatcher___closed__12; +x_1 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__7; +x_2 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__2; x_3 = l_Lean_Name_mkStr2(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__6() { +static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("PSigma", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_inProdLambdaLastArg___closed__0; -x_2 = l_Lean_Tactic_FunInd_rwMatcher___closed__15; +x_2 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__9; x_3 = l_Lean_Name_mkStr2(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__7() { +static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__11() { _start: { lean_object* x_1; @@ -95610,20 +88445,20 @@ x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.FunInd.0.Lean.Tactic.F return x_1; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__8() { +static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___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; x_1 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__17___closed__2; x_2 = lean_unsigned_to_nat(48u); -x_3 = lean_unsigned_to_nat(1116u); -x_4 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__7; +x_3 = lean_unsigned_to_nat(997u); +x_4 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__11; x_5 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__17___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); return x_6; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__9() { +static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__13() { _start: { lean_object* x_1; @@ -95631,30 +88466,30 @@ x_1 = lean_mk_string_unchecked("snd", 3, 3); return x_1; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__10() { +static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__9; -x_2 = l_Lean_Tactic_FunInd_rwMatcher___closed__15; +x_1 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__13; +x_2 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__9; x_3 = l_Lean_Name_mkStr2(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__11() { +static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__17___closed__2; x_2 = lean_unsigned_to_nat(51u); -x_3 = lean_unsigned_to_nat(1110u); -x_4 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__7; +x_3 = lean_unsigned_to_nat(991u); +x_4 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__11; x_5 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__17___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); return x_6; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__12() { +static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__16() { _start: { lean_object* x_1; @@ -95662,37 +88497,47 @@ x_1 = lean_mk_string_unchecked("fst", 3, 3); return x_1; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__13() { +static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__12; -x_2 = l_Lean_Tactic_FunInd_rwMatcher___closed__15; +x_1 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__16; +x_2 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__9; x_3 = l_Lean_Name_mkStr2(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__14() { +static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__17___closed__2; x_2 = lean_unsigned_to_nat(51u); -x_3 = lean_unsigned_to_nat(1102u); -x_4 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__7; +x_3 = lean_unsigned_to_nat(983u); +x_4 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__11; x_5 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__17___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); return x_6; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__15() { +static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__3; +x_2 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__9; +x_3 = l_Lean_Name_mkStr2(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__17___closed__2; x_2 = lean_unsigned_to_nat(50u); -x_3 = lean_unsigned_to_nat(1093u); -x_4 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__7; +x_3 = lean_unsigned_to_nat(974u); +x_4 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__11; x_5 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__17___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); return x_6; @@ -95723,7 +88568,7 @@ else { lean_object* x_253; uint8_t x_254; lean_dec(x_5); -x_253 = l_Lean_Tactic_FunInd_rwMatcher___closed__16; +x_253 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__19; x_254 = l_Lean_Expr_isConstOf(x_3, x_253); if (x_254 == 0) { @@ -95756,7 +88601,7 @@ x_258 = l_Lean_instInhabitedExpr; x_259 = lean_unsigned_to_nat(3u); x_260 = lean_array_get_borrowed(x_258, x_4, x_259); x_261 = lean_unsigned_to_nat(4u); -x_262 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__6; +x_262 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__10; x_263 = l_Lean_Expr_isAppOfArity(x_260, x_262, x_261); if (x_263 == 0) { @@ -95785,7 +88630,7 @@ if (x_271 == 0) { lean_object* x_272; lean_object* x_273; lean_dec_ref(x_269); -x_272 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__15; +x_272 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__20; lean_inc(x_9); lean_inc_ref(x_8); lean_inc(x_7); @@ -95875,17 +88720,17 @@ return x_14; block_77: { lean_object* x_27; -lean_inc_ref(x_24); -x_27 = l_Lean_Meta_ArgsPacker_unpack(x_16, x_24); -lean_dec_ref(x_16); +lean_inc_ref(x_19); +x_27 = l_Lean_Meta_ArgsPacker_unpack(x_25, x_19); +lean_dec_ref(x_25); if (lean_obj_tag(x_27) == 1) { uint8_t x_28; -lean_dec_ref(x_24); -lean_dec_ref(x_22); +lean_dec(x_23); +lean_dec_ref(x_21); lean_dec_ref(x_20); +lean_dec_ref(x_19); lean_dec(x_18); -lean_dec(x_17); x_28 = !lean_is_exclusive(x_27); if (x_28 == 0) { @@ -95903,15 +88748,15 @@ lean_inc(x_33); lean_dec_ref(x_26); x_34 = l_Array_toSubarray___redArg(x_4, x_32, x_33); x_35 = l_panic___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__26___closed__1; -x_36 = lean_array_get(x_35, x_23, x_30); -lean_dec_ref(x_23); -x_37 = l_Subarray_toArray___redArg(x_19); +x_36 = lean_array_get(x_35, x_24, x_30); +lean_dec_ref(x_24); +x_37 = l_Subarray_toArray___redArg(x_17); x_38 = l_Lean_Elab_FixedParamPerm_buildArgs___redArg(x_36, x_37, x_31); lean_dec_ref(x_37); x_39 = lean_box(0); -x_40 = lean_array_get(x_39, x_21, x_30); +x_40 = lean_array_get(x_39, x_16, x_30); lean_dec(x_30); -lean_dec_ref(x_21); +lean_dec_ref(x_16); x_41 = l_Lean_Expr_getAppFn(x_1); lean_dec_ref(x_1); x_42 = l_Lean_Expr_constLevels_x21(x_41); @@ -95947,15 +88792,15 @@ lean_inc(x_53); lean_dec_ref(x_26); x_54 = l_Array_toSubarray___redArg(x_4, x_52, x_53); x_55 = l_panic___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__26___closed__1; -x_56 = lean_array_get(x_55, x_23, x_50); -lean_dec_ref(x_23); -x_57 = l_Subarray_toArray___redArg(x_19); +x_56 = lean_array_get(x_55, x_24, x_50); +lean_dec_ref(x_24); +x_57 = l_Subarray_toArray___redArg(x_17); x_58 = l_Lean_Elab_FixedParamPerm_buildArgs___redArg(x_56, x_57, x_51); lean_dec_ref(x_57); x_59 = lean_box(0); -x_60 = lean_array_get(x_59, x_21, x_50); +x_60 = lean_array_get(x_59, x_16, x_50); lean_dec(x_50); -lean_dec_ref(x_21); +lean_dec_ref(x_16); x_61 = l_Lean_Expr_getAppFn(x_1); lean_dec_ref(x_1); x_62 = l_Lean_Expr_constLevels_x21(x_61); @@ -95980,21 +88825,21 @@ else lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; uint8_t x_74; lean_dec(x_27); lean_dec_ref(x_26); -lean_dec_ref(x_23); -lean_dec_ref(x_21); -lean_dec_ref(x_19); +lean_dec_ref(x_24); +lean_dec_ref(x_17); +lean_dec_ref(x_16); lean_dec_ref(x_4); lean_dec_ref(x_1); x_70 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__1; -x_71 = l_Lean_indentExpr(x_24); +x_71 = l_Lean_indentExpr(x_19); x_72 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_72, 0, x_70); lean_ctor_set(x_72, 1, x_71); -x_73 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___redArg(x_72, x_22, x_18, x_20, x_17); -lean_dec(x_17); -lean_dec_ref(x_20); +x_73 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___redArg(x_72, x_21, x_23, x_20, x_18); lean_dec(x_18); -lean_dec_ref(x_22); +lean_dec_ref(x_20); +lean_dec(x_23); +lean_dec_ref(x_21); x_74 = !lean_is_exclusive(x_73); if (x_74 == 0) { @@ -96084,16 +88929,16 @@ x_98 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_98, 0, x_91); lean_ctor_set(x_98, 1, x_92); lean_inc(x_96); -x_16 = x_85; -x_17 = x_81; -x_18 = x_79; -x_19 = x_95; +x_16 = x_83; +x_17 = x_95; +x_18 = x_81; +x_19 = x_96; x_20 = x_80; -x_21 = x_83; -x_22 = x_78; -x_23 = x_89; -x_24 = x_96; -x_25 = lean_box(0); +x_21 = x_78; +x_22 = lean_box(0); +x_23 = x_79; +x_24 = x_89; +x_25 = x_85; x_26 = x_98; goto block_77; } @@ -96105,16 +88950,16 @@ x_99 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_99, 0, x_94); lean_ctor_set(x_99, 1, x_92); lean_inc(x_96); -x_16 = x_85; -x_17 = x_81; -x_18 = x_79; -x_19 = x_95; +x_16 = x_83; +x_17 = x_95; +x_18 = x_81; +x_19 = x_96; x_20 = x_80; -x_21 = x_83; -x_22 = x_78; -x_23 = x_89; -x_24 = x_96; -x_25 = lean_box(0); +x_21 = x_78; +x_22 = lean_box(0); +x_23 = x_79; +x_24 = x_89; +x_25 = x_85; x_26 = x_99; goto block_77; } @@ -96124,7 +88969,7 @@ goto block_77; block_139: { lean_object* x_106; uint8_t x_107; -x_106 = l_Lean_Tactic_FunInd_rwMatcher___closed__14; +x_106 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__4; x_107 = l_Lean_Expr_isConstOf(x_3, x_106); if (x_107 == 0) { @@ -96158,12 +89003,12 @@ x_112 = lean_unsigned_to_nat(3u); x_113 = lean_array_get_borrowed(x_111, x_4, x_112); lean_inc_ref(x_4); x_114 = l_Array_toSubarray___redArg(x_4, x_108, x_109); -x_115 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__3; +x_115 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__6; x_116 = l_Lean_Expr_isAppOfArity(x_113, x_115, x_112); if (x_116 == 0) { lean_object* x_117; uint8_t x_118; -x_117 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__5; +x_117 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__8; x_118 = l_Lean_Expr_isAppOfArity(x_113, x_117, x_112); if (x_118 == 0) { @@ -96260,7 +89105,7 @@ else { lean_object* x_152; lean_object* x_153; lean_object* x_154; uint8_t x_155; x_152 = l_Lean_Expr_projExpr_x21(x_1); -x_153 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__6; +x_153 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__10; x_154 = lean_unsigned_to_nat(4u); x_155 = l_Lean_Expr_isAppOfArity(x_152, x_153, x_154); if (x_155 == 0) @@ -96290,7 +89135,7 @@ if (x_163 == 0) { lean_object* x_164; lean_object* x_165; lean_dec_ref(x_161); -x_164 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__8; +x_164 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__12; lean_inc(x_149); lean_inc_ref(x_148); lean_inc(x_147); @@ -96375,7 +89220,7 @@ goto block_145; block_211: { lean_object* x_182; uint8_t x_183; -x_182 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__10; +x_182 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__14; x_183 = l_Lean_Expr_isConstOf(x_3, x_182); if (x_183 == 0) { @@ -96406,7 +89251,7 @@ else lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; uint8_t x_191; x_187 = lean_unsigned_to_nat(2u); x_188 = lean_array_fget_borrowed(x_4, x_187); -x_189 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__6; +x_189 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__10; x_190 = lean_unsigned_to_nat(4u); x_191 = l_Lean_Expr_isAppOfArity(x_188, x_189, x_190); if (x_191 == 0) @@ -96436,7 +89281,7 @@ if (x_199 == 0) { lean_object* x_200; lean_object* x_201; lean_dec_ref(x_197); -x_200 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__11; +x_200 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__15; lean_inc(x_180); lean_inc_ref(x_179); lean_inc(x_178); @@ -96508,7 +89353,7 @@ return x_210; block_246: { lean_object* x_217; uint8_t x_218; -x_217 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__13; +x_217 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__17; x_218 = l_Lean_Expr_isConstOf(x_3, x_217); if (x_218 == 0) { @@ -96539,7 +89384,7 @@ else lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; uint8_t x_226; x_222 = lean_unsigned_to_nat(2u); x_223 = lean_array_fget_borrowed(x_4, x_222); -x_224 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__6; +x_224 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__10; x_225 = lean_unsigned_to_nat(4u); x_226 = l_Lean_Expr_isAppOfArity(x_223, x_224, x_225); if (x_226 == 0) @@ -96569,7 +89414,7 @@ if (x_234 == 0) { lean_object* x_235; lean_object* x_236; lean_dec_ref(x_232); -x_235 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__14; +x_235 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__18; lean_inc(x_215); lean_inc_ref(x_214); lean_inc(x_213); @@ -96661,7 +89506,7 @@ return x_252; else { lean_object* x_253; uint8_t x_254; -x_253 = l_Lean_Tactic_FunInd_rwMatcher___closed__16; +x_253 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__19; x_254 = l_Lean_Expr_isConstOf(x_3, x_253); if (x_254 == 0) { @@ -96694,7 +89539,7 @@ x_258 = l_Lean_instInhabitedExpr; x_259 = lean_unsigned_to_nat(3u); x_260 = lean_array_get_borrowed(x_258, x_4, x_259); x_261 = lean_unsigned_to_nat(4u); -x_262 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__6; +x_262 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__10; x_263 = l_Lean_Expr_isAppOfArity(x_260, x_262, x_261); if (x_263 == 0) { @@ -96723,7 +89568,7 @@ if (x_271 == 0) { lean_object* x_272; lean_object* x_273; lean_dec_ref(x_269); -x_272 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__15; +x_272 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__20; lean_inc(x_9); lean_inc_ref(x_8); lean_inc(x_7); @@ -96813,17 +89658,17 @@ return x_14; block_77: { lean_object* x_27; -lean_inc_ref(x_17); -x_27 = l_Lean_Meta_ArgsPacker_unpack(x_18, x_17); -lean_dec_ref(x_18); +lean_inc_ref(x_20); +x_27 = l_Lean_Meta_ArgsPacker_unpack(x_25, x_20); +lean_dec_ref(x_25); if (lean_obj_tag(x_27) == 1) { uint8_t x_28; -lean_dec_ref(x_24); lean_dec_ref(x_23); lean_dec(x_22); -lean_dec(x_19); -lean_dec_ref(x_17); +lean_dec_ref(x_20); +lean_dec_ref(x_18); +lean_dec(x_17); x_28 = !lean_is_exclusive(x_27); if (x_28 == 0) { @@ -96841,15 +89686,15 @@ lean_inc(x_33); lean_dec_ref(x_26); x_34 = l_Array_toSubarray___redArg(x_4, x_32, x_33); x_35 = l_panic___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__26___closed__1; -x_36 = lean_array_get(x_35, x_21, x_30); -lean_dec_ref(x_21); -x_37 = l_Subarray_toArray___redArg(x_25); +x_36 = lean_array_get(x_35, x_19, x_30); +lean_dec_ref(x_19); +x_37 = l_Subarray_toArray___redArg(x_21); x_38 = l_Lean_Elab_FixedParamPerm_buildArgs___redArg(x_36, x_37, x_31); lean_dec_ref(x_37); x_39 = lean_box(0); -x_40 = lean_array_get(x_39, x_20, x_30); +x_40 = lean_array_get(x_39, x_24, x_30); lean_dec(x_30); -lean_dec_ref(x_20); +lean_dec_ref(x_24); x_41 = l_Lean_Expr_getAppFn(x_1); lean_dec_ref(x_1); x_42 = l_Lean_Expr_constLevels_x21(x_41); @@ -96885,15 +89730,15 @@ lean_inc(x_53); lean_dec_ref(x_26); x_54 = l_Array_toSubarray___redArg(x_4, x_52, x_53); x_55 = l_panic___at___00Lean_Meta_MatcherApp_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__12_spec__26___closed__1; -x_56 = lean_array_get(x_55, x_21, x_50); -lean_dec_ref(x_21); -x_57 = l_Subarray_toArray___redArg(x_25); +x_56 = lean_array_get(x_55, x_19, x_50); +lean_dec_ref(x_19); +x_57 = l_Subarray_toArray___redArg(x_21); x_58 = l_Lean_Elab_FixedParamPerm_buildArgs___redArg(x_56, x_57, x_51); lean_dec_ref(x_57); x_59 = lean_box(0); -x_60 = lean_array_get(x_59, x_20, x_50); +x_60 = lean_array_get(x_59, x_24, x_50); lean_dec(x_50); -lean_dec_ref(x_20); +lean_dec_ref(x_24); x_61 = l_Lean_Expr_getAppFn(x_1); lean_dec_ref(x_1); x_62 = l_Lean_Expr_constLevels_x21(x_61); @@ -96918,21 +89763,21 @@ else lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; uint8_t x_74; lean_dec(x_27); lean_dec_ref(x_26); -lean_dec_ref(x_25); +lean_dec_ref(x_24); lean_dec_ref(x_21); -lean_dec_ref(x_20); +lean_dec_ref(x_19); lean_dec_ref(x_4); lean_dec_ref(x_1); x_70 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__1; -x_71 = l_Lean_indentExpr(x_17); +x_71 = l_Lean_indentExpr(x_20); x_72 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_72, 0, x_70); lean_ctor_set(x_72, 1, x_71); -x_73 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___redArg(x_72, x_24, x_19, x_23, x_22); +x_73 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___redArg(x_72, x_18, x_17, x_23, x_22); lean_dec(x_22); lean_dec_ref(x_23); -lean_dec(x_19); -lean_dec_ref(x_24); +lean_dec(x_17); +lean_dec_ref(x_18); x_74 = !lean_is_exclusive(x_73); if (x_74 == 0) { @@ -97023,15 +89868,15 @@ lean_ctor_set(x_98, 0, x_91); lean_ctor_set(x_98, 1, x_92); lean_inc(x_96); x_16 = lean_box(0); -x_17 = x_96; -x_18 = x_85; -x_19 = x_79; -x_20 = x_83; -x_21 = x_89; +x_17 = x_79; +x_18 = x_78; +x_19 = x_89; +x_20 = x_96; +x_21 = x_95; x_22 = x_81; x_23 = x_80; -x_24 = x_78; -x_25 = x_95; +x_24 = x_83; +x_25 = x_85; x_26 = x_98; goto block_77; } @@ -97044,15 +89889,15 @@ lean_ctor_set(x_99, 0, x_94); lean_ctor_set(x_99, 1, x_92); lean_inc(x_96); x_16 = lean_box(0); -x_17 = x_96; -x_18 = x_85; -x_19 = x_79; -x_20 = x_83; -x_21 = x_89; +x_17 = x_79; +x_18 = x_78; +x_19 = x_89; +x_20 = x_96; +x_21 = x_95; x_22 = x_81; x_23 = x_80; -x_24 = x_78; -x_25 = x_95; +x_24 = x_83; +x_25 = x_85; x_26 = x_99; goto block_77; } @@ -97062,7 +89907,7 @@ goto block_77; block_139: { lean_object* x_106; uint8_t x_107; -x_106 = l_Lean_Tactic_FunInd_rwMatcher___closed__14; +x_106 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__4; x_107 = l_Lean_Expr_isConstOf(x_3, x_106); if (x_107 == 0) { @@ -97096,12 +89941,12 @@ x_112 = lean_unsigned_to_nat(3u); x_113 = lean_array_get_borrowed(x_111, x_4, x_112); lean_inc_ref(x_4); x_114 = l_Array_toSubarray___redArg(x_4, x_108, x_109); -x_115 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__3; +x_115 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__6; x_116 = l_Lean_Expr_isAppOfArity(x_113, x_115, x_112); if (x_116 == 0) { lean_object* x_117; uint8_t x_118; -x_117 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__5; +x_117 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__8; x_118 = l_Lean_Expr_isAppOfArity(x_113, x_117, x_112); if (x_118 == 0) { @@ -97198,7 +90043,7 @@ else { lean_object* x_152; lean_object* x_153; lean_object* x_154; uint8_t x_155; x_152 = l_Lean_Expr_projExpr_x21(x_1); -x_153 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__6; +x_153 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__10; x_154 = lean_unsigned_to_nat(4u); x_155 = l_Lean_Expr_isAppOfArity(x_152, x_153, x_154); if (x_155 == 0) @@ -97228,7 +90073,7 @@ if (x_163 == 0) { lean_object* x_164; lean_object* x_165; lean_dec_ref(x_161); -x_164 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__8; +x_164 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__12; lean_inc(x_149); lean_inc_ref(x_148); lean_inc(x_147); @@ -97313,7 +90158,7 @@ goto block_145; block_211: { lean_object* x_182; uint8_t x_183; -x_182 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__10; +x_182 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__14; x_183 = l_Lean_Expr_isConstOf(x_3, x_182); if (x_183 == 0) { @@ -97344,7 +90189,7 @@ else lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; uint8_t x_191; x_187 = lean_unsigned_to_nat(2u); x_188 = lean_array_fget_borrowed(x_4, x_187); -x_189 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__6; +x_189 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__10; x_190 = lean_unsigned_to_nat(4u); x_191 = l_Lean_Expr_isAppOfArity(x_188, x_189, x_190); if (x_191 == 0) @@ -97374,7 +90219,7 @@ if (x_199 == 0) { lean_object* x_200; lean_object* x_201; lean_dec_ref(x_197); -x_200 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__11; +x_200 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__15; lean_inc(x_180); lean_inc_ref(x_179); lean_inc(x_178); @@ -97446,7 +90291,7 @@ return x_210; block_246: { lean_object* x_217; uint8_t x_218; -x_217 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__13; +x_217 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__17; x_218 = l_Lean_Expr_isConstOf(x_3, x_217); if (x_218 == 0) { @@ -97477,7 +90322,7 @@ else lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; uint8_t x_226; x_222 = lean_unsigned_to_nat(2u); x_223 = lean_array_fget_borrowed(x_4, x_222); -x_224 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__6; +x_224 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__10; x_225 = lean_unsigned_to_nat(4u); x_226 = l_Lean_Expr_isAppOfArity(x_223, x_224, x_225); if (x_226 == 0) @@ -97507,7 +90352,7 @@ if (x_234 == 0) { lean_object* x_235; lean_object* x_236; lean_dec_ref(x_232); -x_235 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__14; +x_235 = l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__18; lean_inc(x_215); lean_inc_ref(x_214); lean_inc(x_213); @@ -101472,7 +94317,7 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_Fun lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_withLetDecls___redArg___closed__1; x_2 = lean_unsigned_to_nat(2u); -x_3 = lean_unsigned_to_nat(1206u); +x_3 = lean_unsigned_to_nat(1087u); x_4 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_withLetDecls___redArg___closed__0; x_5 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__17___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -102836,14 +95681,14 @@ goto block_18; block_30: { uint8_t x_29; -x_29 = lean_nat_dec_le(x_28, x_27); +x_29 = lean_nat_dec_le(x_28, x_26); if (x_29 == 0) { -lean_dec(x_27); +lean_dec(x_26); lean_inc(x_28); x_19 = x_25; x_20 = x_28; -x_21 = x_26; +x_21 = x_27; x_22 = x_28; goto block_24; } @@ -102851,8 +95696,8 @@ else { x_19 = x_25; x_20 = x_28; -x_21 = x_26; -x_22 = x_27; +x_21 = x_27; +x_22 = x_26; goto block_24; } } @@ -102872,16 +95717,16 @@ if (x_37 == 0) { lean_inc(x_36); x_25 = x_32; -x_26 = x_31; -x_27 = x_36; +x_26 = x_36; +x_27 = x_31; x_28 = x_36; goto block_30; } else { x_25 = x_32; -x_26 = x_31; -x_27 = x_36; +x_26 = x_36; +x_27 = x_31; x_28 = x_33; goto block_30; } @@ -103650,1698 +96495,550 @@ lean_dec(x_6); lean_dec_ref(x_5); lean_dec(x_4); lean_dec_ref(x_3); -lean_dec_ref(x_2); -x_37 = !lean_is_exclusive(x_9); -if (x_37 == 0) -{ -return x_9; -} -else -{ -lean_object* x_38; lean_object* x_39; -x_38 = lean_ctor_get(x_9, 0); -lean_inc(x_38); -lean_dec(x_9); -x_39 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_39, 0, x_38); -return x_39; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_withErasedFVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__18(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_9; -x_9 = l_Lean_Meta_withErasedFVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__18___redArg(x_2, x_3, x_4, x_5, x_6, x_7); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_withLCtx___at___00Lean_Meta_withErasedFVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__18_spec__33(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_10; -x_10 = l_Lean_Meta_withLCtx___at___00Lean_Meta_withErasedFVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__18_spec__33___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_10; -} -} -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__21___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_7; lean_object* x_8; uint8_t x_13; -x_13 = lean_nat_dec_lt(x_4, x_1); -if (x_13 == 0) -{ -lean_object* x_14; -lean_dec(x_4); -x_14 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_14, 0, x_5); -return x_14; -} -else -{ -lean_object* x_15; -x_15 = lean_array_fget_borrowed(x_2, x_4); -if (lean_obj_tag(x_15) == 1) -{ -lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_16 = lean_ctor_get(x_15, 0); -x_17 = 0; -x_18 = lean_box(x_17); -x_19 = lean_array_get_borrowed(x_18, x_3, x_16); -x_20 = lean_unbox(x_19); -if (x_20 == 0) -{ -uint8_t x_21; lean_object* x_22; lean_object* x_23; -x_21 = 0; -x_22 = lean_box(x_21); -x_23 = lean_array_push(x_5, x_22); -x_7 = x_23; -x_8 = lean_box(0); -goto block_12; -} -else -{ -uint8_t x_24; lean_object* x_25; lean_object* x_26; -x_24 = 1; -x_25 = lean_box(x_24); -x_26 = lean_array_push(x_5, x_25); -x_7 = x_26; -x_8 = lean_box(0); -goto block_12; -} -} -else -{ -uint8_t x_27; lean_object* x_28; lean_object* x_29; -x_27 = 2; -x_28 = lean_box(x_27); -x_29 = lean_array_push(x_5, x_28); -x_7 = x_29; -x_8 = lean_box(0); -goto block_12; -} -} -block_12: -{ -lean_object* x_9; lean_object* x_10; -x_9 = lean_unsigned_to_nat(1u); -x_10 = lean_nat_add(x_4, x_9); -lean_dec(x_4); -x_4 = x_10; -x_5 = x_7; -goto _start; -} -} -} -LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__21(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_14; -x_14 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__21___redArg(x_1, x_2, x_3, x_6, x_7); -return x_14; -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Elab_Structural_Positions_groupAndSort___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__6_spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, size_t x_7, size_t x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; -x_10 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Elab_Structural_Positions_groupAndSort___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__6_spec__7___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_10; -} -} -LEAN_EXPORT uint8_t l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Array_contains___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__14_spec__25(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { -_start: -{ -uint8_t x_5; -x_5 = lean_usize_dec_eq(x_3, x_4); -if (x_5 == 0) -{ -lean_object* x_6; uint8_t x_7; -x_6 = lean_array_uget(x_2, x_3); -x_7 = lean_nat_dec_eq(x_1, x_6); -lean_dec(x_6); -if (x_7 == 0) -{ -size_t x_8; size_t x_9; -x_8 = 1; -x_9 = lean_usize_add(x_3, x_8); -x_3 = x_9; -goto _start; -} -else -{ -return x_7; -} -} -else -{ -uint8_t x_11; -x_11 = 0; -return x_11; -} -} -} -LEAN_EXPORT uint8_t l_Array_contains___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__14(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; lean_object* x_4; uint8_t x_5; -x_3 = lean_unsigned_to_nat(0u); -x_4 = lean_array_get_size(x_1); -x_5 = lean_nat_dec_lt(x_3, x_4); -if (x_5 == 0) -{ -return x_5; -} -else -{ -if (x_5 == 0) -{ -return x_5; -} -else -{ -size_t x_6; size_t x_7; uint8_t x_8; -x_6 = 0; -x_7 = lean_usize_of_nat(x_4); -x_8 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Array_contains___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__14_spec__25(x_2, x_1, x_6, x_7); -return x_8; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__15_spec__27(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { -_start: -{ -uint8_t x_8; -x_8 = lean_usize_dec_lt(x_6, x_5); -if (x_8 == 0) -{ -lean_inc_ref(x_7); -return x_7; -} -else -{ -lean_object* x_9; uint8_t x_10; -x_9 = lean_array_uget(x_4, x_6); -x_10 = l_Array_contains___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__14(x_9, x_1); -if (x_10 == 0) -{ -size_t x_11; size_t x_12; -lean_dec(x_9); -x_11 = 1; -x_12 = lean_usize_add(x_6, x_11); -{ -size_t _tmp_5 = x_12; -lean_object* _tmp_6 = x_2; -x_6 = _tmp_5; -x_7 = _tmp_6; -} -goto _start; -} -else -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_14, 0, x_9); -x_15 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_15, 0, x_14); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_3); -return x_16; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__15___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { -_start: -{ -uint8_t x_8; -x_8 = lean_usize_dec_lt(x_6, x_5); -if (x_8 == 0) -{ -lean_inc_ref(x_7); -return x_7; -} -else -{ -lean_object* x_9; uint8_t x_10; -x_9 = lean_array_uget(x_4, x_6); -x_10 = l_Array_contains___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__14(x_9, x_1); -if (x_10 == 0) -{ -size_t x_11; size_t x_12; lean_object* x_13; -lean_dec(x_9); -x_11 = 1; -x_12 = lean_usize_add(x_6, x_11); -x_13 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__15_spec__27(x_1, x_2, x_3, x_4, x_5, x_12, x_2); -return x_13; -} -else -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_14, 0, x_9); -x_15 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_15, 0, x_14); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_3); -return x_16; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__15(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__15___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_9; -} -} -LEAN_EXPORT uint8_t l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__16___lam__0(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; -x_3 = lean_nat_dec_eq(x_2, x_1); -return x_3; -} -} -LEAN_EXPORT uint8_t l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__16___lam__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; -x_3 = l_Array_contains___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__14(x_2, x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Array_idxOfAux___at___00Array_finIdxOf_x3f___at___00Array_idxOf_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__11_spec__18_spec__21(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; uint8_t x_5; -x_4 = lean_array_get_size(x_1); -x_5 = lean_nat_dec_lt(x_3, x_4); -if (x_5 == 0) -{ -lean_object* x_6; -lean_dec(x_3); -x_6 = lean_box(0); -return x_6; -} -else -{ -lean_object* x_7; uint8_t x_8; -x_7 = lean_array_fget_borrowed(x_1, x_3); -x_8 = lean_expr_eqv(x_7, x_2); -if (x_8 == 0) -{ -lean_object* x_9; lean_object* x_10; -x_9 = lean_unsigned_to_nat(1u); -x_10 = lean_nat_add(x_3, x_9); -lean_dec(x_3); -x_3 = x_10; -goto _start; -} -else -{ -lean_object* x_12; -x_12 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_12, 0, x_3); -return x_12; -} -} -} -} -LEAN_EXPORT lean_object* l_Array_finIdxOf_x3f___at___00Array_idxOf_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__11_spec__18(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; lean_object* x_4; -x_3 = lean_unsigned_to_nat(0u); -x_4 = l_Array_idxOfAux___at___00Array_finIdxOf_x3f___at___00Array_idxOf_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__11_spec__18_spec__21(x_1, x_2, x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Array_idxOf_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__11(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_Array_finIdxOf_x3f___at___00Array_idxOf_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__11_spec__18(x_1, x_2); -if (lean_obj_tag(x_3) == 0) -{ -lean_object* x_4; -x_4 = lean_box(0); -return x_4; -} -else -{ -uint8_t x_5; -x_5 = !lean_is_exclusive(x_3); -if (x_5 == 0) -{ -return x_3; -} -else -{ -lean_object* x_6; lean_object* x_7; -x_6 = lean_ctor_get(x_3, 0); -lean_inc(x_6); -lean_dec(x_3); -x_7 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_7, 0, x_6); -return x_7; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, uint8_t x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; lean_object* x_11; -x_10 = l_Lean_Expr_getAppFn(x_9); -x_11 = l_Array_idxOf_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__11(x_1, x_10); -lean_dec_ref(x_10); -if (lean_obj_tag(x_11) == 1) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -if (lean_is_exclusive(x_11)) { - lean_ctor_release(x_11, 0); - x_13 = x_11; -} else { - lean_dec_ref(x_11); - x_13 = lean_box(0); -} -x_14 = l_Lean_Expr_getAppNumArgs(x_9); -x_15 = lean_unsigned_to_nat(0u); -x_16 = lean_array_get_borrowed(x_15, x_2, x_12); -x_17 = lean_nat_dec_eq(x_14, x_16); -if (x_17 == 0) -{ -lean_object* x_18; -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec_ref(x_9); -lean_dec_ref(x_5); -lean_dec_ref(x_3); -x_18 = lean_box(0); -return x_18; -} -else -{ -lean_object* x_19; lean_object* x_20; -x_19 = lean_array_get_borrowed(x_3, x_4, x_12); -if (x_8 == 0) -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_32 = l___private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0___lam__1___closed__0; -lean_inc(x_14); -x_33 = lean_mk_array(x_14, x_32); -x_34 = lean_unsigned_to_nat(1u); -x_35 = lean_nat_sub(x_14, x_34); -lean_dec(x_14); -x_36 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_9, x_33, x_35); -x_20 = x_36; -goto block_31; -} -else -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_37 = l___private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0___lam__1___closed__0; -lean_inc(x_14); -x_38 = lean_mk_array(x_14, x_37); -x_39 = lean_unsigned_to_nat(1u); -x_40 = lean_nat_sub(x_14, x_39); -lean_dec(x_14); -x_41 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_9, x_38, x_40); -x_42 = lean_array_pop(x_41); -x_20 = x_42; -goto block_31; -} -block_31: -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_21 = lean_ctor_get(x_19, 0); -x_22 = lean_ctor_get(x_21, 0); -x_23 = lean_ctor_get(x_21, 1); -x_24 = lean_array_get_borrowed(x_5, x_6, x_12); -lean_dec(x_12); -lean_inc(x_24); -x_25 = l_Lean_Elab_FixedParamPerm_buildArgs___redArg(x_24, x_7, x_20); -x_26 = lean_box(0); -lean_inc(x_23); -x_27 = l_List_mapTR_loop___at___00Lean_mkConstWithLevelParams___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__2_spec__4(x_23, x_26); -lean_inc(x_22); -x_28 = l_Lean_Expr_const___override(x_22, x_27); -x_29 = l_Lean_mkAppN(x_28, x_25); -lean_dec_ref(x_25); -if (lean_is_scalar(x_13)) { - x_30 = lean_alloc_ctor(1, 1, 0); -} else { - x_30 = x_13; -} -lean_ctor_set(x_30, 0, x_29); -return x_30; -} -} -} -else -{ -lean_object* x_43; -lean_dec(x_11); -lean_dec_ref(x_9); -lean_dec_ref(x_5); -lean_dec_ref(x_3); -x_43 = lean_box(0); -return x_43; -} -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -uint8_t x_10; lean_object* x_11; -x_10 = lean_unbox(x_8); -x_11 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_10, x_9); -lean_dec_ref(x_7); -lean_dec_ref(x_6); -lean_dec_ref(x_4); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -return x_11; -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_14 = l_Lean_Expr_fvarId_x21(x_1); -x_15 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__13___closed__2; -lean_inc(x_2); -x_16 = lean_array_push(x_15, x_2); -lean_inc(x_14); -x_17 = lean_array_push(x_16, x_14); -x_18 = lean_mk_empty_array_with_capacity(x_3); -x_19 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody(x_17, x_18, x_6, x_2, x_14, x_4, x_5, x_7, x_8, x_9, x_10, x_11, x_12); -return x_19; -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: +lean_dec_ref(x_2); +x_37 = !lean_is_exclusive(x_9); +if (x_37 == 0) { -lean_object* x_14; -x_14 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_3); -lean_dec_ref(x_1); -return x_14; -} +return x_9; } -static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__0() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("replacing ", 10, 10); -return x_1; +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_9, 0); +lean_inc(x_38); +lean_dec(x_9); +x_39 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_39, 0, x_38); +return x_39; } } -static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__0; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; } } -static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__2() { +LEAN_EXPORT lean_object* l_Lean_Meta_withErasedFVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__18(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked(" with ", 6, 6); -return x_1; +lean_object* x_9; +x_9 = l_Lean_Meta_withErasedFVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__18___redArg(x_2, x_3, x_4, x_5, x_6, x_7); +return x_9; } } -static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__3() { +LEAN_EXPORT lean_object* l_Lean_Meta_withLCtx___at___00Lean_Meta_withErasedFVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__18_spec__33(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__2; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_10; +x_10 = l_Lean_Meta_withLCtx___at___00Lean_Meta_withErasedFVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__18_spec__33___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_10; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__21___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_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_130; -lean_inc(x_1); -x_130 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__1___redArg(x_1, x_12, x_13, x_16); -if (lean_obj_tag(x_130) == 0) -{ -lean_object* x_131; lean_object* x_132; lean_object* x_133; uint8_t x_134; -x_131 = lean_ctor_get(x_130, 0); -lean_inc(x_131); -lean_dec_ref(x_130); -x_132 = lean_ctor_get(x_131, 0); -lean_inc(x_132); -x_133 = lean_ctor_get(x_132, 0); -x_134 = lean_unbox(x_133); -if (x_134 == 0) +lean_object* x_7; lean_object* x_8; uint8_t x_13; +x_13 = lean_nat_dec_lt(x_4, x_1); +if (x_13 == 0) { -lean_object* x_135; lean_object* x_136; -lean_dec(x_1); -x_135 = lean_ctor_get(x_131, 1); -lean_inc(x_135); -lean_dec(x_131); -x_136 = lean_ctor_get(x_132, 1); -lean_inc(x_136); -lean_dec(x_132); -x_49 = x_136; -x_50 = x_135; -x_51 = x_14; -x_52 = x_15; -x_53 = x_16; -x_54 = x_17; -x_55 = lean_box(0); -goto block_129; +lean_object* x_14; +lean_dec(x_4); +x_14 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_14, 0, x_5); +return x_14; } else { -uint8_t x_137; -x_137 = !lean_is_exclusive(x_131); -if (x_137 == 0) -{ -lean_object* x_138; lean_object* x_139; uint8_t x_140; -x_138 = lean_ctor_get(x_131, 1); -x_139 = lean_ctor_get(x_131, 0); -lean_dec(x_139); -x_140 = !lean_is_exclusive(x_132); -if (x_140 == 0) -{ -lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; -x_141 = lean_ctor_get(x_132, 1); -x_142 = lean_ctor_get(x_132, 0); -lean_dec(x_142); -x_143 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__1; -lean_inc(x_10); -x_144 = l_Lean_Expr_fvar___override(x_10); -x_145 = l_Lean_MessageData_ofExpr(x_144); -lean_ctor_set_tag(x_132, 7); -lean_ctor_set(x_132, 1, x_145); -lean_ctor_set(x_132, 0, x_143); -x_146 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__3; -lean_ctor_set_tag(x_131, 7); -lean_ctor_set(x_131, 1, x_146); -lean_inc_ref(x_4); -x_147 = l_Lean_MessageData_ofExpr(x_4); -x_148 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_148, 0, x_131); -lean_ctor_set(x_148, 1, x_147); -x_149 = l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__2(x_1, x_148, x_141, x_138, x_14, x_15, x_16, x_17); -if (lean_obj_tag(x_149) == 0) -{ -lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; -x_150 = lean_ctor_get(x_149, 0); -lean_inc(x_150); -lean_dec_ref(x_149); -x_151 = lean_ctor_get(x_150, 0); -lean_inc(x_151); -x_152 = lean_ctor_get(x_150, 1); -lean_inc(x_152); -lean_dec(x_150); -x_153 = lean_ctor_get(x_151, 1); -lean_inc(x_153); -lean_dec(x_151); -x_49 = x_153; -x_50 = x_152; -x_51 = x_14; -x_52 = x_15; -x_53 = x_16; -x_54 = x_17; -x_55 = lean_box(0); -goto block_129; -} -else +lean_object* x_15; +x_15 = lean_array_fget_borrowed(x_2, x_4); +if (lean_obj_tag(x_15) == 1) { -uint8_t x_154; -lean_dec(x_17); -lean_dec_ref(x_16); -lean_dec(x_15); -lean_dec_ref(x_14); -lean_dec_ref(x_11); -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec_ref(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -x_154 = !lean_is_exclusive(x_149); -if (x_154 == 0) +lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_16 = lean_ctor_get(x_15, 0); +x_17 = 0; +x_18 = lean_box(x_17); +x_19 = lean_array_get_borrowed(x_18, x_3, x_16); +x_20 = lean_unbox(x_19); +if (x_20 == 0) { -return x_149; +uint8_t x_21; lean_object* x_22; lean_object* x_23; +x_21 = 0; +x_22 = lean_box(x_21); +x_23 = lean_array_push(x_5, x_22); +x_7 = x_23; +x_8 = lean_box(0); +goto block_12; } else { -lean_object* x_155; lean_object* x_156; -x_155 = lean_ctor_get(x_149, 0); -lean_inc(x_155); -lean_dec(x_149); -x_156 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_156, 0, x_155); -return x_156; -} +uint8_t x_24; lean_object* x_25; lean_object* x_26; +x_24 = 1; +x_25 = lean_box(x_24); +x_26 = lean_array_push(x_5, x_25); +x_7 = x_26; +x_8 = lean_box(0); +goto block_12; } } else { -lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; -x_157 = lean_ctor_get(x_132, 1); -lean_inc(x_157); -lean_dec(x_132); -x_158 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__1; -lean_inc(x_10); -x_159 = l_Lean_Expr_fvar___override(x_10); -x_160 = l_Lean_MessageData_ofExpr(x_159); -x_161 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_161, 0, x_158); -lean_ctor_set(x_161, 1, x_160); -x_162 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__3; -lean_ctor_set_tag(x_131, 7); -lean_ctor_set(x_131, 1, x_162); -lean_ctor_set(x_131, 0, x_161); -lean_inc_ref(x_4); -x_163 = l_Lean_MessageData_ofExpr(x_4); -x_164 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_164, 0, x_131); -lean_ctor_set(x_164, 1, x_163); -x_165 = l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__2(x_1, x_164, x_157, x_138, x_14, x_15, x_16, x_17); -if (lean_obj_tag(x_165) == 0) -{ -lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; -x_166 = lean_ctor_get(x_165, 0); -lean_inc(x_166); -lean_dec_ref(x_165); -x_167 = lean_ctor_get(x_166, 0); -lean_inc(x_167); -x_168 = lean_ctor_get(x_166, 1); -lean_inc(x_168); -lean_dec(x_166); -x_169 = lean_ctor_get(x_167, 1); -lean_inc(x_169); -lean_dec(x_167); -x_49 = x_169; -x_50 = x_168; -x_51 = x_14; -x_52 = x_15; -x_53 = x_16; -x_54 = x_17; -x_55 = lean_box(0); -goto block_129; +uint8_t x_27; lean_object* x_28; lean_object* x_29; +x_27 = 2; +x_28 = lean_box(x_27); +x_29 = lean_array_push(x_5, x_28); +x_7 = x_29; +x_8 = lean_box(0); +goto block_12; } -else +} +block_12: { -lean_object* x_170; lean_object* x_171; lean_object* x_172; -lean_dec(x_17); -lean_dec_ref(x_16); -lean_dec(x_15); -lean_dec_ref(x_14); -lean_dec_ref(x_11); -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec_ref(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -x_170 = lean_ctor_get(x_165, 0); -lean_inc(x_170); -if (lean_is_exclusive(x_165)) { - lean_ctor_release(x_165, 0); - x_171 = x_165; -} else { - lean_dec_ref(x_165); - x_171 = lean_box(0); +lean_object* x_9; lean_object* x_10; +x_9 = lean_unsigned_to_nat(1u); +x_10 = lean_nat_add(x_4, x_9); +lean_dec(x_4); +x_4 = x_10; +x_5 = x_7; +goto _start; } -if (lean_is_scalar(x_171)) { - x_172 = lean_alloc_ctor(1, 1, 0); -} else { - x_172 = x_171; } -lean_ctor_set(x_172, 0, x_170); -return x_172; } +LEAN_EXPORT lean_object* l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__21(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_14; +x_14 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__21___redArg(x_1, x_2, x_3, x_6, x_7); +return x_14; } } -else +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Elab_Structural_Positions_groupAndSort___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__6_spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, size_t x_7, size_t x_8, lean_object* x_9) { +_start: { -lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; -x_173 = lean_ctor_get(x_131, 1); -lean_inc(x_173); -lean_dec(x_131); -x_174 = lean_ctor_get(x_132, 1); -lean_inc(x_174); -if (lean_is_exclusive(x_132)) { - lean_ctor_release(x_132, 0); - lean_ctor_release(x_132, 1); - x_175 = x_132; -} else { - lean_dec_ref(x_132); - x_175 = lean_box(0); +lean_object* x_10; +x_10 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Elab_Structural_Positions_groupAndSort___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__6_spec__7___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_10; } -x_176 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__1; -lean_inc(x_10); -x_177 = l_Lean_Expr_fvar___override(x_10); -x_178 = l_Lean_MessageData_ofExpr(x_177); -if (lean_is_scalar(x_175)) { - x_179 = lean_alloc_ctor(7, 2, 0); -} else { - x_179 = x_175; - lean_ctor_set_tag(x_179, 7); } -lean_ctor_set(x_179, 0, x_176); -lean_ctor_set(x_179, 1, x_178); -x_180 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__3; -x_181 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_181, 0, x_179); -lean_ctor_set(x_181, 1, x_180); -lean_inc_ref(x_4); -x_182 = l_Lean_MessageData_ofExpr(x_4); -x_183 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_183, 0, x_181); -lean_ctor_set(x_183, 1, x_182); -x_184 = l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionCase_spec__2(x_1, x_183, x_174, x_173, x_14, x_15, x_16, x_17); -if (lean_obj_tag(x_184) == 0) +LEAN_EXPORT uint8_t l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Array_contains___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__14_spec__25(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { +_start: { -lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; -x_185 = lean_ctor_get(x_184, 0); -lean_inc(x_185); -lean_dec_ref(x_184); -x_186 = lean_ctor_get(x_185, 0); -lean_inc(x_186); -x_187 = lean_ctor_get(x_185, 1); -lean_inc(x_187); -lean_dec(x_185); -x_188 = lean_ctor_get(x_186, 1); -lean_inc(x_188); -lean_dec(x_186); -x_49 = x_188; -x_50 = x_187; -x_51 = x_14; -x_52 = x_15; -x_53 = x_16; -x_54 = x_17; -x_55 = lean_box(0); -goto block_129; +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_3, x_4); +if (x_5 == 0) +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_array_uget(x_2, x_3); +x_7 = lean_nat_dec_eq(x_1, x_6); +lean_dec(x_6); +if (x_7 == 0) +{ +size_t x_8; size_t x_9; +x_8 = 1; +x_9 = lean_usize_add(x_3, x_8); +x_3 = x_9; +goto _start; } else { -lean_object* x_189; lean_object* x_190; lean_object* x_191; -lean_dec(x_17); -lean_dec_ref(x_16); -lean_dec(x_15); -lean_dec_ref(x_14); -lean_dec_ref(x_11); -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec_ref(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -x_189 = lean_ctor_get(x_184, 0); -lean_inc(x_189); -if (lean_is_exclusive(x_184)) { - lean_ctor_release(x_184, 0); - x_190 = x_184; -} else { - lean_dec_ref(x_184); - x_190 = lean_box(0); +return x_7; } -if (lean_is_scalar(x_190)) { - x_191 = lean_alloc_ctor(1, 1, 0); -} else { - x_191 = x_190; } -lean_ctor_set(x_191, 0, x_189); -return x_191; +else +{ +uint8_t x_11; +x_11 = 0; +return x_11; } } } +LEAN_EXPORT uint8_t l_Array_contains___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__14(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = lean_unsigned_to_nat(0u); +x_4 = lean_array_get_size(x_1); +x_5 = lean_nat_dec_lt(x_3, x_4); +if (x_5 == 0) +{ +return x_5; } else { -uint8_t x_192; -lean_dec(x_17); -lean_dec_ref(x_16); -lean_dec(x_15); -lean_dec_ref(x_14); -lean_dec_ref(x_11); -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec_ref(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -lean_dec(x_1); -x_192 = !lean_is_exclusive(x_130); -if (x_192 == 0) +if (x_5 == 0) { -return x_130; +return x_5; } else { -lean_object* x_193; lean_object* x_194; -x_193 = lean_ctor_get(x_130, 0); -lean_inc(x_193); -lean_dec(x_130); -x_194 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_194, 0, x_193); -return x_194; +size_t x_6; size_t x_7; uint8_t x_8; +x_6 = 0; +x_7 = lean_usize_of_nat(x_4); +x_8 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Array_contains___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__14_spec__25(x_2, x_1, x_6, x_7); +return x_8; } } -block_48: -{ -uint8_t x_28; uint8_t x_29; lean_object* x_30; -x_28 = 0; -x_29 = 1; -x_30 = l_Lean_Meta_mkLambdaFVars(x_20, x_19, x_28, x_2, x_28, x_2, x_29, x_23, x_24, x_25, x_26); -lean_dec_ref(x_20); -if (lean_obj_tag(x_30) == 0) -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -lean_dec_ref(x_30); -x_32 = lean_array_push(x_3, x_4); -x_33 = l_Lean_Meta_mkLambdaFVars(x_32, x_31, x_28, x_2, x_28, x_2, x_29, x_23, x_24, x_25, x_26); -lean_dec(x_26); -lean_dec_ref(x_25); -lean_dec(x_24); -lean_dec_ref(x_23); -lean_dec_ref(x_32); -if (lean_obj_tag(x_33) == 0) +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__15_spec__27(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { +_start: { -uint8_t x_34; -x_34 = !lean_is_exclusive(x_33); -if (x_34 == 0) +uint8_t x_8; +x_8 = lean_usize_dec_lt(x_6, x_5); +if (x_8 == 0) { -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_33, 0); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_21); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_22); -lean_ctor_set(x_33, 0, x_37); -return x_33; +lean_inc_ref(x_7); +return x_7; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_38 = lean_ctor_get(x_33, 0); -lean_inc(x_38); -lean_dec(x_33); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_21); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_22); -x_41 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_41, 0, x_40); -return x_41; +lean_object* x_9; uint8_t x_10; +x_9 = lean_array_uget(x_4, x_6); +x_10 = l_Array_contains___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__14(x_9, x_1); +if (x_10 == 0) +{ +size_t x_11; size_t x_12; +lean_dec(x_9); +x_11 = 1; +x_12 = lean_usize_add(x_6, x_11); +{ +size_t _tmp_5 = x_12; +lean_object* _tmp_6 = x_2; +x_6 = _tmp_5; +x_7 = _tmp_6; } +goto _start; } else { -uint8_t x_42; -lean_dec_ref(x_22); -lean_dec_ref(x_21); -x_42 = !lean_is_exclusive(x_33); -if (x_42 == 0) -{ -return x_33; +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_9); +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_14); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_3); +return x_16; +} } -else -{ -lean_object* x_43; lean_object* x_44; -x_43 = lean_ctor_get(x_33, 0); -lean_inc(x_43); -lean_dec(x_33); -x_44 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_44, 0, x_43); -return x_44; } } +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__15___redArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { +_start: +{ +uint8_t x_8; +x_8 = lean_usize_dec_lt(x_6, x_5); +if (x_8 == 0) +{ +lean_inc_ref(x_7); +return x_7; } else { -uint8_t x_45; -lean_dec(x_26); -lean_dec_ref(x_25); -lean_dec(x_24); -lean_dec_ref(x_23); -lean_dec_ref(x_22); -lean_dec_ref(x_21); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -x_45 = !lean_is_exclusive(x_30); -if (x_45 == 0) +lean_object* x_9; uint8_t x_10; +x_9 = lean_array_uget(x_4, x_6); +x_10 = l_Array_contains___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__14(x_9, x_1); +if (x_10 == 0) { -return x_30; +size_t x_11; size_t x_12; lean_object* x_13; +lean_dec(x_9); +x_11 = 1; +x_12 = lean_usize_add(x_6, x_11); +x_13 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__15_spec__27(x_1, x_2, x_3, x_4, x_5, x_12, x_2); +return x_13; } else { -lean_object* x_46; lean_object* x_47; -x_46 = lean_ctor_get(x_30, 0); -lean_inc(x_46); -lean_dec(x_30); -x_47 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_47, 0, x_46); -return x_47; +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_9); +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_14); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_3); +return x_16; } } } -block_129: -{ -lean_object* x_56; lean_object* x_57; -x_56 = l_Subarray_toArray___redArg(x_5); -lean_inc(x_54); -lean_inc_ref(x_53); -lean_inc(x_52); -lean_inc_ref(x_51); -x_57 = l_Lean_Meta_instantiateLambda(x_11, x_56, x_51, x_52, x_53, 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; lean_object* x_62; -x_58 = lean_ctor_get(x_57, 0); -lean_inc(x_58); -lean_dec_ref(x_57); -lean_inc(x_58); -lean_inc(x_10); -lean_inc_ref(x_4); -x_59 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__1___boxed), 13, 5); -lean_closure_set(x_59, 0, x_4); -lean_closure_set(x_59, 1, x_10); -lean_closure_set(x_59, 2, x_6); -lean_closure_set(x_59, 3, x_7); -lean_closure_set(x_59, 4, x_58); -x_60 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun___boxed), 7, 1); -lean_closure_set(x_60, 0, x_8); -x_61 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_inProdLambdaLastArg___boxed), 7, 1); -lean_closure_set(x_61, 0, x_60); -lean_inc(x_54); -lean_inc_ref(x_53); -lean_inc(x_52); -lean_inc_ref(x_51); -x_62 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_withRewrittenMotive(x_9, x_61, x_59, x_49, x_50, x_51, x_52, x_53, x_54); -if (lean_obj_tag(x_62) == 0) -{ -lean_object* x_63; uint8_t x_64; -x_63 = lean_ctor_get(x_62, 0); -lean_inc(x_63); -lean_dec_ref(x_62); -x_64 = !lean_is_exclusive(x_63); -if (x_64 == 0) +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__15(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, lean_object* x_8) { +_start: { -lean_object* x_65; uint8_t x_66; -x_65 = lean_ctor_get(x_63, 0); -x_66 = !lean_is_exclusive(x_65); -if (x_66 == 0) +lean_object* x_9; +x_9 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__15___redArg(x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_9; +} +} +LEAN_EXPORT uint8_t l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__16___lam__0(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_67; lean_object* x_68; lean_object* x_69; uint8_t x_70; -x_67 = lean_ctor_get(x_63, 1); -x_68 = lean_ctor_get(x_65, 0); -x_69 = lean_ctor_get(x_65, 1); -x_70 = l_Lean_Expr_containsFVar(x_68, x_10); -if (x_70 == 0) +uint8_t x_3; +x_3 = lean_nat_dec_eq(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT uint8_t l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__16___lam__1(lean_object* x_1, lean_object* x_2) { +_start: { -lean_free_object(x_65); -lean_free_object(x_63); -lean_dec(x_58); -lean_dec(x_10); -x_19 = x_68; -x_20 = x_56; -x_21 = x_69; -x_22 = x_67; -x_23 = x_51; -x_24 = x_52; -x_25 = x_53; -x_26 = x_54; -x_27 = lean_box(0); -goto block_48; +uint8_t x_3; +x_3 = l_Array_contains___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__14(x_2, x_1); +return x_3; } -else +} +LEAN_EXPORT lean_object* l_Array_idxOfAux___at___00Array_finIdxOf_x3f___at___00Array_idxOf_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__11_spec__18_spec__21(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; -lean_dec(x_69); -lean_dec(x_67); -x_71 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__5___closed__1; -x_72 = l_Lean_mkFVar(x_10); -x_73 = l_Lean_MessageData_ofExpr(x_72); -lean_ctor_set_tag(x_65, 7); -lean_ctor_set(x_65, 1, x_73); -lean_ctor_set(x_65, 0, x_71); -x_74 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__5___closed__3; -lean_ctor_set_tag(x_63, 7); -lean_ctor_set(x_63, 1, x_74); -x_75 = l_Lean_indentExpr(x_58); -x_76 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_76, 0, x_63); -lean_ctor_set(x_76, 1, x_75); -x_77 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__0___redArg(x_76, x_51, x_52, x_53, x_54); -if (lean_obj_tag(x_77) == 0) +lean_object* x_4; uint8_t x_5; +x_4 = lean_array_get_size(x_1); +x_5 = lean_nat_dec_lt(x_3, x_4); +if (x_5 == 0) { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_78 = lean_ctor_get(x_77, 0); -lean_inc(x_78); -lean_dec_ref(x_77); -x_79 = lean_ctor_get(x_78, 0); -lean_inc(x_79); -x_80 = lean_ctor_get(x_78, 1); -lean_inc(x_80); -lean_dec(x_78); -x_81 = lean_ctor_get(x_79, 1); -lean_inc(x_81); -lean_dec(x_79); -x_19 = x_68; -x_20 = x_56; -x_21 = x_81; -x_22 = x_80; -x_23 = x_51; -x_24 = x_52; -x_25 = x_53; -x_26 = x_54; -x_27 = lean_box(0); -goto block_48; +lean_object* x_6; +lean_dec(x_3); +x_6 = lean_box(0); +return x_6; } else { -uint8_t x_82; -lean_dec(x_68); -lean_dec_ref(x_56); -lean_dec(x_54); -lean_dec_ref(x_53); -lean_dec(x_52); -lean_dec_ref(x_51); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -x_82 = !lean_is_exclusive(x_77); -if (x_82 == 0) +lean_object* x_7; uint8_t x_8; +x_7 = lean_array_fget_borrowed(x_1, x_3); +x_8 = lean_expr_eqv(x_7, x_2); +if (x_8 == 0) { -return x_77; +lean_object* x_9; lean_object* x_10; +x_9 = lean_unsigned_to_nat(1u); +x_10 = lean_nat_add(x_3, x_9); +lean_dec(x_3); +x_3 = x_10; +goto _start; } else { -lean_object* x_83; lean_object* x_84; -x_83 = lean_ctor_get(x_77, 0); -lean_inc(x_83); -lean_dec(x_77); -x_84 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_84, 0, x_83); -return x_84; +lean_object* x_12; +x_12 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_12, 0, x_3); +return x_12; } } } } -else +LEAN_EXPORT lean_object* l_Array_finIdxOf_x3f___at___00Array_idxOf_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__11_spec__18(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_85; lean_object* x_86; lean_object* x_87; uint8_t x_88; -x_85 = lean_ctor_get(x_63, 1); -x_86 = lean_ctor_get(x_65, 0); -x_87 = lean_ctor_get(x_65, 1); -lean_inc(x_87); -lean_inc(x_86); -lean_dec(x_65); -x_88 = l_Lean_Expr_containsFVar(x_86, x_10); -if (x_88 == 0) +lean_object* x_3; lean_object* x_4; +x_3 = lean_unsigned_to_nat(0u); +x_4 = l_Array_idxOfAux___at___00Array_finIdxOf_x3f___at___00Array_idxOf_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__11_spec__18_spec__21(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Array_idxOf_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__11(lean_object* x_1, lean_object* x_2) { +_start: { -lean_free_object(x_63); -lean_dec(x_58); -lean_dec(x_10); -x_19 = x_86; -x_20 = x_56; -x_21 = x_87; -x_22 = x_85; -x_23 = x_51; -x_24 = x_52; -x_25 = x_53; -x_26 = x_54; -x_27 = lean_box(0); -goto block_48; +lean_object* x_3; +x_3 = l_Array_finIdxOf_x3f___at___00Array_idxOf_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__11_spec__18(x_1, x_2); +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; +x_4 = lean_box(0); +return x_4; } else { -lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; -lean_dec(x_87); -lean_dec(x_85); -x_89 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__5___closed__1; -x_90 = l_Lean_mkFVar(x_10); -x_91 = l_Lean_MessageData_ofExpr(x_90); -x_92 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_92, 0, x_89); -lean_ctor_set(x_92, 1, x_91); -x_93 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__5___closed__3; -lean_ctor_set_tag(x_63, 7); -lean_ctor_set(x_63, 1, x_93); -lean_ctor_set(x_63, 0, x_92); -x_94 = l_Lean_indentExpr(x_58); -x_95 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_95, 0, x_63); -lean_ctor_set(x_95, 1, x_94); -x_96 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__0___redArg(x_95, x_51, x_52, x_53, x_54); -if (lean_obj_tag(x_96) == 0) +uint8_t x_5; +x_5 = !lean_is_exclusive(x_3); +if (x_5 == 0) { -lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_97 = lean_ctor_get(x_96, 0); -lean_inc(x_97); -lean_dec_ref(x_96); -x_98 = lean_ctor_get(x_97, 0); -lean_inc(x_98); -x_99 = lean_ctor_get(x_97, 1); -lean_inc(x_99); -lean_dec(x_97); -x_100 = lean_ctor_get(x_98, 1); -lean_inc(x_100); -lean_dec(x_98); -x_19 = x_86; -x_20 = x_56; -x_21 = x_100; -x_22 = x_99; -x_23 = x_51; -x_24 = x_52; -x_25 = x_53; -x_26 = x_54; -x_27 = lean_box(0); -goto block_48; +return x_3; } else { -lean_object* x_101; lean_object* x_102; lean_object* x_103; -lean_dec(x_86); -lean_dec_ref(x_56); -lean_dec(x_54); -lean_dec_ref(x_53); -lean_dec(x_52); -lean_dec_ref(x_51); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -x_101 = lean_ctor_get(x_96, 0); -lean_inc(x_101); -if (lean_is_exclusive(x_96)) { - lean_ctor_release(x_96, 0); - x_102 = x_96; -} else { - lean_dec_ref(x_96); - x_102 = lean_box(0); -} -if (lean_is_scalar(x_102)) { - x_103 = lean_alloc_ctor(1, 1, 0); -} else { - x_103 = x_102; -} -lean_ctor_set(x_103, 0, x_101); -return x_103; +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_3, 0); +lean_inc(x_6); +lean_dec(x_3); +x_7 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_7, 0, x_6); +return x_7; } } } } -else +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, uint8_t x_8, lean_object* x_9) { +_start: { -lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; uint8_t x_109; -x_104 = lean_ctor_get(x_63, 0); -x_105 = lean_ctor_get(x_63, 1); -lean_inc(x_105); -lean_inc(x_104); -lean_dec(x_63); -x_106 = lean_ctor_get(x_104, 0); -lean_inc(x_106); -x_107 = lean_ctor_get(x_104, 1); -lean_inc(x_107); -if (lean_is_exclusive(x_104)) { - lean_ctor_release(x_104, 0); - lean_ctor_release(x_104, 1); - x_108 = x_104; +lean_object* x_10; lean_object* x_11; +x_10 = l_Lean_Expr_getAppFn(x_9); +x_11 = l_Array_idxOf_x3f___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__11(x_1, x_10); +lean_dec_ref(x_10); +if (lean_obj_tag(x_11) == 1) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +if (lean_is_exclusive(x_11)) { + lean_ctor_release(x_11, 0); + x_13 = x_11; } else { - lean_dec_ref(x_104); - x_108 = lean_box(0); + lean_dec_ref(x_11); + x_13 = lean_box(0); } -x_109 = l_Lean_Expr_containsFVar(x_106, x_10); -if (x_109 == 0) +x_14 = l_Lean_Expr_getAppNumArgs(x_9); +x_15 = lean_unsigned_to_nat(0u); +x_16 = lean_array_get_borrowed(x_15, x_2, x_12); +x_17 = lean_nat_dec_eq(x_14, x_16); +if (x_17 == 0) { -lean_dec(x_108); -lean_dec(x_58); -lean_dec(x_10); -x_19 = x_106; -x_20 = x_56; -x_21 = x_107; -x_22 = x_105; -x_23 = x_51; -x_24 = x_52; -x_25 = x_53; -x_26 = x_54; -x_27 = lean_box(0); -goto block_48; +lean_object* x_18; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec_ref(x_9); +lean_dec_ref(x_5); +lean_dec_ref(x_3); +x_18 = lean_box(0); +return x_18; } else { -lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; -lean_dec(x_107); -lean_dec(x_105); -x_110 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__5___closed__1; -x_111 = l_Lean_mkFVar(x_10); -x_112 = l_Lean_MessageData_ofExpr(x_111); -if (lean_is_scalar(x_108)) { - x_113 = lean_alloc_ctor(7, 2, 0); -} else { - x_113 = x_108; - lean_ctor_set_tag(x_113, 7); -} -lean_ctor_set(x_113, 0, x_110); -lean_ctor_set(x_113, 1, x_112); -x_114 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize___lam__5___closed__3; -x_115 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_115, 0, x_113); -lean_ctor_set(x_115, 1, x_114); -x_116 = l_Lean_indentExpr(x_58); -x_117 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_117, 0, x_115); -lean_ctor_set(x_117, 1, x_116); -x_118 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__0___redArg(x_117, x_51, x_52, x_53, x_54); -if (lean_obj_tag(x_118) == 0) +lean_object* x_19; lean_object* x_20; +x_19 = lean_array_get_borrowed(x_3, x_4, x_12); +if (x_8 == 0) { -lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; -x_119 = lean_ctor_get(x_118, 0); -lean_inc(x_119); -lean_dec_ref(x_118); -x_120 = lean_ctor_get(x_119, 0); -lean_inc(x_120); -x_121 = lean_ctor_get(x_119, 1); -lean_inc(x_121); -lean_dec(x_119); -x_122 = lean_ctor_get(x_120, 1); -lean_inc(x_122); -lean_dec(x_120); -x_19 = x_106; -x_20 = x_56; -x_21 = x_122; -x_22 = x_121; -x_23 = x_51; -x_24 = x_52; -x_25 = x_53; -x_26 = x_54; -x_27 = lean_box(0); -goto block_48; +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_32 = l___private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0___lam__1___closed__0; +lean_inc(x_14); +x_33 = lean_mk_array(x_14, x_32); +x_34 = lean_unsigned_to_nat(1u); +x_35 = lean_nat_sub(x_14, x_34); +lean_dec(x_14); +x_36 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_9, x_33, x_35); +x_20 = x_36; +goto block_31; } else { -lean_object* x_123; lean_object* x_124; lean_object* x_125; -lean_dec(x_106); -lean_dec_ref(x_56); -lean_dec(x_54); -lean_dec_ref(x_53); -lean_dec(x_52); -lean_dec_ref(x_51); -lean_dec_ref(x_4); -lean_dec_ref(x_3); -x_123 = lean_ctor_get(x_118, 0); -lean_inc(x_123); -if (lean_is_exclusive(x_118)) { - lean_ctor_release(x_118, 0); - x_124 = x_118; -} else { - lean_dec_ref(x_118); - x_124 = lean_box(0); +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_37 = l___private_Lean_Meta_Transform_0__Lean_Core_transform_visit___at___00Lean_Core_transform___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_elimTypeAnnotations_spec__0_spec__0___lam__1___closed__0; +lean_inc(x_14); +x_38 = lean_mk_array(x_14, x_37); +x_39 = lean_unsigned_to_nat(1u); +x_40 = lean_nat_sub(x_14, x_39); +lean_dec(x_14); +x_41 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_9, x_38, x_40); +x_42 = lean_array_pop(x_41); +x_20 = x_42; +goto block_31; } -if (lean_is_scalar(x_124)) { - x_125 = lean_alloc_ctor(1, 1, 0); +block_31: +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_21 = lean_ctor_get(x_19, 0); +x_22 = lean_ctor_get(x_21, 0); +x_23 = lean_ctor_get(x_21, 1); +x_24 = lean_array_get_borrowed(x_5, x_6, x_12); +lean_dec(x_12); +lean_inc(x_24); +x_25 = l_Lean_Elab_FixedParamPerm_buildArgs___redArg(x_24, x_7, x_20); +x_26 = lean_box(0); +lean_inc(x_23); +x_27 = l_List_mapTR_loop___at___00Lean_mkConstWithLevelParams___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__2_spec__4(x_23, x_26); +lean_inc(x_22); +x_28 = l_Lean_Expr_const___override(x_22, x_27); +x_29 = l_Lean_mkAppN(x_28, x_25); +lean_dec_ref(x_25); +if (lean_is_scalar(x_13)) { + x_30 = lean_alloc_ctor(1, 1, 0); } else { - x_125 = x_124; -} -lean_ctor_set(x_125, 0, x_123); -return x_125; + x_30 = x_13; } +lean_ctor_set(x_30, 0, x_29); +return x_30; } } } else { -lean_dec(x_58); -lean_dec_ref(x_56); -lean_dec(x_54); -lean_dec_ref(x_53); -lean_dec(x_52); -lean_dec_ref(x_51); -lean_dec(x_10); -lean_dec_ref(x_4); +lean_object* x_43; +lean_dec(x_11); +lean_dec_ref(x_9); +lean_dec_ref(x_5); lean_dec_ref(x_3); -return x_62; +x_43 = lean_box(0); +return x_43; } } -else +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__0___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_126; -lean_dec_ref(x_56); -lean_dec(x_54); -lean_dec_ref(x_53); -lean_dec(x_52); -lean_dec_ref(x_51); -lean_dec_ref(x_50); -lean_dec_ref(x_49); -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec_ref(x_8); +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_8); +x_11 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__0(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_10, x_9); lean_dec_ref(x_7); -lean_dec(x_6); +lean_dec_ref(x_6); lean_dec_ref(x_4); -lean_dec_ref(x_3); -x_126 = !lean_is_exclusive(x_57); -if (x_126 == 0) -{ -return x_57; -} -else -{ -lean_object* x_127; lean_object* x_128; -x_127 = lean_ctor_get(x_57, 0); -lean_inc(x_127); -lean_dec(x_57); -x_128 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_128, 0, x_127); -return x_128; -} +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_11; } } +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_14 = l_Lean_Expr_fvarId_x21(x_1); +x_15 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__13___closed__2; +lean_inc(x_2); +x_16 = lean_array_push(x_15, x_2); +lean_inc(x_14); +x_17 = lean_array_push(x_16, x_14); +x_18 = lean_mk_empty_array_with_capacity(x_3); +x_19 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody(x_17, x_18, x_6, x_2, x_14, x_4, x_5, x_7, x_8, x_9, x_10, x_11, x_12); +return x_19; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; -lean_object* x_18 = _args[17]; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -uint8_t x_19; lean_object* x_20; -x_19 = lean_unbox(x_2); -x_20 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2(x_1, x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -return x_20; +lean_object* x_14; +x_14 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_3); +lean_dec_ref(x_1); +return x_14; } } -static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___closed__0() { +static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___closed__0() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked(".brecOn argument has too few parameters, expected at least ", 59, 59); +x_1 = lean_mk_string_unchecked("replacing ", 10, 10); return x_1; } } -static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___closed__1() { +static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___closed__1() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___closed__0; +x_1 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___closed__0; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___closed__2() { +static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___closed__2() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked(": ", 2, 2); +x_1 = lean_mk_string_unchecked(" with ", 6, 6); return x_1; } } -static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___closed__3() { +static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___closed__2; +x_1 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___closed__2; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { -_start: -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_50; uint8_t x_51; -x_50 = lean_array_get_size(x_9); -x_51 = lean_nat_dec_le(x_1, x_50); -if (x_51 == 0) -{ -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -lean_dec_ref(x_12); -lean_dec_ref(x_11); -x_52 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___closed__1; -lean_inc(x_1); -x_53 = l_Nat_reprFast(x_1); -x_54 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_54, 0, x_53); -x_55 = l_Lean_MessageData_ofFormat(x_54); -x_56 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_56, 0, x_52); -lean_ctor_set(x_56, 1, x_55); -x_57 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___closed__3; -x_58 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_58, 0, x_56); -lean_ctor_set(x_58, 1, x_57); -lean_inc_ref(x_9); -x_59 = lean_array_to_list(x_9); -x_60 = lean_box(0); -x_61 = l_List_mapTR_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveUnaryInduction_doRealize_spec__4(x_59, x_60); -x_62 = l_Lean_MessageData_ofList(x_61); -x_63 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_63, 0, x_58); -lean_ctor_set(x_63, 1, x_62); -x_64 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__0___redArg(x_63, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_64) == 0) -{ -lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_65 = lean_ctor_get(x_64, 0); -lean_inc(x_65); -lean_dec_ref(x_64); -x_66 = lean_ctor_get(x_65, 0); -lean_inc(x_66); -x_67 = lean_ctor_get(x_65, 1); -lean_inc(x_67); -lean_dec(x_65); -x_68 = lean_ctor_get(x_66, 1); -lean_inc(x_68); -lean_dec(x_66); -x_37 = x_68; -x_38 = x_67; -x_39 = lean_box(0); -goto block_49; -} -else -{ -uint8_t x_69; -lean_dec(x_16); -lean_dec_ref(x_15); -lean_dec(x_14); -lean_dec_ref(x_13); -lean_dec_ref(x_10); -lean_dec_ref(x_9); -lean_dec_ref(x_7); -lean_dec_ref(x_6); -lean_dec(x_4); -lean_dec_ref(x_3); -lean_dec_ref(x_2); -lean_dec(x_1); -x_69 = !lean_is_exclusive(x_64); -if (x_69 == 0) -{ -return x_64; -} -else -{ -lean_object* x_70; lean_object* x_71; -x_70 = lean_ctor_get(x_64, 0); -lean_inc(x_70); -lean_dec(x_64); -x_71 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_71, 0, x_70); -return x_71; -} -} -} -else -{ -x_37 = x_11; -x_38 = x_12; -x_39 = lean_box(0); -goto block_49; -} -block_36: -{ -lean_object* x_25; -lean_inc(x_16); -lean_inc_ref(x_15); -lean_inc(x_14); -lean_inc_ref(x_13); -x_25 = l_Lean_Meta_instantiateLambda(x_3, x_22, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_25) == 0) -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -lean_dec_ref(x_25); -x_27 = lean_ctor_get(x_24, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_24, 1); -lean_inc(x_28); -lean_dec_ref(x_24); -x_29 = l_Array_toSubarray___redArg(x_9, x_27, x_28); -x_30 = lean_box(x_5); -x_31 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___boxed), 18, 9); -lean_closure_set(x_31, 0, x_4); -lean_closure_set(x_31, 1, x_30); -lean_closure_set(x_31, 2, x_22); -lean_closure_set(x_31, 3, x_18); -lean_closure_set(x_31, 4, x_29); -lean_closure_set(x_31, 5, x_19); -lean_closure_set(x_31, 6, x_6); -lean_closure_set(x_31, 7, x_7); -lean_closure_set(x_31, 8, x_10); -x_32 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__9___redArg(x_26, x_31, x_23, x_21, x_13, x_14, x_15, x_16); -return x_32; -} -else -{ -uint8_t x_33; -lean_dec_ref(x_24); -lean_dec_ref(x_23); -lean_dec_ref(x_22); -lean_dec_ref(x_21); -lean_dec(x_19); -lean_dec_ref(x_18); -lean_dec(x_16); -lean_dec_ref(x_15); -lean_dec(x_14); -lean_dec_ref(x_13); -lean_dec_ref(x_10); -lean_dec_ref(x_9); -lean_dec_ref(x_7); -lean_dec_ref(x_6); -lean_dec(x_4); -x_33 = !lean_is_exclusive(x_25); -if (x_33 == 0) -{ -return x_25; -} -else -{ -lean_object* x_34; lean_object* x_35; -x_34 = lean_ctor_get(x_25, 0); -lean_inc(x_34); -lean_dec(x_25); -x_35 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_35, 0, x_34); -return x_35; -} -} -} -block_49: -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; -x_40 = lean_unsigned_to_nat(0u); -lean_inc(x_1); -lean_inc_ref(x_9); -x_41 = l_Array_toSubarray___redArg(x_9, x_40, x_1); -x_42 = l_Subarray_toArray___redArg(x_41); -x_43 = lean_array_get_borrowed(x_2, x_9, x_1); -x_44 = lean_array_get_size(x_9); -x_45 = lean_nat_add(x_1, x_8); -lean_dec(x_1); -x_46 = lean_nat_dec_le(x_45, x_40); -if (x_46 == 0) -{ -lean_object* x_47; -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_44); -lean_inc(x_43); -x_18 = x_43; -x_19 = x_40; -x_20 = lean_box(0); -x_21 = x_38; -x_22 = x_42; -x_23 = x_37; -x_24 = x_47; -goto block_36; -} -else -{ -lean_object* x_48; -lean_dec(x_45); -x_48 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_48, 0, x_40); -lean_ctor_set(x_48, 1, x_44); -lean_inc(x_43); -x_18 = x_43; -x_19 = x_40; -x_20 = lean_box(0); -x_21 = x_38; -x_22 = x_42; -x_23 = x_37; -x_24 = x_48; -goto block_36; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; -_start: -{ -uint8_t x_18; lean_object* x_19; -x_18 = lean_unbox(x_5); -x_19 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3(x_1, x_2, x_3, x_4, x_18, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -lean_dec(x_8); -return x_19; -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { _start: { lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_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_130; @@ -105393,14 +97090,14 @@ lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; x_141 = lean_ctor_get(x_132, 1); x_142 = lean_ctor_get(x_132, 0); lean_dec(x_142); -x_143 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__1; +x_143 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___closed__1; lean_inc(x_10); x_144 = l_Lean_Expr_fvar___override(x_10); x_145 = l_Lean_MessageData_ofExpr(x_144); lean_ctor_set_tag(x_132, 7); lean_ctor_set(x_132, 1, x_145); lean_ctor_set(x_132, 0, x_143); -x_146 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__3; +x_146 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___closed__3; lean_ctor_set_tag(x_131, 7); lean_ctor_set(x_131, 1, x_146); lean_inc_ref(x_4); @@ -105471,14 +97168,14 @@ lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; x_157 = lean_ctor_get(x_132, 1); lean_inc(x_157); lean_dec(x_132); -x_158 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__1; +x_158 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___closed__1; lean_inc(x_10); x_159 = l_Lean_Expr_fvar___override(x_10); x_160 = l_Lean_MessageData_ofExpr(x_159); x_161 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_161, 0, x_158); lean_ctor_set(x_161, 1, x_160); -x_162 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__3; +x_162 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___closed__3; lean_ctor_set_tag(x_131, 7); lean_ctor_set(x_131, 1, x_162); lean_ctor_set(x_131, 0, x_161); @@ -105562,7 +97259,7 @@ if (lean_is_exclusive(x_132)) { lean_dec_ref(x_132); x_175 = lean_box(0); } -x_176 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__1; +x_176 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___closed__1; lean_inc(x_10); x_177 = l_Lean_Expr_fvar___override(x_10); x_178 = l_Lean_MessageData_ofExpr(x_177); @@ -105574,7 +97271,7 @@ if (lean_is_scalar(x_175)) { } lean_ctor_set(x_179, 0, x_176); lean_ctor_set(x_179, 1, x_178); -x_180 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__3; +x_180 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___closed__3; x_181 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_181, 0, x_179); lean_ctor_set(x_181, 1, x_180); @@ -105681,8 +97378,8 @@ return x_194; uint8_t x_28; uint8_t x_29; lean_object* x_30; x_28 = 0; x_29 = 1; -x_30 = l_Lean_Meta_mkLambdaFVars(x_19, x_20, x_28, x_2, x_28, x_2, x_29, x_23, x_24, x_25, x_26); -lean_dec_ref(x_19); +x_30 = l_Lean_Meta_mkLambdaFVars(x_20, x_19, x_28, x_2, x_28, x_2, x_29, x_23, x_24, x_25, x_26); +lean_dec_ref(x_20); if (lean_obj_tag(x_30) == 0) { lean_object* x_31; lean_object* x_32; lean_object* x_33; @@ -105838,8 +97535,8 @@ lean_free_object(x_65); lean_free_object(x_63); lean_dec(x_58); lean_dec(x_10); -x_19 = x_56; -x_20 = x_68; +x_19 = x_68; +x_20 = x_56; x_21 = x_69; x_22 = x_67; x_23 = x_51; @@ -105882,8 +97579,8 @@ lean_dec(x_78); x_81 = lean_ctor_get(x_79, 1); lean_inc(x_81); lean_dec(x_79); -x_19 = x_56; -x_20 = x_68; +x_19 = x_68; +x_20 = x_56; x_21 = x_81; x_22 = x_80; x_23 = x_51; @@ -105937,8 +97634,8 @@ if (x_88 == 0) lean_free_object(x_63); lean_dec(x_58); lean_dec(x_10); -x_19 = x_56; -x_20 = x_86; +x_19 = x_86; +x_20 = x_56; x_21 = x_87; x_22 = x_85; x_23 = x_51; @@ -105982,8 +97679,8 @@ lean_dec(x_97); x_100 = lean_ctor_get(x_98, 1); lean_inc(x_100); lean_dec(x_98); -x_19 = x_56; -x_20 = x_86; +x_19 = x_86; +x_20 = x_56; x_21 = x_100; x_22 = x_99; x_23 = x_51; @@ -106050,8 +97747,8 @@ if (x_109 == 0) lean_dec(x_108); lean_dec(x_58); lean_dec(x_10); -x_19 = x_56; -x_20 = x_106; +x_19 = x_106; +x_20 = x_56; x_21 = x_107; x_22 = x_105; x_23 = x_51; @@ -106100,8 +97797,8 @@ lean_dec(x_119); x_122 = lean_ctor_get(x_120, 1); lean_inc(x_122); lean_dec(x_120); -x_19 = x_56; -x_20 = x_106; +x_19 = x_106; +x_20 = x_56; x_21 = x_122; x_22 = x_121; x_23 = x_51; @@ -106192,7 +97889,7 @@ return x_128; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___boxed(lean_object** _args) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -106215,11 +97912,45 @@ lean_object* x_18 = _args[17]; { uint8_t x_19; lean_object* x_20; x_19 = lean_unbox(x_2); -x_20 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2(x_1, x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +x_20 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2(x_1, x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); return x_20; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(".brecOn argument has too few parameters, expected at least ", 59, 59); +return x_1; +} +} +static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___closed__0; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(": ", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_50; uint8_t x_51; @@ -106230,7 +97961,7 @@ if (x_51 == 0) lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_dec_ref(x_12); lean_dec_ref(x_11); -x_52 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___closed__1; +x_52 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___closed__1; lean_inc(x_1); x_53 = l_Nat_reprFast(x_1); x_54 = lean_alloc_ctor(3, 1, 0); @@ -106239,7 +97970,7 @@ x_55 = l_Lean_MessageData_ofFormat(x_54); x_56 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_56, 0, x_52); lean_ctor_set(x_56, 1, x_55); -x_57 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___closed__3; +x_57 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___closed__3; x_58 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_58, 0, x_56); lean_ctor_set(x_58, 1, x_57); @@ -106317,7 +98048,7 @@ lean_inc(x_16); lean_inc_ref(x_15); lean_inc(x_14); lean_inc_ref(x_13); -x_25 = l_Lean_Meta_instantiateLambda(x_3, x_18, x_13, x_14, x_15, x_16); +x_25 = l_Lean_Meta_instantiateLambda(x_3, x_23, x_13, x_14, x_15, x_16); if (lean_obj_tag(x_25) == 0) { lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; @@ -106331,17 +98062,17 @@ lean_inc(x_28); lean_dec_ref(x_24); x_29 = l_Array_toSubarray___redArg(x_9, x_27, x_28); x_30 = lean_box(x_5); -x_31 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___boxed), 18, 9); +x_31 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___boxed), 18, 9); lean_closure_set(x_31, 0, x_4); lean_closure_set(x_31, 1, x_30); -lean_closure_set(x_31, 2, x_18); -lean_closure_set(x_31, 3, x_20); +lean_closure_set(x_31, 2, x_23); +lean_closure_set(x_31, 3, x_21); lean_closure_set(x_31, 4, x_29); -lean_closure_set(x_31, 5, x_21); +lean_closure_set(x_31, 5, x_20); lean_closure_set(x_31, 6, x_6); lean_closure_set(x_31, 7, x_7); lean_closure_set(x_31, 8, x_10); -x_32 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__9___redArg(x_26, x_31, x_23, x_22, x_13, x_14, x_15, x_16); +x_32 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_lambdaTelescope1___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody_spec__9___redArg(x_26, x_31, x_18, x_19, x_13, x_14, x_15, x_16); return x_32; } else @@ -106349,9 +98080,9 @@ else uint8_t x_33; lean_dec_ref(x_24); lean_dec_ref(x_23); -lean_dec_ref(x_22); -lean_dec(x_21); -lean_dec_ref(x_20); +lean_dec_ref(x_21); +lean_dec(x_20); +lean_dec_ref(x_19); lean_dec_ref(x_18); lean_dec(x_16); lean_dec_ref(x_15); @@ -106399,12 +98130,12 @@ x_47 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_47, 0, x_45); lean_ctor_set(x_47, 1, x_44); lean_inc(x_43); -x_18 = x_42; -x_19 = lean_box(0); -x_20 = x_43; -x_21 = x_40; -x_22 = x_38; -x_23 = x_37; +x_18 = x_37; +x_19 = x_38; +x_20 = x_40; +x_21 = x_43; +x_22 = lean_box(0); +x_23 = x_42; x_24 = x_47; goto block_36; } @@ -106416,19 +98147,19 @@ x_48 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_48, 0, x_40); lean_ctor_set(x_48, 1, x_44); lean_inc(x_43); -x_18 = x_42; -x_19 = lean_box(0); -x_20 = x_43; -x_21 = x_40; -x_22 = x_38; -x_23 = x_37; +x_18 = x_37; +x_19 = x_38; +x_20 = x_40; +x_21 = x_43; +x_22 = lean_box(0); +x_23 = x_42; x_24 = x_48; goto block_36; } } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___boxed(lean_object** _args) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -106450,7 +98181,7 @@ lean_object* x_17 = _args[16]; { uint8_t x_18; lean_object* x_19; x_18 = lean_unbox(x_5); -x_19 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0(x_1, x_2, x_3, x_4, x_18, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +x_19 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3(x_1, x_2, x_3, x_4, x_18, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); lean_dec(x_8); return x_19; } @@ -106623,7 +98354,7 @@ x_65 = lean_box(x_56); lean_inc_ref(x_10); lean_inc(x_11); lean_inc_ref(x_1); -x_66 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___boxed), 17, 8); +x_66 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___boxed), 17, 8); lean_closure_set(x_66, 0, x_64); lean_closure_set(x_66, 1, x_1); lean_closure_set(x_66, 2, x_63); @@ -106739,7 +98470,7 @@ x_87 = lean_box(x_56); lean_inc_ref(x_10); lean_inc(x_11); lean_inc_ref(x_1); -x_88 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___boxed), 17, 8); +x_88 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___boxed), 17, 8); lean_closure_set(x_88, 0, x_86); lean_closure_set(x_88, 1, x_1); lean_closure_set(x_88, 2, x_85); @@ -106912,7 +98643,7 @@ x_118 = lean_box(x_112); lean_inc_ref(x_10); lean_inc(x_11); lean_inc_ref(x_1); -x_119 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___boxed), 17, 8); +x_119 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___boxed), 17, 8); lean_closure_set(x_119, 0, x_117); lean_closure_set(x_119, 1, x_1); lean_closure_set(x_119, 2, x_116); @@ -107148,7 +98879,7 @@ x_157 = lean_box(x_150); lean_inc_ref(x_10); lean_inc(x_11); lean_inc_ref(x_1); -x_158 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___boxed), 17, 8); +x_158 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___boxed), 17, 8); lean_closure_set(x_158, 0, x_156); lean_closure_set(x_158, 1, x_1); lean_closure_set(x_158, 2, x_155); @@ -107407,7 +99138,7 @@ x_201 = lean_box(x_193); lean_inc_ref(x_10); lean_inc(x_11); lean_inc_ref(x_1); -x_202 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___boxed), 17, 8); +x_202 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___boxed), 17, 8); lean_closure_set(x_202, 0, x_200); lean_closure_set(x_202, 1, x_1); lean_closure_set(x_202, 2, x_199); @@ -108755,7 +100486,7 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_Fun lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__6___closed__0; x_2 = lean_unsigned_to_nat(6u); -x_3 = lean_unsigned_to_nat(1302u); +x_3 = lean_unsigned_to_nat(1183u); x_4 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__16___lam__2___closed__0; x_5 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__17___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -108892,21 +100623,21 @@ goto block_41; uint8_t x_20; uint8_t x_21; lean_object* x_22; x_20 = 0; x_21 = 1; -x_22 = l_Lean_Meta_mkForallFVars(x_17, x_19, x_20, x_1, x_1, x_21, x_12, x_13, x_14, x_15); -lean_dec_ref(x_17); +x_22 = l_Lean_Meta_mkForallFVars(x_18, x_19, x_20, x_1, x_1, x_21, x_12, x_13, x_14, x_15); +lean_dec_ref(x_18); if (lean_obj_tag(x_22) == 0) { lean_object* x_23; lean_object* x_24; x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); lean_dec_ref(x_22); -x_24 = l_Lean_Meta_mkLambdaFVars(x_18, x_23, x_20, x_1, x_20, x_1, x_21, x_12, x_13, x_14, x_15); -lean_dec_ref(x_18); +x_24 = l_Lean_Meta_mkLambdaFVars(x_17, x_23, x_20, x_1, x_20, x_1, x_21, x_12, x_13, x_14, x_15); +lean_dec_ref(x_17); return x_24; } else { -lean_dec_ref(x_18); +lean_dec_ref(x_17); return x_22; } } @@ -108926,8 +100657,8 @@ if (x_4 == 0) lean_dec_ref(x_26); lean_dec_ref(x_6); lean_dec_ref(x_5); -x_17 = x_29; -x_18 = x_28; +x_17 = x_28; +x_18 = x_29; x_19 = x_30; goto block_25; } @@ -108951,8 +100682,8 @@ x_38 = l_Lean_Expr_const___override(x_32, x_37); x_39 = l_Lean_mkAppN(x_38, x_35); lean_dec_ref(x_35); x_40 = l_Lean_Expr_app___override(x_30, x_39); -x_17 = x_29; -x_18 = x_28; +x_17 = x_28; +x_18 = x_29; x_19 = x_40; goto block_25; } @@ -110291,7 +102022,7 @@ static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Un 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_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__16___lam__2___closed__1; x_2 = lean_unsigned_to_nat(67u); -x_3 = lean_unsigned_to_nat(1399u); +x_3 = lean_unsigned_to_nat(1280u); x_4 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__16___lam__2___closed__0; x_5 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__17___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -110333,7 +102064,7 @@ static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Un 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_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__16___lam__2___closed__1; x_2 = lean_unsigned_to_nat(62u); -x_3 = lean_unsigned_to_nat(1400u); +x_3 = lean_unsigned_to_nat(1281u); x_4 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__16___lam__2___closed__0; x_5 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__17___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -110346,7 +102077,7 @@ static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Un 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_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__16___lam__2___closed__1; x_2 = lean_unsigned_to_nat(73u); -x_3 = lean_unsigned_to_nat(1398u); +x_3 = lean_unsigned_to_nat(1279u); x_4 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__16___lam__2___closed__0; x_5 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__17___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -110483,7 +102214,7 @@ x_59 = l_Lean_MessageData_ofName(x_13); lean_ctor_set_tag(x_35, 7); lean_ctor_set(x_35, 1, x_59); lean_ctor_set(x_35, 0, x_58); -x_60 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___closed__3; +x_60 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___closed__3; lean_ctor_set_tag(x_24, 7); lean_ctor_set(x_24, 1, x_60); lean_ctor_set(x_24, 0, x_35); @@ -110571,7 +102302,7 @@ x_74 = l_Lean_MessageData_ofName(x_13); lean_ctor_set_tag(x_35, 7); lean_ctor_set(x_35, 1, x_74); lean_ctor_set(x_35, 0, x_73); -x_75 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___closed__3; +x_75 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___closed__3; lean_ctor_set_tag(x_24, 7); lean_ctor_set(x_24, 1, x_75); lean_ctor_set(x_24, 0, x_35); @@ -110832,7 +102563,7 @@ x_111 = l_Lean_MessageData_ofName(x_13); x_112 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_112, 0, x_110); lean_ctor_set(x_112, 1, x_111); -x_113 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___closed__3; +x_113 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___closed__3; lean_ctor_set_tag(x_24, 7); lean_ctor_set(x_24, 1, x_113); lean_ctor_set(x_24, 0, x_112); @@ -111136,7 +102867,7 @@ if (lean_is_scalar(x_140)) { } lean_ctor_set(x_163, 0, x_161); lean_ctor_set(x_163, 1, x_162); -x_164 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___closed__3; +x_164 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___closed__3; x_165 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_165, 0, x_163); lean_ctor_set(x_165, 1, x_164); @@ -114854,7 +106585,7 @@ lean_inc(x_72); lean_inc_ref(x_71); lean_inc(x_70); lean_inc_ref(x_69); -x_75 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_withLetDecls___redArg(x_74, x_68, x_66, x_65, x_69, x_70, x_71, x_72); +x_75 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_withLetDecls___redArg(x_74, x_68, x_67, x_65, x_69, x_70, x_71, x_72); if (lean_obj_tag(x_75) == 0) { lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; @@ -114878,7 +106609,7 @@ lean_inc(x_72); lean_inc_ref(x_71); lean_inc(x_70); lean_inc_ref(x_69); -x_82 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars(x_67, x_81, x_76, x_69, x_70, x_71, x_72); +x_82 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars(x_66, x_81, x_76, x_69, x_70, x_71, x_72); if (lean_obj_tag(x_82) == 0) { lean_object* x_83; uint8_t x_84; uint8_t x_85; lean_object* x_86; @@ -114909,7 +106640,7 @@ if (x_90 == 0) lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; uint8_t x_97; x_91 = lean_ctor_get(x_89, 0); x_92 = lean_ctor_get(x_89, 1); -x_93 = l_Lean_instantiateMVars___at___00Lean_Tactic_FunInd_rwMatcher_spec__3___redArg(x_92, x_70); +x_93 = l_Lean_instantiateMVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__4___redArg(x_92, x_70); x_94 = lean_ctor_get(x_93, 0); lean_inc(x_94); lean_dec_ref(x_93); @@ -114928,8 +106659,8 @@ lean_dec_ref(x_71); lean_dec(x_70); lean_dec_ref(x_69); lean_dec(x_13); -x_32 = x_94; -x_33 = x_91; +x_32 = x_91; +x_33 = x_94; x_34 = lean_box(0); goto block_38; } @@ -114950,8 +106681,8 @@ lean_dec_ref(x_69); if (lean_obj_tag(x_100) == 0) { lean_dec_ref(x_100); -x_32 = x_94; -x_33 = x_91; +x_32 = x_91; +x_33 = x_94; x_34 = lean_box(0); goto block_38; } @@ -114987,7 +106718,7 @@ x_105 = lean_ctor_get(x_89, 1); lean_inc(x_105); lean_inc(x_104); lean_dec(x_89); -x_106 = l_Lean_instantiateMVars___at___00Lean_Tactic_FunInd_rwMatcher_spec__3___redArg(x_105, x_70); +x_106 = l_Lean_instantiateMVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__4___redArg(x_105, x_70); x_107 = lean_ctor_get(x_106, 0); lean_inc(x_107); lean_dec_ref(x_106); @@ -115005,8 +106736,8 @@ lean_dec_ref(x_71); lean_dec(x_70); lean_dec_ref(x_69); lean_dec(x_13); -x_32 = x_107; -x_33 = x_104; +x_32 = x_104; +x_33 = x_107; x_34 = lean_box(0); goto block_38; } @@ -115027,8 +106758,8 @@ lean_dec_ref(x_69); if (lean_obj_tag(x_114) == 0) { lean_dec_ref(x_114); -x_32 = x_107; -x_33 = x_104; +x_32 = x_104; +x_33 = x_107; x_34 = lean_box(0); goto block_38; } @@ -115147,7 +106878,7 @@ lean_dec(x_72); lean_dec_ref(x_71); lean_dec(x_70); lean_dec_ref(x_69); -lean_dec_ref(x_67); +lean_dec_ref(x_66); lean_dec_ref(x_26); lean_dec_ref(x_14); lean_dec(x_13); @@ -115176,7 +106907,7 @@ lean_dec(x_72); lean_dec_ref(x_71); lean_dec(x_70); lean_dec_ref(x_69); -lean_dec_ref(x_67); +lean_dec_ref(x_66); lean_dec_ref(x_26); lean_dec_ref(x_14); lean_dec(x_13); @@ -115274,8 +107005,8 @@ lean_dec(x_157); if (x_158 == 0) { lean_free_object(x_152); -x_66 = x_154; -x_67 = x_155; +x_66 = x_155; +x_67 = x_154; x_68 = x_134; x_69 = x_135; x_70 = x_136; @@ -115301,8 +107032,8 @@ x_164 = l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tacti if (lean_obj_tag(x_164) == 0) { lean_dec_ref(x_164); -x_66 = x_154; -x_67 = x_155; +x_66 = x_155; +x_67 = x_154; x_68 = x_134; x_69 = x_135; x_70 = x_136; @@ -115362,8 +107093,8 @@ x_172 = lean_unbox(x_171); lean_dec(x_171); if (x_172 == 0) { -x_66 = x_168; -x_67 = x_169; +x_66 = x_169; +x_67 = x_168; x_68 = x_134; x_69 = x_135; x_70 = x_136; @@ -115389,8 +107120,8 @@ x_179 = l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tacti if (lean_obj_tag(x_179) == 0) { lean_dec_ref(x_179); -x_66 = x_168; -x_67 = x_169; +x_66 = x_169; +x_67 = x_168; x_68 = x_134; x_69 = x_135; x_70 = x_136; @@ -115717,10 +107448,10 @@ return x_229; { lean_object* x_35; lean_object* x_36; lean_object* x_37; x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 0, x_32); lean_ctor_set(x_35, 1, x_14); x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_32); +lean_ctor_set(x_36, 0, x_33); lean_ctor_set(x_36, 1, x_35); x_37 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_37, 0, x_36); @@ -115955,7 +107686,7 @@ x_50 = l_Array_zip___redArg(x_41, x_32); lean_dec(x_32); lean_dec_ref(x_41); x_51 = 0; -x_52 = l_Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__5___redArg(x_49, x_50, x_46, x_51, x_33, x_34, x_35, x_36); +x_52 = l_Lean_Meta_withLocalDeclsDND___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__6___redArg(x_49, x_50, x_46, x_51, x_33, x_34, x_35, x_36); return x_52; } else @@ -117020,7 +108751,7 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_Fun lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__17___closed__2; x_2 = lean_unsigned_to_nat(41u); -x_3 = lean_unsigned_to_nat(1253u); +x_3 = lean_unsigned_to_nat(1134u); x_4 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__16___lam__2___closed__0; x_5 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__17___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -117184,15 +108915,15 @@ goto block_353; lean_object* x_52; lean_object* x_53; uint8_t x_54; x_52 = lean_array_get_size(x_46); lean_dec_ref(x_46); -x_53 = lean_array_get_size(x_41); +x_53 = lean_array_get_size(x_43); x_54 = lean_nat_dec_eq(x_52, x_53); if (x_54 == 0) { lean_object* x_55; lean_object* x_56; lean_dec_ref(x_45); -lean_dec(x_44); +lean_dec_ref(x_44); lean_dec_ref(x_43); -lean_dec_ref(x_42); +lean_dec(x_42); lean_dec_ref(x_41); lean_dec_ref(x_40); lean_dec(x_39); @@ -117228,27 +108959,27 @@ x_63 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Ta lean_closure_set(x_63, 0, x_4); lean_closure_set(x_63, 1, x_5); lean_closure_set(x_63, 2, x_6); -lean_closure_set(x_63, 3, x_44); +lean_closure_set(x_63, 3, x_35); lean_closure_set(x_63, 4, x_60); lean_closure_set(x_63, 5, x_7); lean_closure_set(x_63, 6, x_1); lean_closure_set(x_63, 7, x_2); lean_closure_set(x_63, 8, x_59); -lean_closure_set(x_63, 9, x_35); +lean_closure_set(x_63, 9, x_42); lean_closure_set(x_63, 10, x_61); lean_closure_set(x_63, 11, x_33); -lean_closure_set(x_63, 12, x_36); -lean_closure_set(x_63, 13, x_41); +lean_closure_set(x_63, 12, x_41); +lean_closure_set(x_63, 13, x_43); lean_closure_set(x_63, 14, x_38); -lean_closure_set(x_63, 15, x_43); -lean_closure_set(x_63, 16, x_40); +lean_closure_set(x_63, 15, x_40); +lean_closure_set(x_63, 16, x_37); lean_closure_set(x_63, 17, x_62); -lean_closure_set(x_63, 18, x_42); -lean_closure_set(x_63, 19, x_37); +lean_closure_set(x_63, 18, x_45); +lean_closure_set(x_63, 19, x_34); lean_closure_set(x_63, 20, x_10); lean_closure_set(x_63, 21, x_11); -lean_closure_set(x_63, 22, x_34); -lean_closure_set(x_63, 23, x_45); +lean_closure_set(x_63, 22, x_36); +lean_closure_set(x_63, 23, x_44); lean_closure_set(x_63, 24, x_39); x_64 = l_Lean_Meta_withErasedFVars___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__18___redArg(x_58, x_63, x_47, x_48, x_49, x_50); lean_dec_ref(x_58); @@ -117282,7 +109013,7 @@ lean_inc(x_79); lean_inc(x_85); lean_ctor_set(x_26, 0, x_85); x_86 = l_Lean_Expr_const___override(x_82, x_26); -x_87 = l_Lean_mkAppN(x_86, x_72); +x_87 = l_Lean_mkAppN(x_86, x_70); lean_inc(x_76); lean_inc_ref(x_75); lean_inc(x_74); @@ -117307,9 +109038,9 @@ lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; x_92 = lean_ctor_get(x_91, 0); lean_inc(x_92); lean_dec_ref(x_91); -lean_inc_ref(x_70); -x_93 = l_Lean_Elab_Structural_IndGroupInfo_ofInductiveVal(x_70); -lean_ctor_set(x_71, 2, x_72); +lean_inc_ref(x_72); +x_93 = l_Lean_Elab_Structural_IndGroupInfo_ofInductiveVal(x_72); +lean_ctor_set(x_71, 2, x_70); lean_ctor_set(x_71, 1, x_79); lean_ctor_set(x_71, 0, x_93); x_94 = lean_mk_empty_array_with_capacity(x_1); @@ -117350,8 +109081,8 @@ lean_inc(x_102); lean_dec_ref(x_101); x_103 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__6___closed__2; x_104 = lean_alloc_closure((void*)(l_instDecidableEqNat___boxed), 2, 0); -x_105 = l_Lean_InductiveVal_numTypeFormers(x_70); -lean_dec_ref(x_70); +x_105 = l_Lean_InductiveVal_numTypeFormers(x_72); +lean_dec_ref(x_72); x_106 = l_Array_range(x_105); x_107 = l_Lean_Elab_Structural_Positions_groupAndSort___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__6___redArg(x_103, x_8, x_96, x_106); x_108 = lean_unbox(x_102); @@ -117360,17 +109091,17 @@ if (x_108 == 0) { lean_inc(x_99); x_34 = x_66; -x_35 = x_100; -x_36 = x_96; -x_37 = x_67; +x_35 = x_85; +x_36 = x_67; +x_37 = x_71; x_38 = x_99; x_39 = x_68; -x_40 = x_71; -x_41 = x_107; -x_42 = x_92; -x_43 = x_104; -x_44 = x_85; -x_45 = x_97; +x_40 = x_104; +x_41 = x_96; +x_42 = x_100; +x_43 = x_107; +x_44 = x_97; +x_45 = x_92; x_46 = x_99; x_47 = x_73; x_48 = x_74; @@ -117397,17 +109128,17 @@ if (lean_obj_tag(x_115) == 0) lean_dec_ref(x_115); lean_inc(x_99); x_34 = x_66; -x_35 = x_100; -x_36 = x_96; -x_37 = x_67; +x_35 = x_85; +x_36 = x_67; +x_37 = x_71; x_38 = x_99; x_39 = x_68; -x_40 = x_71; -x_41 = x_107; -x_42 = x_92; -x_43 = x_104; -x_44 = x_85; -x_45 = x_97; +x_40 = x_104; +x_41 = x_96; +x_42 = x_100; +x_43 = x_107; +x_44 = x_97; +x_45 = x_92; x_46 = x_99; x_47 = x_73; x_48 = x_74; @@ -117474,7 +109205,7 @@ lean_dec(x_76); lean_dec_ref(x_75); lean_dec(x_74); lean_dec_ref(x_73); -lean_dec_ref(x_70); +lean_dec_ref(x_72); lean_dec(x_68); lean_dec_ref(x_67); lean_dec_ref(x_66); @@ -117516,7 +109247,7 @@ lean_dec(x_76); lean_dec_ref(x_75); lean_dec(x_74); lean_dec_ref(x_73); -lean_dec_ref(x_70); +lean_dec_ref(x_72); lean_dec(x_69); lean_dec(x_68); lean_dec_ref(x_67); @@ -117648,7 +109379,7 @@ lean_inc(x_79); lean_inc(x_132); lean_ctor_set(x_26, 0, x_132); x_133 = l_Lean_Expr_const___override(x_131, x_26); -x_134 = l_Lean_mkAppN(x_133, x_72); +x_134 = l_Lean_mkAppN(x_133, x_70); lean_inc(x_76); lean_inc_ref(x_75); lean_inc(x_74); @@ -117673,12 +109404,12 @@ lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; x_139 = lean_ctor_get(x_138, 0); lean_inc(x_139); lean_dec_ref(x_138); -lean_inc_ref(x_70); -x_140 = l_Lean_Elab_Structural_IndGroupInfo_ofInductiveVal(x_70); +lean_inc_ref(x_72); +x_140 = l_Lean_Elab_Structural_IndGroupInfo_ofInductiveVal(x_72); x_141 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_141, 0, x_140); lean_ctor_set(x_141, 1, x_79); -lean_ctor_set(x_141, 2, x_72); +lean_ctor_set(x_141, 2, x_70); x_142 = lean_mk_empty_array_with_capacity(x_1); lean_inc(x_76); lean_inc_ref(x_75); @@ -117717,8 +109448,8 @@ lean_inc(x_150); lean_dec_ref(x_149); x_151 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__6___closed__2; x_152 = lean_alloc_closure((void*)(l_instDecidableEqNat___boxed), 2, 0); -x_153 = l_Lean_InductiveVal_numTypeFormers(x_70); -lean_dec_ref(x_70); +x_153 = l_Lean_InductiveVal_numTypeFormers(x_72); +lean_dec_ref(x_72); x_154 = l_Array_range(x_153); x_155 = l_Lean_Elab_Structural_Positions_groupAndSort___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__6___redArg(x_151, x_8, x_144, x_154); x_156 = lean_unbox(x_150); @@ -117727,17 +109458,17 @@ if (x_156 == 0) { lean_inc(x_147); x_34 = x_66; -x_35 = x_148; -x_36 = x_144; -x_37 = x_67; +x_35 = x_132; +x_36 = x_67; +x_37 = x_141; x_38 = x_147; x_39 = x_68; -x_40 = x_141; -x_41 = x_155; -x_42 = x_139; -x_43 = x_152; -x_44 = x_132; -x_45 = x_145; +x_40 = x_152; +x_41 = x_144; +x_42 = x_148; +x_43 = x_155; +x_44 = x_145; +x_45 = x_139; x_46 = x_147; x_47 = x_73; x_48 = x_74; @@ -117764,17 +109495,17 @@ if (lean_obj_tag(x_163) == 0) lean_dec_ref(x_163); lean_inc(x_147); x_34 = x_66; -x_35 = x_148; -x_36 = x_144; -x_37 = x_67; +x_35 = x_132; +x_36 = x_67; +x_37 = x_141; x_38 = x_147; x_39 = x_68; -x_40 = x_141; -x_41 = x_155; -x_42 = x_139; -x_43 = x_152; -x_44 = x_132; -x_45 = x_145; +x_40 = x_152; +x_41 = x_144; +x_42 = x_148; +x_43 = x_155; +x_44 = x_145; +x_45 = x_139; x_46 = x_147; x_47 = x_73; x_48 = x_74; @@ -117842,7 +109573,7 @@ lean_dec(x_76); lean_dec_ref(x_75); lean_dec(x_74); lean_dec_ref(x_73); -lean_dec_ref(x_70); +lean_dec_ref(x_72); lean_dec(x_68); lean_dec_ref(x_67); lean_dec_ref(x_66); @@ -117885,7 +109616,7 @@ lean_dec(x_76); lean_dec_ref(x_75); lean_dec(x_74); lean_dec_ref(x_73); -lean_dec_ref(x_70); +lean_dec_ref(x_72); lean_dec(x_69); lean_dec(x_68); lean_dec_ref(x_67); @@ -118032,7 +109763,7 @@ x_183 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_183, 0, x_182); lean_ctor_set(x_183, 1, x_179); x_184 = l_Lean_Expr_const___override(x_180, x_183); -x_185 = l_Lean_mkAppN(x_184, x_72); +x_185 = l_Lean_mkAppN(x_184, x_70); lean_inc(x_76); lean_inc_ref(x_75); lean_inc(x_74); @@ -118057,8 +109788,8 @@ lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; x_190 = lean_ctor_get(x_189, 0); lean_inc(x_190); lean_dec_ref(x_189); -lean_inc_ref(x_70); -x_191 = l_Lean_Elab_Structural_IndGroupInfo_ofInductiveVal(x_70); +lean_inc_ref(x_72); +x_191 = l_Lean_Elab_Structural_IndGroupInfo_ofInductiveVal(x_72); if (lean_is_scalar(x_181)) { x_192 = lean_alloc_ctor(0, 3, 0); } else { @@ -118066,7 +109797,7 @@ if (lean_is_scalar(x_181)) { } lean_ctor_set(x_192, 0, x_191); lean_ctor_set(x_192, 1, x_179); -lean_ctor_set(x_192, 2, x_72); +lean_ctor_set(x_192, 2, x_70); x_193 = lean_mk_empty_array_with_capacity(x_1); lean_inc(x_76); lean_inc_ref(x_75); @@ -118105,8 +109836,8 @@ lean_inc(x_201); lean_dec_ref(x_200); x_202 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__6___closed__2; x_203 = lean_alloc_closure((void*)(l_instDecidableEqNat___boxed), 2, 0); -x_204 = l_Lean_InductiveVal_numTypeFormers(x_70); -lean_dec_ref(x_70); +x_204 = l_Lean_InductiveVal_numTypeFormers(x_72); +lean_dec_ref(x_72); x_205 = l_Array_range(x_204); x_206 = l_Lean_Elab_Structural_Positions_groupAndSort___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__6___redArg(x_202, x_8, x_195, x_205); x_207 = lean_unbox(x_201); @@ -118115,17 +109846,17 @@ if (x_207 == 0) { lean_inc(x_198); x_34 = x_66; -x_35 = x_199; -x_36 = x_195; -x_37 = x_67; +x_35 = x_182; +x_36 = x_67; +x_37 = x_192; x_38 = x_198; x_39 = x_68; -x_40 = x_192; -x_41 = x_206; -x_42 = x_190; -x_43 = x_203; -x_44 = x_182; -x_45 = x_196; +x_40 = x_203; +x_41 = x_195; +x_42 = x_199; +x_43 = x_206; +x_44 = x_196; +x_45 = x_190; x_46 = x_198; x_47 = x_73; x_48 = x_74; @@ -118152,17 +109883,17 @@ if (lean_obj_tag(x_214) == 0) lean_dec_ref(x_214); lean_inc(x_198); x_34 = x_66; -x_35 = x_199; -x_36 = x_195; -x_37 = x_67; +x_35 = x_182; +x_36 = x_67; +x_37 = x_192; x_38 = x_198; x_39 = x_68; -x_40 = x_192; -x_41 = x_206; -x_42 = x_190; -x_43 = x_203; -x_44 = x_182; -x_45 = x_196; +x_40 = x_203; +x_41 = x_195; +x_42 = x_199; +x_43 = x_206; +x_44 = x_196; +x_45 = x_190; x_46 = x_198; x_47 = x_73; x_48 = x_74; @@ -118230,7 +109961,7 @@ lean_dec(x_76); lean_dec_ref(x_75); lean_dec(x_74); lean_dec_ref(x_73); -lean_dec_ref(x_70); +lean_dec_ref(x_72); lean_dec(x_68); lean_dec_ref(x_67); lean_dec_ref(x_66); @@ -118273,7 +110004,7 @@ lean_dec(x_76); lean_dec_ref(x_75); lean_dec(x_74); lean_dec_ref(x_73); -lean_dec_ref(x_70); +lean_dec_ref(x_72); lean_dec(x_69); lean_dec(x_68); lean_dec_ref(x_67); @@ -118436,17 +110167,17 @@ return x_233; block_259: { lean_object* x_248; uint8_t x_249; -x_248 = lean_array_get_size(x_241); +x_248 = lean_array_get_size(x_240); x_249 = lean_nat_dec_lt(x_2, x_248); if (x_249 == 0) { -lean_dec_ref(x_241); +lean_dec_ref(x_240); x_66 = x_235; x_67 = x_236; x_68 = x_237; x_69 = x_238; x_70 = x_239; -x_71 = x_240; +x_71 = x_241; x_72 = x_242; x_73 = x_243; x_74 = x_244; @@ -118459,13 +110190,13 @@ else { if (x_249 == 0) { -lean_dec_ref(x_241); +lean_dec_ref(x_240); x_66 = x_235; x_67 = x_236; x_68 = x_237; x_69 = x_238; x_70 = x_239; -x_71 = x_240; +x_71 = x_241; x_72 = x_242; x_73 = x_243; x_74 = x_244; @@ -118478,8 +110209,8 @@ else { size_t x_250; uint8_t x_251; x_250 = lean_usize_of_nat(x_248); -x_251 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__19(x_241, x_3, x_250); -lean_dec_ref(x_241); +x_251 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__19(x_240, x_3, x_250); +lean_dec_ref(x_240); if (x_251 == 0) { x_66 = x_235; @@ -118487,7 +110218,7 @@ x_67 = x_236; x_68 = x_237; x_69 = x_238; x_70 = x_239; -x_71 = x_240; +x_71 = x_241; x_72 = x_242; x_73 = x_243; x_74 = x_244; @@ -118500,7 +110231,7 @@ else { lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; uint8_t x_256; lean_dec_ref(x_242); -lean_dec_ref(x_240); +lean_dec_ref(x_241); lean_dec_ref(x_239); lean_dec(x_238); lean_dec(x_237); @@ -118552,23 +110283,23 @@ return x_258; lean_object* x_280; lean_object* x_281; lean_object* x_282; uint8_t x_283; x_280 = l_Array_toSubarray___redArg(x_266, x_278, x_279); x_281 = l_Subarray_toArray___redArg(x_280); -x_282 = lean_array_get_size(x_267); +x_282 = lean_array_get_size(x_269); x_283 = lean_nat_dec_lt(x_2, x_282); if (x_283 == 0) { -lean_inc(x_269); -x_235 = x_267; -x_236 = x_268; -x_237 = x_269; -x_238 = x_269; -x_239 = x_272; -x_240 = x_271; -x_241 = x_281; -x_242 = x_270; -x_243 = x_274; -x_244 = x_277; -x_245 = x_276; -x_246 = x_273; +lean_inc(x_271); +x_235 = x_268; +x_236 = x_269; +x_237 = x_271; +x_238 = x_271; +x_239 = x_270; +x_240 = x_281; +x_241 = x_267; +x_242 = x_277; +x_243 = x_273; +x_244 = x_276; +x_245 = x_274; +x_246 = x_272; x_247 = lean_box(0); goto block_259; } @@ -118576,19 +110307,19 @@ else { if (x_283 == 0) { -lean_inc(x_269); -x_235 = x_267; -x_236 = x_268; -x_237 = x_269; -x_238 = x_269; -x_239 = x_272; -x_240 = x_271; -x_241 = x_281; -x_242 = x_270; -x_243 = x_274; -x_244 = x_277; -x_245 = x_276; -x_246 = x_273; +lean_inc(x_271); +x_235 = x_268; +x_236 = x_269; +x_237 = x_271; +x_238 = x_271; +x_239 = x_270; +x_240 = x_281; +x_241 = x_267; +x_242 = x_277; +x_243 = x_273; +x_244 = x_276; +x_245 = x_274; +x_246 = x_272; x_247 = lean_box(0); goto block_259; } @@ -118596,22 +110327,22 @@ else { size_t x_284; uint8_t x_285; x_284 = lean_usize_of_nat(x_282); -x_285 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__19(x_267, x_3, x_284); +x_285 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__19(x_269, x_3, x_284); if (x_285 == 0) { -lean_inc(x_269); -x_235 = x_267; -x_236 = x_268; -x_237 = x_269; -x_238 = x_269; -x_239 = x_272; -x_240 = x_271; -x_241 = x_281; -x_242 = x_270; -x_243 = x_274; -x_244 = x_277; -x_245 = x_276; -x_246 = x_273; +lean_inc(x_271); +x_235 = x_268; +x_236 = x_269; +x_237 = x_271; +x_238 = x_271; +x_239 = x_270; +x_240 = x_281; +x_241 = x_267; +x_242 = x_277; +x_243 = x_273; +x_244 = x_276; +x_245 = x_274; +x_246 = x_272; x_247 = lean_box(0); goto block_259; } @@ -118619,10 +110350,10 @@ else { lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; uint8_t x_290; lean_dec_ref(x_281); -lean_dec_ref(x_272); -lean_dec_ref(x_271); +lean_dec_ref(x_277); +lean_dec(x_271); lean_dec_ref(x_270); -lean_dec(x_269); +lean_dec_ref(x_269); lean_dec_ref(x_268); lean_dec_ref(x_267); lean_dec_ref(x_33); @@ -118642,11 +110373,11 @@ x_287 = l_Lean_indentExpr(x_24); x_288 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_288, 0, x_286); lean_ctor_set(x_288, 1, x_287); -x_289 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___redArg(x_288, x_274, x_277, x_276, x_273); -lean_dec(x_273); -lean_dec_ref(x_276); -lean_dec(x_277); +x_289 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect_spec__8___redArg(x_288, x_273, x_276, x_274, x_272); +lean_dec(x_272); lean_dec_ref(x_274); +lean_dec(x_276); +lean_dec_ref(x_273); x_290 = !lean_is_exclusive(x_289); if (x_290 == 0) { @@ -118669,22 +110400,22 @@ return x_292; block_316: { lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; uint8_t x_315; -lean_inc(x_296); +lean_inc(x_295); lean_inc(x_2); lean_inc_ref(x_266); -x_304 = l_Array_toSubarray___redArg(x_266, x_2, x_296); +x_304 = l_Array_toSubarray___redArg(x_266, x_2, x_295); x_305 = l_Subarray_toArray___redArg(x_304); -x_306 = lean_nat_add(x_296, x_298); +x_306 = lean_nat_add(x_295, x_297); +lean_dec(x_295); +x_307 = lean_nat_add(x_306, x_296); lean_dec(x_296); -x_307 = lean_nat_add(x_306, x_297); -lean_dec(x_297); x_308 = lean_nat_add(x_307, x_32); lean_dec(x_307); lean_inc(x_308); lean_inc_ref(x_266); x_309 = l_Array_toSubarray___redArg(x_266, x_306, x_308); x_310 = l_Subarray_toArray___redArg(x_309); -x_311 = lean_nat_add(x_308, x_298); +x_311 = lean_nat_add(x_308, x_297); lean_inc(x_311); lean_inc_ref(x_266); x_312 = l_Array_toSubarray___redArg(x_266, x_308, x_311); @@ -118693,17 +110424,17 @@ x_314 = lean_array_get_size(x_266); x_315 = lean_nat_dec_le(x_311, x_2); if (x_315 == 0) { -x_267 = x_310; +x_267 = x_294; x_268 = x_313; -x_269 = x_298; +x_269 = x_310; x_270 = x_305; -x_271 = x_295; -x_272 = x_294; -x_273 = x_302; -x_274 = x_299; +x_271 = x_297; +x_272 = x_302; +x_273 = x_299; +x_274 = x_301; x_275 = lean_box(0); -x_276 = x_301; -x_277 = x_300; +x_276 = x_300; +x_277 = x_298; x_278 = x_311; x_279 = x_314; goto block_293; @@ -118712,17 +110443,17 @@ else { lean_dec(x_311); lean_inc(x_2); -x_267 = x_310; +x_267 = x_294; x_268 = x_313; -x_269 = x_298; +x_269 = x_310; x_270 = x_305; -x_271 = x_295; -x_272 = x_294; -x_273 = x_302; -x_274 = x_299; +x_271 = x_297; +x_272 = x_302; +x_273 = x_299; +x_274 = x_301; x_275 = lean_box(0); -x_276 = x_301; -x_277 = x_300; +x_276 = x_300; +x_277 = x_298; x_278 = x_2; x_279 = x_314; goto block_293; @@ -118777,11 +110508,11 @@ x_337 = lean_nat_dec_lt(x_332, x_336); lean_dec(x_336); if (x_337 == 0) { -x_294 = x_324; -x_295 = x_328; -x_296 = x_329; -x_297 = x_330; -x_298 = x_331; +x_294 = x_328; +x_295 = x_329; +x_296 = x_330; +x_297 = x_331; +x_298 = x_324; x_299 = x_317; x_300 = x_318; x_301 = x_319; @@ -119167,7 +110898,7 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_Fun lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___closed__0; x_2 = lean_unsigned_to_nat(2u); -x_3 = lean_unsigned_to_nat(1231u); +x_3 = lean_unsigned_to_nat(1112u); x_4 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__16___lam__2___closed__0; x_5 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__17___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -121622,7 +113353,7 @@ goto block_64; block_39: { lean_object* x_18; uint8_t x_19; -x_18 = l_Lean_Tactic_FunInd_rwIfWith___closed__1; +x_18 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__20; x_19 = l_Lean_Expr_isConstOf(x_3, x_18); lean_dec_ref(x_3); if (x_19 == 0) @@ -121717,7 +113448,7 @@ x_47 = lean_unsigned_to_nat(5u); lean_inc_ref(x_4); x_48 = l_Array_toSubarray___redArg(x_4, x_46, x_47); x_49 = lean_box(0); -x_50 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_refinedArguments_go_spec__0___redArg(x_1, x_49, x_48, x_49, x_41, x_44, x_40, x_43, x_45); +x_50 = l_WellFounded_opaqueFix_u2083___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_refinedArguments_go_spec__0___redArg(x_1, x_49, x_48, x_49, x_44, x_40, x_42, x_45, x_43); if (lean_obj_tag(x_50) == 0) { lean_object* x_51; lean_object* x_52; @@ -121728,10 +113459,10 @@ x_52 = lean_ctor_get(x_51, 1); lean_inc(x_52); lean_dec(x_51); x_12 = x_52; -x_13 = x_44; -x_14 = x_40; -x_15 = x_43; -x_16 = x_45; +x_13 = x_40; +x_14 = x_42; +x_15 = x_45; +x_16 = x_43; x_17 = lean_box(0); goto block_39; } @@ -121745,12 +113476,12 @@ return x_50; block_64: { lean_object* x_60; uint8_t x_61; -x_60 = l_Lean_Tactic_FunInd_rwIfWith___closed__5; +x_60 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__24; x_61 = l_Lean_Expr_isConstOf(x_3, x_60); if (x_61 == 0) { lean_object* x_62; uint8_t x_63; -x_62 = l_Lean_Tactic_FunInd_rwIfWith___closed__3; +x_62 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__22; x_63 = l_Lean_Expr_isConstOf(x_3, x_62); if (x_63 == 0) { @@ -121764,23 +113495,23 @@ goto block_39; } else { -x_40 = x_56; -x_41 = x_54; -x_42 = lean_box(0); -x_43 = x_57; -x_44 = x_55; -x_45 = x_58; +x_40 = x_55; +x_41 = lean_box(0); +x_42 = x_56; +x_43 = x_58; +x_44 = x_54; +x_45 = x_57; goto block_53; } } else { -x_40 = x_56; -x_41 = x_54; -x_42 = lean_box(0); -x_43 = x_57; -x_44 = x_55; -x_45 = x_58; +x_40 = x_55; +x_41 = lean_box(0); +x_42 = x_56; +x_43 = x_58; +x_44 = x_54; +x_45 = x_57; goto block_53; } } @@ -122225,7 +113956,7 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_Fun lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_refinedArguments___closed__1; x_2 = lean_unsigned_to_nat(2u); -x_3 = lean_unsigned_to_nat(1471u); +x_3 = lean_unsigned_to_nat(1352u); x_4 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_refinedArguments___closed__0; x_5 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__17___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -123185,7 +114916,7 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_Fun _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("` not an equality:", 18, 18); +x_1 = lean_mk_string_unchecked("Type of `", 9, 9); return x_1; } } @@ -123198,6 +114929,23 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__1___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("` not an equality:", 18, 18); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__1___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__1(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { @@ -123208,12 +114956,12 @@ x_12 = l_Lean_Expr_isAppOfArity(x_4, x_10, x_11); if (x_12 == 0) { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_13 = l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__19; +x_13 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__1___closed__1; x_14 = l_Lean_MessageData_ofName(x_1); x_15 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_15, 0, x_13); lean_ctor_set(x_15, 1, x_14); -x_16 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__1___closed__1; +x_16 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__1___closed__3; x_17 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_17, 0, x_15); lean_ctor_set(x_17, 1, x_16); @@ -123614,7 +115362,7 @@ static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Un 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_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases_spec__3___closed__1; x_2 = lean_unsigned_to_nat(8u); -x_3 = lean_unsigned_to_nat(1589u); +x_3 = lean_unsigned_to_nat(1470u); x_4 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases_spec__3___closed__0; x_5 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_foldAndCollect___lam__17___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -124006,14 +115754,14 @@ lean_dec(x_4); if (x_13 == 0) { lean_object* x_14; -x_14 = l_Lean_Tactic_FunInd_rwIfWith___closed__19; +x_14 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__31; x_7 = x_14; goto block_12; } else { lean_object* x_15; -x_15 = l_Lean_Tactic_FunInd_rwIfWith___closed__16; +x_15 = l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__28; x_7 = x_15; goto block_12; } @@ -126827,6 +118575,7 @@ return x_2; } lean_object* initialize_Lean_Meta_Tactic_Simp_Types(uint8_t builtin); lean_object* initialize_Lean_Meta_Match_MatcherApp_Transform(uint8_t builtin); +lean_object* initialize_Lean_Meta_Match_Rewrite(uint8_t builtin); lean_object* initialize_Lean_Meta_Injective(uint8_t builtin); lean_object* initialize_Lean_Meta_ArgsPacker(uint8_t builtin); lean_object* initialize_Lean_Elab_PreDefinition_WF_Eqns(uint8_t builtin); @@ -126836,7 +118585,6 @@ lean_object* initialize_Lean_Meta_Tactic_ElimInfo(uint8_t builtin); lean_object* initialize_Lean_Meta_Tactic_FunIndInfo(uint8_t builtin); lean_object* initialize_Lean_Data_Array(uint8_t builtin); lean_object* initialize_Lean_Meta_Tactic_Simp_Rewrite(uint8_t builtin); -lean_object* initialize_Lean_Meta_Tactic_Refl(uint8_t builtin); lean_object* initialize_Lean_Meta_Tactic_Replace(uint8_t builtin); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_FunInd(uint8_t builtin) { @@ -126849,6 +118597,9 @@ lean_dec_ref(res); res = initialize_Lean_Meta_Match_MatcherApp_Transform(builtin); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Meta_Match_Rewrite(builtin); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Lean_Meta_Injective(builtin); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); @@ -126876,9 +118627,6 @@ lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Simp_Rewrite(builtin); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = initialize_Lean_Meta_Tactic_Refl(builtin); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Replace(builtin); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); @@ -127220,60 +118968,6 @@ l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_inProdLambdaLastArg___ lean_mark_persistent(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_inProdLambdaLastArg___closed__0); l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_inProdLambdaLastArg___closed__1 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_inProdLambdaLastArg___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_inProdLambdaLastArg___closed__1); -l_Lean_Tactic_FunInd_rwIfWith___closed__0 = _init_l_Lean_Tactic_FunInd_rwIfWith___closed__0(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwIfWith___closed__0); -l_Lean_Tactic_FunInd_rwIfWith___closed__1 = _init_l_Lean_Tactic_FunInd_rwIfWith___closed__1(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwIfWith___closed__1); -l_Lean_Tactic_FunInd_rwIfWith___closed__2 = _init_l_Lean_Tactic_FunInd_rwIfWith___closed__2(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwIfWith___closed__2); -l_Lean_Tactic_FunInd_rwIfWith___closed__3 = _init_l_Lean_Tactic_FunInd_rwIfWith___closed__3(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwIfWith___closed__3); -l_Lean_Tactic_FunInd_rwIfWith___closed__4 = _init_l_Lean_Tactic_FunInd_rwIfWith___closed__4(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwIfWith___closed__4); -l_Lean_Tactic_FunInd_rwIfWith___closed__5 = _init_l_Lean_Tactic_FunInd_rwIfWith___closed__5(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwIfWith___closed__5); -l_Lean_Tactic_FunInd_rwIfWith___closed__6 = _init_l_Lean_Tactic_FunInd_rwIfWith___closed__6(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwIfWith___closed__6); -l_Lean_Tactic_FunInd_rwIfWith___closed__7 = _init_l_Lean_Tactic_FunInd_rwIfWith___closed__7(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwIfWith___closed__7); -l_Lean_Tactic_FunInd_rwIfWith___closed__8 = _init_l_Lean_Tactic_FunInd_rwIfWith___closed__8(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwIfWith___closed__8); -l_Lean_Tactic_FunInd_rwIfWith___closed__9 = _init_l_Lean_Tactic_FunInd_rwIfWith___closed__9(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwIfWith___closed__9); -l_Lean_Tactic_FunInd_rwIfWith___closed__10 = _init_l_Lean_Tactic_FunInd_rwIfWith___closed__10(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwIfWith___closed__10); -l_Lean_Tactic_FunInd_rwIfWith___closed__11 = _init_l_Lean_Tactic_FunInd_rwIfWith___closed__11(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwIfWith___closed__11); -l_Lean_Tactic_FunInd_rwIfWith___closed__12 = _init_l_Lean_Tactic_FunInd_rwIfWith___closed__12(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwIfWith___closed__12); -l_Lean_Tactic_FunInd_rwIfWith___closed__13 = _init_l_Lean_Tactic_FunInd_rwIfWith___closed__13(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwIfWith___closed__13); -l_Lean_Tactic_FunInd_rwIfWith___closed__14 = _init_l_Lean_Tactic_FunInd_rwIfWith___closed__14(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwIfWith___closed__14); -l_Lean_Tactic_FunInd_rwIfWith___closed__15 = _init_l_Lean_Tactic_FunInd_rwIfWith___closed__15(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwIfWith___closed__15); -l_Lean_Tactic_FunInd_rwIfWith___closed__16 = _init_l_Lean_Tactic_FunInd_rwIfWith___closed__16(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwIfWith___closed__16); -l_Lean_Tactic_FunInd_rwIfWith___closed__17 = _init_l_Lean_Tactic_FunInd_rwIfWith___closed__17(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwIfWith___closed__17); -l_Lean_Tactic_FunInd_rwIfWith___closed__18 = _init_l_Lean_Tactic_FunInd_rwIfWith___closed__18(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwIfWith___closed__18); -l_Lean_Tactic_FunInd_rwIfWith___closed__19 = _init_l_Lean_Tactic_FunInd_rwIfWith___closed__19(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwIfWith___closed__19); -l_Lean_Tactic_FunInd_rwIfWith___closed__20 = _init_l_Lean_Tactic_FunInd_rwIfWith___closed__20(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwIfWith___closed__20); -l_Lean_Tactic_FunInd_rwIfWith___closed__21 = _init_l_Lean_Tactic_FunInd_rwIfWith___closed__21(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwIfWith___closed__21); -l_Lean_Tactic_FunInd_rwIfWith___closed__22 = _init_l_Lean_Tactic_FunInd_rwIfWith___closed__22(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwIfWith___closed__22); -l_Lean_Tactic_FunInd_rwIfWith___closed__23 = _init_l_Lean_Tactic_FunInd_rwIfWith___closed__23(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwIfWith___closed__23); -l_Lean_Tactic_FunInd_rwIfWith___closed__24 = _init_l_Lean_Tactic_FunInd_rwIfWith___closed__24(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwIfWith___closed__24); -l_Lean_Tactic_FunInd_rwIfWith___closed__25 = _init_l_Lean_Tactic_FunInd_rwIfWith___closed__25(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwIfWith___closed__25); -l_Lean_Tactic_FunInd_rwIfWith___closed__26 = _init_l_Lean_Tactic_FunInd_rwIfWith___closed__26(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwIfWith___closed__26); l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___closed__0 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___closed__0(); lean_mark_persistent(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___closed__0); l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___closed__1 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwLetWith___closed__1(); @@ -127294,106 +118988,6 @@ l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic lean_mark_persistent(l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__2); l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__3 = _init_l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__3(); lean_mark_persistent(l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_rwFun_spec__1___closed__3); -l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3___redArg___closed__0 = _init_l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3___redArg___closed__0(); -l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3___redArg___closed__1 = _init_l_Lean_PersistentHashMap_containsAux___at___00Lean_PersistentHashMap_contains___at___00Lean_MVarId_isAssigned___at___00Lean_Tactic_FunInd_rwMatcher_spec__0_spec__0_spec__3___redArg___closed__1(); -l_Lean_Tactic_FunInd_rwMatcher___lam__0___closed__0 = _init_l_Lean_Tactic_FunInd_rwMatcher___lam__0___closed__0(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___lam__0___closed__0); -l_Lean_Tactic_FunInd_rwMatcher___lam__0___closed__1 = _init_l_Lean_Tactic_FunInd_rwMatcher___lam__0___closed__1(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___lam__0___closed__1); -l_Lean_Tactic_FunInd_rwMatcher___lam__0___closed__2 = _init_l_Lean_Tactic_FunInd_rwMatcher___lam__0___closed__2(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___lam__0___closed__2); -l_Lean_Tactic_FunInd_rwMatcher___lam__0___closed__3 = _init_l_Lean_Tactic_FunInd_rwMatcher___lam__0___closed__3(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___lam__0___closed__3); -l_Lean_Tactic_FunInd_rwMatcher___closed__0 = _init_l_Lean_Tactic_FunInd_rwMatcher___closed__0(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___closed__0); -l_Lean_Tactic_FunInd_rwMatcher___closed__1 = _init_l_Lean_Tactic_FunInd_rwMatcher___closed__1(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___closed__1); -l_Lean_Tactic_FunInd_rwMatcher___closed__2 = _init_l_Lean_Tactic_FunInd_rwMatcher___closed__2(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___closed__2); -l_Lean_Tactic_FunInd_rwMatcher___closed__3 = _init_l_Lean_Tactic_FunInd_rwMatcher___closed__3(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___closed__3); -l_Lean_Tactic_FunInd_rwMatcher___closed__4 = _init_l_Lean_Tactic_FunInd_rwMatcher___closed__4(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___closed__4); -l_Lean_Tactic_FunInd_rwMatcher___closed__5 = _init_l_Lean_Tactic_FunInd_rwMatcher___closed__5(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___closed__5); -l_Lean_Tactic_FunInd_rwMatcher___closed__6 = _init_l_Lean_Tactic_FunInd_rwMatcher___closed__6(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___closed__6); -l_Lean_Tactic_FunInd_rwMatcher___closed__7 = _init_l_Lean_Tactic_FunInd_rwMatcher___closed__7(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___closed__7); -l_Lean_Tactic_FunInd_rwMatcher___closed__8 = _init_l_Lean_Tactic_FunInd_rwMatcher___closed__8(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___closed__8); -l_Lean_Tactic_FunInd_rwMatcher___closed__9 = _init_l_Lean_Tactic_FunInd_rwMatcher___closed__9(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___closed__9); -l_Lean_Tactic_FunInd_rwMatcher___closed__10 = _init_l_Lean_Tactic_FunInd_rwMatcher___closed__10(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___closed__10); -l_Lean_Tactic_FunInd_rwMatcher___closed__11 = _init_l_Lean_Tactic_FunInd_rwMatcher___closed__11(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___closed__11); -l_Lean_Tactic_FunInd_rwMatcher___closed__12 = _init_l_Lean_Tactic_FunInd_rwMatcher___closed__12(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___closed__12); -l_Lean_Tactic_FunInd_rwMatcher___closed__13 = _init_l_Lean_Tactic_FunInd_rwMatcher___closed__13(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___closed__13); -l_Lean_Tactic_FunInd_rwMatcher___closed__14 = _init_l_Lean_Tactic_FunInd_rwMatcher___closed__14(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___closed__14); -l_Lean_Tactic_FunInd_rwMatcher___closed__15 = _init_l_Lean_Tactic_FunInd_rwMatcher___closed__15(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___closed__15); -l_Lean_Tactic_FunInd_rwMatcher___closed__16 = _init_l_Lean_Tactic_FunInd_rwMatcher___closed__16(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___closed__16); -l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__0 = _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__0(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__0); -l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__1 = _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__1(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__1); -l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__2 = _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__2(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__2); -l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__3 = _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__3(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__3); -l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__4 = _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__4(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__4); -l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__5 = _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__5(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__5); -l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__6 = _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__6(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__6); -l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__7 = _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__7(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__7); -l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__8 = _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__8(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__8); -l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__9 = _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__9(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__9); -l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5___closed__0 = _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5___closed__0(); -lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5___closed__0); -l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5___closed__1 = _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5___closed__1(); -lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5___closed__1); -l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5___closed__2 = _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5___closed__2(); -lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5___closed__2); -l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5___closed__3 = _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5___closed__3(); -lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Tactic_FunInd_rwMatcher_spec__5___closed__3); -l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__10 = _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__10(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__10); -l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__11 = _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__11(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__11); -l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__12 = _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__12(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__12); -l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__13 = _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__13(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__13); -l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__14 = _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__14(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__14); -l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__15 = _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__15(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__15); -l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__16 = _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__16(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__16); -l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__17 = _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__17(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__17); -l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__18 = _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__18(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__18); -l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__19 = _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__19(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__19); -l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__20 = _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__20(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__20); -l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__21 = _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__21(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__21); -l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__22 = _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__22(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__22); -l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__23 = _init_l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__23(); -lean_mark_persistent(l_Lean_Tactic_FunInd_rwMatcher___lam__1___closed__23); l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__1___closed__0 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__1___closed__0(); lean_mark_persistent(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__1___closed__0); l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__1___closed__1 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__1___closed__1(); @@ -127402,6 +118996,10 @@ l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___l lean_mark_persistent(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__1___closed__2); l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__1___closed__3 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__1___closed__3(); lean_mark_persistent(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__1___closed__3); +l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__1___closed__4 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__1___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__1___closed__4); +l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__1___closed__5 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__1___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__1___closed__5); l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__0 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__0(); lean_mark_persistent(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__0); l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__1 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__1(); @@ -127456,10 +119054,38 @@ l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___l lean_mark_persistent(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__21); l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__22 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__22(); lean_mark_persistent(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__22); -l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6___redArg___closed__0 = _init_l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6___redArg___closed__0(); -lean_mark_persistent(l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__6___redArg___closed__0); -l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__1___closed__0 = _init_l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__1___closed__0(); -lean_mark_persistent(l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__1___closed__0); +l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__23 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__23(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__23); +l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__24 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__24(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__24); +l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__25 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__25(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__25); +l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__26 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__26(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__26); +l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__28 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__28(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__28); +l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__27 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__27(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__27); +l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__29 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__29(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__29); +l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__30 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__30(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__30); +l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__31 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__31(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__31); +l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__32 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__32(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__32); +l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__33 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__33(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__33); +l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__34 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__34(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__34); +l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__35 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__35(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_buildInductionBody___lam__28___closed__35); +l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7___redArg___closed__0 = _init_l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7___redArg___closed__0(); +l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7___redArg___closed__1 = _init_l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7___redArg___closed__1(); +l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7___redArg___closed__2 = _init_l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7___redArg___closed__2(); +lean_mark_persistent(l_Lean_PersistentHashMap_insertAux___at___00Lean_PersistentHashMap_insert___at___00Lean_MVarId_assign___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__1_spec__2_spec__7___redArg___closed__2); +l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2___closed__0 = _init_l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2___closed__0(); +lean_mark_persistent(l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at___00Lean_PersistentArray_foldlM___at___00Lean_LocalContext_foldlM___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars_spec__0_spec__0_spec__2___closed__0); l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars___lam__0___closed__0 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars___lam__0___closed__0(); lean_mark_persistent(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars___lam__0___closed__0); l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars___closed__0 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_abstractIndependentMVars___closed__0(); @@ -127594,10 +119220,10 @@ l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Met lean_mark_persistent(l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__0); l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__1 = _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__1(); lean_mark_persistent(l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__1); -l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__2 = _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__2(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__2); l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__3 = _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__3(); lean_mark_persistent(l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__3); +l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__2 = _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__2(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__2); l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__4 = _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__4(); lean_mark_persistent(l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__4); l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__5 = _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__5(); @@ -127622,6 +119248,16 @@ l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Met lean_mark_persistent(l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__14); l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__15 = _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__15(); lean_mark_persistent(l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__15); +l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__16 = _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__16(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__16); +l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__17 = _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__17(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__17); +l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__18 = _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__18(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__18); +l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__19 = _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__19(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__19); +l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__20 = _init_l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__20(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at___00Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_cleanPackedArgs_spec__0_spec__0___closed__20); l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_unpackMutualInduction_doRealize_spec__0___closed__0 = _init_l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_unpackMutualInduction_doRealize_spec__0___closed__0(); lean_mark_persistent(l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_unpackMutualInduction_doRealize_spec__0___closed__0); l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_unpackMutualInduction_doRealize_spec__0___closed__1 = _init_l_Lean_Expr_withAppAux___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_unpackMutualInduction_doRealize_spec__0___closed__1(); @@ -127700,22 +119336,22 @@ l_Lean_Elab_Structural_Positions_mapMwith___at___00__private_Lean_Meta_Tactic_Fu lean_mark_persistent(l_Lean_Elab_Structural_Positions_mapMwith___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__13___redArg___closed__5); l_Lean_Elab_Structural_Positions_mapMwith___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__13___redArg___closed__6 = _init_l_Lean_Elab_Structural_Positions_mapMwith___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__13___redArg___closed__6(); lean_mark_persistent(l_Lean_Elab_Structural_Positions_mapMwith___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__13___redArg___closed__6); -l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__0 = _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__0(); -lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__0); -l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__1 = _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__1(); -lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__1); -l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__2 = _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__2(); -lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__2); -l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__3 = _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__3(); -lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__2___closed__3); -l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___closed__0 = _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___closed__0(); -lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___closed__0); -l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___closed__1 = _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___closed__1(); -lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___closed__1); -l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___closed__2 = _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___closed__2(); -lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___closed__2); -l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___closed__3 = _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___closed__3(); -lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17_spec__31___lam__0___closed__3); +l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___closed__0 = _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___closed__0(); +lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___closed__0); +l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___closed__1 = _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___closed__1(); +lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___closed__1); +l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___closed__2 = _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___closed__2(); +lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___closed__2); +l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___closed__3 = _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___closed__3(); +lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__2___closed__3); +l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___closed__0 = _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___closed__0(); +lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___closed__0); +l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___closed__1 = _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___closed__1(); +lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___closed__1); +l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___closed__2 = _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___closed__2(); +lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___closed__2); +l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___closed__3 = _init_l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___closed__3(); +lean_mark_persistent(l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize_spec__17___redArg___lam__3___closed__3); l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__1___closed__0 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__1___closed__0(); lean_mark_persistent(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__1___closed__0); l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__1___closed__1 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveInductionStructural_doRealize___lam__1___closed__1(); @@ -127848,6 +119484,10 @@ l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__1__ lean_mark_persistent(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__1___closed__0); l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__1___closed__1 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__1___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__1___closed__1); +l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__1___closed__2 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__1___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__1___closed__2); +l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__1___closed__3 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__1___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__1___closed__3); l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__10___closed__0 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__10___closed__0(); lean_mark_persistent(l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__10___closed__0); l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__10___closed__1 = _init_l___private_Lean_Meta_Tactic_FunInd_0__Lean_Tactic_FunInd_deriveCases___lam__10___closed__1(); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Simp/Diagnostics.c b/stage0/stdlib/Lean/Meta/Tactic/Simp/Diagnostics.c index ad4f12047d79..8c21e98dfbf7 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Simp/Diagnostics.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Simp/Diagnostics.c @@ -4623,7 +4623,7 @@ return x_9; LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logInfo___at___00Lean_Meta_Simp_reportDiag_spec__0_spec__0_spec__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -uint8_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_50; uint8_t x_51; uint8_t x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; lean_object* x_57; lean_object* x_77; uint8_t x_78; uint8_t x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; uint8_t x_83; lean_object* x_84; lean_object* x_88; uint8_t x_89; uint8_t x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; uint8_t x_94; uint8_t x_100; lean_object* x_101; uint8_t x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; uint8_t x_106; uint8_t x_107; uint8_t x_109; uint8_t x_125; +uint8_t x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_50; lean_object* x_51; uint8_t x_52; uint8_t x_53; lean_object* x_54; uint8_t x_55; lean_object* x_56; lean_object* x_57; lean_object* x_77; lean_object* x_78; uint8_t x_79; lean_object* x_80; uint8_t x_81; lean_object* x_82; uint8_t x_83; lean_object* x_84; lean_object* x_88; lean_object* x_89; uint8_t x_90; uint8_t x_91; lean_object* x_92; lean_object* x_93; uint8_t x_94; uint8_t x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; uint8_t x_104; lean_object* x_105; uint8_t x_106; uint8_t x_107; uint8_t x_109; uint8_t x_125; x_100 = 2; x_125 = l_Lean_instBEqMessageSeverity_beq(x_3, x_100); if (x_125 == 0) @@ -4658,15 +4658,15 @@ lean_ctor_set(x_25, 0, x_21); lean_ctor_set(x_25, 1, x_22); x_26 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_11); +lean_ctor_set(x_26, 1, x_14); x_27 = lean_alloc_ctor(0, 5, 3); -lean_ctor_set(x_27, 0, x_12); -lean_ctor_set(x_27, 1, x_14); -lean_ctor_set(x_27, 2, x_16); -lean_ctor_set(x_27, 3, x_13); +lean_ctor_set(x_27, 0, x_16); +lean_ctor_set(x_27, 1, x_11); +lean_ctor_set(x_27, 2, x_13); +lean_ctor_set(x_27, 3, x_15); lean_ctor_set(x_27, 4, x_26); -lean_ctor_set_uint8(x_27, sizeof(void*)*5, x_10); -lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 1, x_15); +lean_ctor_set_uint8(x_27, sizeof(void*)*5, x_12); +lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 1, x_10); lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 2, x_4); x_28 = l_Lean_MessageLog_add(x_27, x_24); lean_ctor_set(x_20, 6, x_28); @@ -4703,15 +4703,15 @@ lean_ctor_set(x_41, 0, x_21); lean_ctor_set(x_41, 1, x_22); x_42 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_11); +lean_ctor_set(x_42, 1, x_14); x_43 = lean_alloc_ctor(0, 5, 3); -lean_ctor_set(x_43, 0, x_12); -lean_ctor_set(x_43, 1, x_14); -lean_ctor_set(x_43, 2, x_16); -lean_ctor_set(x_43, 3, x_13); +lean_ctor_set(x_43, 0, x_16); +lean_ctor_set(x_43, 1, x_11); +lean_ctor_set(x_43, 2, x_13); +lean_ctor_set(x_43, 3, x_15); lean_ctor_set(x_43, 4, x_42); -lean_ctor_set_uint8(x_43, sizeof(void*)*5, x_10); -lean_ctor_set_uint8(x_43, sizeof(void*)*5 + 1, x_15); +lean_ctor_set_uint8(x_43, sizeof(void*)*5, x_12); +lean_ctor_set_uint8(x_43, sizeof(void*)*5 + 1, x_10); lean_ctor_set_uint8(x_43, sizeof(void*)*5 + 2, x_4); x_44 = l_Lean_MessageLog_add(x_43, x_38); x_45 = lean_alloc_ctor(0, 9, 0); @@ -4741,25 +4741,25 @@ if (x_60 == 0) { lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; x_61 = lean_ctor_get(x_59, 0); -lean_inc_ref(x_55); -x_62 = l_Lean_FileMap_toPosition(x_55, x_53); -lean_dec(x_53); -x_63 = l_Lean_FileMap_toPosition(x_55, x_57); +lean_inc_ref(x_51); +x_62 = l_Lean_FileMap_toPosition(x_51, x_54); +lean_dec(x_54); +x_63 = l_Lean_FileMap_toPosition(x_51, x_57); lean_dec(x_57); x_64 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_64, 0, x_63); x_65 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_Simp_mkSimpDiagSummary_spec__3___redArg___closed__3; -if (x_52 == 0) +if (x_55 == 0) { lean_free_object(x_59); lean_dec_ref(x_50); -x_10 = x_51; -x_11 = x_61; -x_12 = x_54; -x_13 = x_65; -x_14 = x_62; -x_15 = x_56; -x_16 = x_64; +x_10 = x_52; +x_11 = x_62; +x_12 = x_53; +x_13 = x_64; +x_14 = x_61; +x_15 = x_65; +x_16 = x_56; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -4776,7 +4776,7 @@ lean_object* x_67; lean_dec_ref(x_64); lean_dec_ref(x_62); lean_dec(x_61); -lean_dec_ref(x_54); +lean_dec_ref(x_56); lean_dec_ref(x_7); x_67 = lean_box(0); lean_ctor_set(x_59, 0, x_67); @@ -4785,13 +4785,13 @@ return x_59; else { lean_free_object(x_59); -x_10 = x_51; -x_11 = x_61; -x_12 = x_54; -x_13 = x_65; -x_14 = x_62; -x_15 = x_56; -x_16 = x_64; +x_10 = x_52; +x_11 = x_62; +x_12 = x_53; +x_13 = x_64; +x_14 = x_61; +x_15 = x_65; +x_16 = x_56; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -4805,24 +4805,24 @@ lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean x_68 = lean_ctor_get(x_59, 0); lean_inc(x_68); lean_dec(x_59); -lean_inc_ref(x_55); -x_69 = l_Lean_FileMap_toPosition(x_55, x_53); -lean_dec(x_53); -x_70 = l_Lean_FileMap_toPosition(x_55, x_57); +lean_inc_ref(x_51); +x_69 = l_Lean_FileMap_toPosition(x_51, x_54); +lean_dec(x_54); +x_70 = l_Lean_FileMap_toPosition(x_51, x_57); lean_dec(x_57); x_71 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_71, 0, x_70); x_72 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Meta_Simp_mkSimpDiagSummary_spec__3___redArg___closed__3; -if (x_52 == 0) +if (x_55 == 0) { lean_dec_ref(x_50); -x_10 = x_51; -x_11 = x_68; -x_12 = x_54; -x_13 = x_72; -x_14 = x_69; -x_15 = x_56; -x_16 = x_71; +x_10 = x_52; +x_11 = x_69; +x_12 = x_53; +x_13 = x_71; +x_14 = x_68; +x_15 = x_72; +x_16 = x_56; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -4839,7 +4839,7 @@ lean_object* x_74; lean_object* x_75; lean_dec_ref(x_71); lean_dec_ref(x_69); lean_dec(x_68); -lean_dec_ref(x_54); +lean_dec_ref(x_56); lean_dec_ref(x_7); x_74 = lean_box(0); x_75 = lean_alloc_ctor(0, 1, 0); @@ -4848,13 +4848,13 @@ return x_75; } else { -x_10 = x_51; -x_11 = x_68; -x_12 = x_54; -x_13 = x_72; -x_14 = x_69; -x_15 = x_56; -x_16 = x_71; +x_10 = x_52; +x_11 = x_69; +x_12 = x_53; +x_13 = x_71; +x_14 = x_68; +x_15 = x_72; +x_16 = x_56; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -4866,18 +4866,18 @@ goto block_49; block_87: { lean_object* x_85; -x_85 = l_Lean_Syntax_getTailPos_x3f(x_81, x_78); -lean_dec(x_81); +x_85 = l_Lean_Syntax_getTailPos_x3f(x_80, x_81); +lean_dec(x_80); if (lean_obj_tag(x_85) == 0) { lean_inc(x_84); x_50 = x_77; x_51 = x_78; x_52 = x_79; -x_53 = x_84; -x_54 = x_80; -x_55 = x_82; -x_56 = x_83; +x_53 = x_81; +x_54 = x_84; +x_55 = x_83; +x_56 = x_82; x_57 = x_84; goto block_76; } @@ -4890,10 +4890,10 @@ lean_dec_ref(x_85); x_50 = x_77; x_51 = x_78; x_52 = x_79; -x_53 = x_84; -x_54 = x_80; -x_55 = x_82; -x_56 = x_83; +x_53 = x_81; +x_54 = x_84; +x_55 = x_83; +x_56 = x_82; x_57 = x_86; goto block_76; } @@ -4901,20 +4901,20 @@ goto block_76; block_99: { lean_object* x_95; lean_object* x_96; -x_95 = l_Lean_replaceRef(x_1, x_91); -lean_dec(x_91); -x_96 = l_Lean_Syntax_getPos_x3f(x_95, x_89); +x_95 = l_Lean_replaceRef(x_1, x_93); +lean_dec(x_93); +x_96 = l_Lean_Syntax_getPos_x3f(x_95, x_90); if (lean_obj_tag(x_96) == 0) { lean_object* x_97; x_97 = lean_unsigned_to_nat(0u); x_77 = x_88; x_78 = x_89; -x_79 = x_90; -x_80 = x_92; -x_81 = x_95; -x_82 = x_93; -x_83 = x_94; +x_79 = x_94; +x_80 = x_95; +x_81 = x_90; +x_82 = x_92; +x_83 = x_91; x_84 = x_97; goto block_87; } @@ -4926,11 +4926,11 @@ lean_inc(x_98); lean_dec_ref(x_96); x_77 = x_88; x_78 = x_89; -x_79 = x_90; -x_80 = x_92; -x_81 = x_95; -x_82 = x_93; -x_83 = x_94; +x_79 = x_94; +x_80 = x_95; +x_81 = x_90; +x_82 = x_92; +x_83 = x_91; x_84 = x_98; goto block_87; } @@ -4939,22 +4939,22 @@ goto block_87; { if (x_107 == 0) { -x_88 = x_103; -x_89 = x_106; -x_90 = x_102; -x_91 = x_101; -x_92 = x_104; +x_88 = x_102; +x_89 = x_101; +x_90 = x_106; +x_91 = x_104; +x_92 = x_103; x_93 = x_105; x_94 = x_3; goto block_99; } else { -x_88 = x_103; -x_89 = x_106; -x_90 = x_102; -x_91 = x_101; -x_92 = x_104; +x_88 = x_102; +x_89 = x_101; +x_90 = x_106; +x_91 = x_104; +x_92 = x_103; x_93 = x_105; x_94 = x_100; goto block_99; @@ -4979,14 +4979,14 @@ x_118 = 1; x_119 = l_Lean_instBEqMessageSeverity_beq(x_3, x_118); if (x_119 == 0) { -lean_inc_ref(x_111); -lean_inc_ref(x_110); lean_inc(x_113); -x_101 = x_113; -x_102 = x_114; -x_103 = x_117; -x_104 = x_110; -x_105 = x_111; +lean_inc_ref(x_110); +lean_inc_ref(x_111); +x_101 = x_111; +x_102 = x_117; +x_103 = x_110; +x_104 = x_114; +x_105 = x_113; x_106 = x_109; x_107 = x_119; goto block_108; @@ -4996,14 +4996,14 @@ else lean_object* x_120; uint8_t x_121; x_120 = l_Lean_logAt___at___00Lean_log___at___00Lean_logInfo___at___00Lean_Meta_Simp_reportDiag_spec__0_spec__0_spec__1___closed__0; x_121 = l_Lean_Option_get___at___00Lean_logAt___at___00Lean_log___at___00Lean_logInfo___at___00Lean_Meta_Simp_reportDiag_spec__0_spec__0_spec__1_spec__3(x_112, x_120); -lean_inc_ref(x_111); -lean_inc_ref(x_110); lean_inc(x_113); -x_101 = x_113; -x_102 = x_114; -x_103 = x_117; -x_104 = x_110; -x_105 = x_111; +lean_inc_ref(x_110); +lean_inc_ref(x_111); +x_101 = x_111; +x_102 = x_117; +x_103 = x_110; +x_104 = x_114; +x_105 = x_113; x_106 = x_109; x_107 = x_121; goto block_108; diff --git a/stage0/stdlib/Lean/Meta/Tactic/SolveByElim.c b/stage0/stdlib/Lean/Meta/Tactic/SolveByElim.c index d7adcf0c33d5..e8b1e9dded69 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/SolveByElim.c +++ b/stage0/stdlib/Lean/Meta/Tactic/SolveByElim.c @@ -10761,8 +10761,8 @@ if (x_24 == 0) { if (x_1 == 0) { -lean_dec_ref(x_19); -x_11 = x_17; +lean_dec_ref(x_21); +x_11 = x_20; x_12 = x_23; x_13 = lean_box(0); goto block_16; @@ -10773,10 +10773,10 @@ if (x_2 == 0) { lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_dec(x_23); -lean_dec_ref(x_17); +lean_dec_ref(x_20); x_25 = l_Lean_Meta_SolveByElim_mkAssumptionSet___closed__1; -x_26 = l_Lean_throwError___at___00Lean_Meta_SolveByElim_SolveByElimConfig_testPartialSolutions_spec__3___redArg(x_25, x_21, x_18, x_19, x_20); -lean_dec_ref(x_19); +x_26 = l_Lean_throwError___at___00Lean_Meta_SolveByElim_SolveByElimConfig_testPartialSolutions_spec__3___redArg(x_25, x_19, x_18, x_21, x_22); +lean_dec_ref(x_21); x_27 = !lean_is_exclusive(x_26); if (x_27 == 0) { @@ -10795,8 +10795,8 @@ return x_29; } else { -lean_dec_ref(x_19); -x_11 = x_17; +lean_dec_ref(x_21); +x_11 = x_20; x_12 = x_23; x_13 = lean_box(0); goto block_16; @@ -10805,8 +10805,8 @@ goto block_16; } else { -lean_dec_ref(x_19); -x_11 = x_17; +lean_dec_ref(x_21); +x_11 = x_20; x_12 = x_23; x_13 = lean_box(0); goto block_16; @@ -10816,35 +10816,35 @@ goto block_16; { lean_object* x_42; lean_object* x_43; x_42 = lean_array_to_list(x_41); -lean_inc(x_39); -x_43 = l_List_mapTR_loop___at___00Lean_Meta_SolveByElim_mkAssumptionSet_spec__4(x_42, x_39); +lean_inc(x_40); +x_43 = l_List_mapTR_loop___at___00Lean_Meta_SolveByElim_mkAssumptionSet_spec__4(x_42, x_40); if (x_1 == 0) { lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = l_List_mapTR_loop___at___00Lean_Meta_SolveByElim_mkAssumptionSet_spec__3(x_3, x_39); +x_44 = l_List_mapTR_loop___at___00Lean_Meta_SolveByElim_mkAssumptionSet_spec__3(x_3, x_40); x_45 = l_List_appendTR___redArg(x_44, x_43); -x_46 = l_List_appendTR___redArg(x_45, x_34); -x_17 = x_33; -x_18 = x_35; -x_19 = x_36; -x_20 = x_38; -x_21 = x_37; -x_22 = lean_box(0); +x_46 = l_List_appendTR___redArg(x_45, x_35); +x_17 = lean_box(0); +x_18 = x_36; +x_19 = x_37; +x_20 = x_33; +x_21 = x_38; +x_22 = x_39; x_23 = x_46; goto block_30; } else { lean_object* x_47; lean_object* x_48; -lean_dec(x_34); -x_47 = l_List_mapTR_loop___at___00Lean_Meta_SolveByElim_mkAssumptionSet_spec__3(x_3, x_39); +lean_dec(x_35); +x_47 = l_List_mapTR_loop___at___00Lean_Meta_SolveByElim_mkAssumptionSet_spec__3(x_3, x_40); x_48 = l_List_appendTR___redArg(x_47, x_43); -x_17 = x_33; -x_18 = x_35; -x_19 = x_36; -x_20 = x_38; -x_21 = x_37; -x_22 = lean_box(0); +x_17 = lean_box(0); +x_18 = x_36; +x_19 = x_37; +x_20 = x_33; +x_21 = x_38; +x_22 = x_39; x_23 = x_48; goto block_30; } @@ -10943,13 +10943,13 @@ x_99 = lean_nat_dec_lt(x_97, x_98); if (x_99 == 0) { lean_dec(x_64); -x_34 = x_95; -x_35 = x_51; -x_36 = x_52; +x_34 = lean_box(0); +x_35 = x_95; +x_36 = x_51; x_37 = x_50; -x_38 = x_53; -x_39 = x_73; -x_40 = lean_box(0); +x_38 = x_52; +x_39 = x_53; +x_40 = x_73; x_41 = x_96; goto block_49; } @@ -10962,13 +10962,13 @@ if (x_100 == 0) if (x_99 == 0) { lean_dec(x_64); -x_34 = x_95; -x_35 = x_51; -x_36 = x_52; +x_34 = lean_box(0); +x_35 = x_95; +x_36 = x_51; x_37 = x_50; -x_38 = x_53; -x_39 = x_73; -x_40 = lean_box(0); +x_38 = x_52; +x_39 = x_53; +x_40 = x_73; x_41 = x_96; goto block_49; } @@ -10978,13 +10978,13 @@ size_t x_101; lean_object* x_102; x_101 = lean_usize_of_nat(x_98); x_102 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Meta_SolveByElim_mkAssumptionSet_spec__5(x_64, x_62, x_101, x_96); lean_dec(x_64); -x_34 = x_95; -x_35 = x_51; -x_36 = x_52; +x_34 = lean_box(0); +x_35 = x_95; +x_36 = x_51; x_37 = x_50; -x_38 = x_53; -x_39 = x_73; -x_40 = lean_box(0); +x_38 = x_52; +x_39 = x_53; +x_40 = x_73; x_41 = x_102; goto block_49; } @@ -10995,13 +10995,13 @@ size_t x_103; lean_object* x_104; x_103 = lean_usize_of_nat(x_98); x_104 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Meta_SolveByElim_mkAssumptionSet_spec__5(x_64, x_62, x_103, x_96); lean_dec(x_64); -x_34 = x_95; -x_35 = x_51; -x_36 = x_52; +x_34 = lean_box(0); +x_35 = x_95; +x_36 = x_51; x_37 = x_50; -x_38 = x_53; -x_39 = x_73; -x_40 = lean_box(0); +x_38 = x_52; +x_39 = x_53; +x_40 = x_73; x_41 = x_104; goto block_49; } diff --git a/stage0/stdlib/Lean/Meta/Tactic/UnifyEq.c b/stage0/stdlib/Lean/Meta/Tactic/UnifyEq.c index 9763fc14e558..a6772da5d910 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/UnifyEq.c +++ b/stage0/stdlib/Lean/Meta/Tactic/UnifyEq.c @@ -1616,7 +1616,7 @@ lean_inc(x_12); lean_inc_ref(x_11); lean_inc(x_10); lean_inc_ref(x_9); -x_17 = l_Lean_Meta_mkEq(x_15, x_14, x_9, x_10, x_11, x_12); +x_17 = l_Lean_Meta_mkEq(x_16, x_14, x_9, x_10, x_11, x_12); if (lean_obj_tag(x_17) == 0) { lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; @@ -1792,8 +1792,8 @@ lean_dec(x_45); lean_dec_ref(x_8); lean_dec(x_4); x_14 = x_50; -x_15 = x_48; -x_16 = lean_box(0); +x_15 = lean_box(0); +x_16 = x_48; goto block_43; } else @@ -1806,8 +1806,8 @@ if (x_52 == 0) lean_dec(x_45); lean_dec(x_4); x_14 = x_50; -x_15 = x_48; -x_16 = lean_box(0); +x_15 = lean_box(0); +x_16 = x_48; goto block_43; } else diff --git a/stage0/stdlib/Lean/MetavarContext.c b/stage0/stdlib/Lean/MetavarContext.c index cbaf2397147c..23e056a46b0a 100644 --- a/stage0/stdlib/Lean/MetavarContext.c +++ b/stage0/stdlib/Lean/MetavarContext.c @@ -17835,7 +17835,7 @@ goto _start; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimMVar(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t 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_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; lean_object* x_52; uint8_t x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_115; lean_object* x_116; lean_object* x_117; uint8_t x_118; lean_object* x_119; lean_object* x_120; x_41 = lean_ctor_get(x_6, 0); lean_inc_ref(x_41); x_42 = l_Lean_MetavarContext_getDecl(x_41, x_2); @@ -17918,11 +17918,11 @@ x_179 = lean_nat_dec_lt(x_50, x_177); if (x_179 == 0) { lean_dec_ref(x_45); -x_115 = x_167; +x_115 = x_175; x_116 = x_171; -x_117 = x_175; -x_118 = x_176; -x_119 = x_174; +x_117 = x_174; +x_118 = x_167; +x_119 = x_176; x_120 = x_178; goto block_162; } @@ -17935,11 +17935,11 @@ if (x_180 == 0) if (x_179 == 0) { lean_dec_ref(x_45); -x_115 = x_167; +x_115 = x_175; x_116 = x_171; -x_117 = x_175; -x_118 = x_176; -x_119 = x_174; +x_117 = x_174; +x_118 = x_167; +x_119 = x_176; x_120 = x_178; goto block_162; } @@ -17949,11 +17949,11 @@ size_t x_181; lean_object* x_182; x_181 = lean_usize_of_nat(x_177); x_182 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimMVar_spec__12(x_174, x_50, x_45, x_169, x_181, x_178); lean_dec_ref(x_45); -x_115 = x_167; +x_115 = x_175; x_116 = x_171; -x_117 = x_175; -x_118 = x_176; -x_119 = x_174; +x_117 = x_174; +x_118 = x_167; +x_119 = x_176; x_120 = x_182; goto block_162; } @@ -17964,11 +17964,11 @@ size_t x_183; lean_object* x_184; x_183 = lean_usize_of_nat(x_177); x_184 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimMVar_spec__12(x_174, x_50, x_45, x_169, x_183, x_178); lean_dec_ref(x_45); -x_115 = x_167; +x_115 = x_175; x_116 = x_171; -x_117 = x_175; -x_118 = x_176; -x_119 = x_174; +x_117 = x_174; +x_118 = x_167; +x_119 = x_176; x_120 = x_184; goto block_162; } @@ -18163,10 +18163,10 @@ return x_13; block_29: { lean_object* x_22; lean_object* x_23; -lean_inc_ref(x_18); -x_22 = l_Array_append___redArg(x_18, x_20); +lean_inc_ref(x_17); +x_22 = l_Array_append___redArg(x_17, x_20); lean_dec_ref(x_20); -x_23 = l_Lean_assignDelayedMVar___at___00__private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimMVar_spec__9___redArg(x_15, x_22, x_19, x_21); +x_23 = l_Lean_assignDelayedMVar___at___00__private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimMVar_spec__9___redArg(x_18, x_22, x_19, x_21); if (lean_obj_tag(x_23) == 0) { lean_object* x_24; @@ -18174,17 +18174,17 @@ x_24 = lean_ctor_get(x_23, 1); lean_inc(x_24); lean_dec_ref(x_23); x_7 = x_16; -x_8 = x_17; -x_9 = x_18; +x_8 = x_15; +x_9 = x_17; x_10 = x_24; goto block_14; } else { uint8_t x_25; -lean_dec_ref(x_18); lean_dec_ref(x_17); lean_dec_ref(x_16); +lean_dec_ref(x_15); x_25 = !lean_is_exclusive(x_23); if (x_25 == 0) { @@ -18208,26 +18208,26 @@ return x_28; block_40: { lean_object* x_34; -lean_inc_ref(x_32); -x_34 = l_Lean_MVarId_assign___at___00__private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimMVar_spec__10___redArg(x_2, x_32, x_30); +lean_inc_ref(x_31); +x_34 = l_Lean_MVarId_assign___at___00__private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimMVar_spec__10___redArg(x_2, x_31, x_33); if (lean_obj_tag(x_34) == 0) { lean_object* x_35; x_35 = lean_ctor_get(x_34, 1); lean_inc(x_35); lean_dec_ref(x_34); -x_7 = x_31; -x_8 = x_32; -x_9 = x_33; +x_7 = x_30; +x_8 = x_31; +x_9 = x_32; x_10 = x_35; goto block_14; } else { uint8_t x_36; -lean_dec_ref(x_33); lean_dec_ref(x_32); lean_dec_ref(x_31); +lean_dec_ref(x_30); x_36 = !lean_is_exclusive(x_34); if (x_36 == 0) { @@ -18262,14 +18262,14 @@ lean_inc(x_63); x_65 = l_Lean_Name_num___override(x_63, x_64); lean_inc(x_65); x_66 = l_Lean_mkMVar(x_65); -x_67 = l___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_mkMVarApp(x_43, x_66, x_56, x_53); +x_67 = l___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_mkMVarApp(x_43, x_66, x_53, x_54); x_68 = l_Lean_Expr_getAppNumArgs(x_67); x_69 = lean_nat_add(x_47, x_68); lean_dec(x_68); lean_dec(x_47); x_70 = lean_box(0); lean_inc(x_65); -x_71 = l_Lean_MetavarContext_addExprMVarDecl(x_58, x_65, x_70, x_55, x_52, x_57, x_53, x_69); +x_71 = l_Lean_MetavarContext_addExprMVarDecl(x_58, x_65, x_70, x_56, x_55, x_57, x_54, x_69); x_72 = lean_unsigned_to_nat(1u); x_73 = lean_nat_add(x_64, x_72); lean_dec(x_64); @@ -18283,10 +18283,10 @@ x_75 = l_Lean_MetavarKind_isSyntheticOpaque(x_46); if (x_75 == 0) { lean_dec(x_65); -x_30 = x_74; -x_31 = x_54; -x_32 = x_67; -x_33 = x_56; +x_30 = x_52; +x_31 = x_67; +x_32 = x_53; +x_33 = x_74; goto block_40; } else @@ -18307,10 +18307,10 @@ x_78 = lean_ctor_get(x_76, 1); lean_inc(x_78); lean_dec_ref(x_76); x_79 = l___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_getInScope___closed__0; -x_15 = x_65; -x_16 = x_54; -x_17 = x_67; -x_18 = x_56; +x_15 = x_67; +x_16 = x_52; +x_17 = x_53; +x_18 = x_65; x_19 = x_2; x_20 = x_79; x_21 = x_78; @@ -18331,10 +18331,10 @@ lean_inc_ref(x_82); x_83 = lean_ctor_get(x_80, 1); lean_inc(x_83); lean_dec(x_80); -x_15 = x_65; -x_16 = x_54; -x_17 = x_67; -x_18 = x_56; +x_15 = x_67; +x_16 = x_52; +x_17 = x_53; +x_18 = x_65; x_19 = x_83; x_20 = x_82; x_21 = x_81; @@ -18346,8 +18346,8 @@ else uint8_t x_84; lean_dec_ref(x_67); lean_dec(x_65); -lean_dec_ref(x_56); -lean_dec_ref(x_54); +lean_dec_ref(x_53); +lean_dec_ref(x_52); lean_dec(x_2); x_84 = !lean_is_exclusive(x_76); if (x_84 == 0) @@ -18372,10 +18372,10 @@ return x_87; else { lean_dec(x_65); -x_30 = x_74; -x_31 = x_54; -x_32 = x_67; -x_33 = x_56; +x_30 = x_52; +x_31 = x_67; +x_32 = x_53; +x_33 = x_74; goto block_40; } } @@ -18393,14 +18393,14 @@ lean_inc(x_88); x_90 = l_Lean_Name_num___override(x_88, x_89); lean_inc(x_90); x_91 = l_Lean_mkMVar(x_90); -x_92 = l___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_mkMVarApp(x_43, x_91, x_56, x_53); +x_92 = l___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_mkMVarApp(x_43, x_91, x_53, x_54); x_93 = l_Lean_Expr_getAppNumArgs(x_92); x_94 = lean_nat_add(x_47, x_93); lean_dec(x_93); lean_dec(x_47); x_95 = lean_box(0); lean_inc(x_90); -x_96 = l_Lean_MetavarContext_addExprMVarDecl(x_58, x_90, x_95, x_55, x_52, x_57, x_53, x_94); +x_96 = l_Lean_MetavarContext_addExprMVarDecl(x_58, x_90, x_95, x_56, x_55, x_57, x_54, x_94); x_97 = lean_unsigned_to_nat(1u); x_98 = lean_nat_add(x_89, x_97); lean_dec(x_89); @@ -18416,10 +18416,10 @@ x_101 = l_Lean_MetavarKind_isSyntheticOpaque(x_46); if (x_101 == 0) { lean_dec(x_90); -x_30 = x_100; -x_31 = x_54; -x_32 = x_92; -x_33 = x_56; +x_30 = x_52; +x_31 = x_92; +x_32 = x_53; +x_33 = x_100; goto block_40; } else @@ -18440,10 +18440,10 @@ x_104 = lean_ctor_get(x_102, 1); lean_inc(x_104); lean_dec_ref(x_102); x_105 = l___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_getInScope___closed__0; -x_15 = x_90; -x_16 = x_54; -x_17 = x_92; -x_18 = x_56; +x_15 = x_92; +x_16 = x_52; +x_17 = x_53; +x_18 = x_90; x_19 = x_2; x_20 = x_105; x_21 = x_104; @@ -18464,10 +18464,10 @@ lean_inc_ref(x_108); x_109 = lean_ctor_get(x_106, 1); lean_inc(x_109); lean_dec(x_106); -x_15 = x_90; -x_16 = x_54; -x_17 = x_92; -x_18 = x_56; +x_15 = x_92; +x_16 = x_52; +x_17 = x_53; +x_18 = x_90; x_19 = x_109; x_20 = x_108; x_21 = x_107; @@ -18479,8 +18479,8 @@ else lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_dec_ref(x_92); lean_dec(x_90); -lean_dec_ref(x_56); -lean_dec_ref(x_54); +lean_dec_ref(x_53); +lean_dec_ref(x_52); lean_dec(x_2); x_110 = lean_ctor_get(x_102, 0); lean_inc(x_110); @@ -18507,10 +18507,10 @@ return x_113; else { lean_dec(x_90); -x_30 = x_100; -x_31 = x_54; -x_32 = x_92; -x_33 = x_56; +x_30 = x_52; +x_31 = x_92; +x_32 = x_53; +x_33 = x_100; goto block_40; } } @@ -18519,15 +18519,15 @@ goto block_40; block_162: { uint8_t x_121; -x_121 = !lean_is_exclusive(x_117); +x_121 = !lean_is_exclusive(x_115); if (x_121 == 0) { lean_object* x_122; lean_object* x_123; lean_object* x_124; -x_122 = lean_ctor_get(x_117, 3); +x_122 = lean_ctor_get(x_115, 3); x_123 = l_Lean_instantiateMVarsCore___lam__0___closed__1; -lean_ctor_set(x_117, 3, x_123); +lean_ctor_set(x_115, 3, x_123); lean_inc_ref(x_43); -x_124 = l___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_mkAuxMVarType(x_43, x_119, x_115, x_44, x_4, x_5, x_117); +x_124 = l___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_mkAuxMVarType(x_43, x_117, x_118, x_44, x_4, x_5, x_115); if (lean_obj_tag(x_124) == 0) { lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; @@ -18543,10 +18543,10 @@ lean_inc(x_128); x_129 = lean_ctor_get(x_125, 2); lean_inc_ref(x_129); lean_dec(x_125); -x_52 = x_120; -x_53 = x_115; -x_54 = x_116; -x_55 = x_118; +x_52 = x_116; +x_53 = x_117; +x_54 = x_118; +x_55 = x_120; x_56 = x_119; x_57 = x_126; x_58 = x_127; @@ -18575,10 +18575,10 @@ lean_inc_ref(x_134); x_135 = lean_ctor_get(x_130, 3); lean_inc_ref(x_135); lean_dec(x_130); -x_52 = x_120; -x_53 = x_115; -x_54 = x_116; -x_55 = x_118; +x_52 = x_116; +x_53 = x_117; +x_54 = x_118; +x_55 = x_120; x_56 = x_119; x_57 = x_131; x_58 = x_132; @@ -18592,7 +18592,7 @@ else uint8_t x_136; lean_dec_ref(x_120); lean_dec_ref(x_119); -lean_dec_ref(x_118); +lean_dec_ref(x_117); lean_dec_ref(x_116); lean_dec(x_47); lean_dec_ref(x_43); @@ -18621,15 +18621,15 @@ return x_139; else { lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; -x_140 = lean_ctor_get(x_117, 0); -x_141 = lean_ctor_get(x_117, 1); -x_142 = lean_ctor_get(x_117, 2); -x_143 = lean_ctor_get(x_117, 3); +x_140 = lean_ctor_get(x_115, 0); +x_141 = lean_ctor_get(x_115, 1); +x_142 = lean_ctor_get(x_115, 2); +x_143 = lean_ctor_get(x_115, 3); lean_inc(x_143); lean_inc(x_142); lean_inc(x_141); lean_inc(x_140); -lean_dec(x_117); +lean_dec(x_115); x_144 = l_Lean_instantiateMVarsCore___lam__0___closed__1; x_145 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_145, 0, x_140); @@ -18637,7 +18637,7 @@ lean_ctor_set(x_145, 1, x_141); lean_ctor_set(x_145, 2, x_142); lean_ctor_set(x_145, 3, x_144); lean_inc_ref(x_43); -x_146 = l___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_mkAuxMVarType(x_43, x_119, x_115, x_44, x_4, x_5, x_145); +x_146 = l___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_mkAuxMVarType(x_43, x_117, x_118, x_44, x_4, x_5, x_145); if (lean_obj_tag(x_146) == 0) { lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; @@ -18653,10 +18653,10 @@ lean_inc(x_150); x_151 = lean_ctor_get(x_147, 2); lean_inc_ref(x_151); lean_dec(x_147); -x_52 = x_120; -x_53 = x_115; -x_54 = x_116; -x_55 = x_118; +x_52 = x_116; +x_53 = x_117; +x_54 = x_118; +x_55 = x_120; x_56 = x_119; x_57 = x_148; x_58 = x_149; @@ -18685,10 +18685,10 @@ lean_inc_ref(x_156); x_157 = lean_ctor_get(x_152, 3); lean_inc_ref(x_157); lean_dec(x_152); -x_52 = x_120; -x_53 = x_115; -x_54 = x_116; -x_55 = x_118; +x_52 = x_116; +x_53 = x_117; +x_54 = x_118; +x_55 = x_120; x_56 = x_119; x_57 = x_153; x_58 = x_154; @@ -18702,7 +18702,7 @@ else lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_dec_ref(x_120); lean_dec_ref(x_119); -lean_dec_ref(x_118); +lean_dec_ref(x_117); lean_dec_ref(x_116); lean_dec(x_47); lean_dec_ref(x_43); @@ -20868,7 +20868,7 @@ return x_14; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; lean_object* x_26; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_64; uint8_t x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; lean_object* x_69; lean_object* x_70; +lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_26; lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_44; uint8_t x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_64; uint8_t x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; lean_object* x_69; lean_object* x_70; x_15 = lean_unsigned_to_nat(1u); x_16 = lean_nat_sub(x_8, x_15); lean_dec(x_8); @@ -21164,7 +21164,7 @@ goto block_126; if (x_2 == 0) { lean_object* x_21; -x_21 = l_Lean_mkForall(x_17, x_18, x_19, x_9); +x_21 = l_Lean_mkForall(x_18, x_17, x_19, x_9); x_8 = x_16; x_9 = x_21; x_11 = x_20; @@ -21173,7 +21173,7 @@ goto _start; else { lean_object* x_23; -x_23 = l___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_mkLambda_x27(x_17, x_18, x_19, x_9, x_3); +x_23 = l___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_mkLambda_x27(x_18, x_17, x_19, x_9, x_3); x_8 = x_16; x_9 = x_23; x_11 = x_20; @@ -21206,8 +21206,8 @@ return x_26; block_43: { lean_object* x_36; lean_object* x_37; -x_36 = l_Lean_Expr_headBeta(x_34); -x_37 = l_Lean_MetavarContext_MkBinding_elimMVarDeps(x_1, x_36, x_33, x_32); +x_36 = l_Lean_Expr_headBeta(x_35); +x_37 = l_Lean_MetavarContext_MkBinding_elimMVarDeps(x_1, x_36, x_31, x_34); if (lean_obj_tag(x_37) == 0) { lean_object* x_38; lean_object* x_39; lean_object* x_40; @@ -21218,8 +21218,8 @@ lean_inc(x_39); lean_dec_ref(x_37); x_40 = lean_expr_abstract_range(x_38, x_16, x_1); lean_dec(x_38); -x_17 = x_31; -x_18 = x_35; +x_17 = x_32; +x_18 = x_33; x_19 = x_40; x_20 = x_39; goto block_25; @@ -21234,15 +21234,15 @@ lean_inc(x_41); x_42 = lean_ctor_get(x_37, 1); lean_inc(x_42); lean_dec_ref(x_37); -x_17 = x_31; -x_18 = x_35; +x_17 = x_32; +x_18 = x_33; x_19 = x_41; x_20 = x_42; goto block_25; } else { -lean_dec(x_31); +lean_dec(x_33); lean_dec_ref(x_9); x_26 = x_37; goto block_30; @@ -21257,22 +21257,22 @@ if (x_49 == 0) { lean_object* x_50; lean_dec_ref(x_47); -lean_dec_ref(x_45); -lean_dec(x_44); +lean_dec(x_46); +lean_dec_ref(x_44); x_50 = lean_expr_lower_loose_bvars(x_9, x_15, x_15); lean_dec_ref(x_9); x_8 = x_16; x_9 = x_50; -x_11 = x_46; +x_11 = x_48; goto _start; } else { x_31 = x_44; -x_32 = x_46; -x_33 = x_45; -x_34 = x_47; -x_35 = x_48; +x_32 = x_45; +x_33 = x_46; +x_34 = x_48; +x_35 = x_47; goto block_43; } } @@ -21307,30 +21307,30 @@ if (x_4 == 0) { if (x_65 == 0) { -x_44 = x_66; -x_45 = x_69; -x_46 = x_70; +x_44 = x_69; +x_45 = x_68; +x_46 = x_66; x_47 = x_67; -x_48 = x_68; +x_48 = x_70; goto block_52; } else { -x_31 = x_66; -x_32 = x_70; -x_33 = x_69; -x_34 = x_67; -x_35 = x_68; +x_31 = x_69; +x_32 = x_68; +x_33 = x_66; +x_34 = x_70; +x_35 = x_67; goto block_43; } } else { -x_44 = x_66; -x_45 = x_69; -x_46 = x_70; +x_44 = x_69; +x_45 = x_68; +x_46 = x_66; x_47 = x_67; -x_48 = x_68; +x_48 = x_70; goto block_52; } } diff --git a/stage0/stdlib/Lean/Server/InfoUtils.c b/stage0/stdlib/Lean/Server/InfoUtils.c index c590d2c4b212..484b5dd32224 100644 --- a/stage0/stdlib/Lean/Server/InfoUtils.c +++ b/stage0/stdlib/Lean/Server/InfoUtils.c @@ -10446,22 +10446,22 @@ return x_20; } block_28: { -if (lean_obj_tag(x_22) == 1) +if (lean_obj_tag(x_23) == 1) { lean_object* x_26; lean_object* x_27; -x_26 = lean_ctor_get(x_22, 0); +x_26 = lean_ctor_get(x_23, 0); lean_inc(x_26); -lean_dec_ref(x_22); +lean_dec_ref(x_23); x_27 = lean_array_push(x_24, x_26); -x_9 = x_23; +x_9 = x_22; x_10 = x_27; x_11 = lean_box(0); goto block_21; } else { -lean_dec(x_22); -x_9 = x_23; +lean_dec(x_23); +x_9 = x_22; x_10 = x_24; x_11 = lean_box(0); goto block_21; @@ -10490,8 +10490,8 @@ if (x_35 == 0) lean_object* x_36; lean_ctor_set_tag(x_34, 3); x_36 = lean_array_push(x_30, x_34); -x_22 = x_29; -x_23 = x_31; +x_22 = x_31; +x_23 = x_29; x_24 = x_36; x_25 = lean_box(0); goto block_28; @@ -10505,8 +10505,8 @@ lean_dec(x_34); x_38 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_38, 0, x_37); x_39 = lean_array_push(x_30, x_38); -x_22 = x_29; -x_23 = x_31; +x_22 = x_31; +x_23 = x_29; x_24 = x_39; x_25 = lean_box(0); goto block_28; @@ -10515,8 +10515,8 @@ goto block_28; else { lean_dec(x_34); -x_22 = x_29; -x_23 = x_31; +x_22 = x_31; +x_23 = x_29; x_24 = x_30; x_25 = lean_box(0); goto block_28; @@ -12123,7 +12123,7 @@ return x_4; LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_hoverableInfoAtM_x3f___at___00Lean_Elab_InfoTree_termGoalAt_x3f_spec__1___lam__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_8; uint8_t x_9; uint8_t x_10; uint8_t x_11; lean_object* x_17; uint8_t x_18; uint8_t x_19; uint8_t x_20; uint8_t x_21; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +uint8_t x_8; lean_object* x_9; uint8_t x_10; uint8_t x_11; uint8_t x_17; uint8_t x_18; uint8_t x_19; lean_object* x_20; uint8_t x_21; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; x_23 = l_Lean_Elab_InfoTree_hoverableInfoAtM_x3f___redArg___lam__9___closed__0; x_24 = l_List_filterMapTR_go___at___00Lean_Elab_InfoTree_hoverableInfoAtM_x3f___at___00Lean_Elab_InfoTree_termGoalAt_x3f_spec__1_spec__1(x_7, x_23); lean_inc_ref(x_6); @@ -12237,29 +12237,29 @@ x_44 = lean_ctor_get(x_5, 0); x_45 = lean_ctor_get(x_44, 3); if (lean_obj_tag(x_45) == 1) { -x_17 = x_43; -x_18 = x_42; -x_19 = x_32; -x_20 = x_33; +x_17 = x_33; +x_18 = x_32; +x_19 = x_42; +x_20 = x_43; x_21 = x_33; goto block_22; } else { -x_17 = x_43; -x_18 = x_42; -x_19 = x_32; -x_20 = x_33; +x_17 = x_33; +x_18 = x_32; +x_19 = x_42; +x_20 = x_43; x_21 = x_32; goto block_22; } } else { -x_17 = x_43; -x_18 = x_42; -x_19 = x_32; -x_20 = x_33; +x_17 = x_33; +x_18 = x_32; +x_19 = x_42; +x_20 = x_43; x_21 = x_32; goto block_22; } @@ -12345,9 +12345,9 @@ return x_54; { lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; x_12 = lean_alloc_ctor(0, 1, 3); -lean_ctor_set(x_12, 0, x_8); -lean_ctor_set_uint8(x_12, sizeof(void*)*1, x_10); -lean_ctor_set_uint8(x_12, sizeof(void*)*1 + 1, x_9); +lean_ctor_set(x_12, 0, x_9); +lean_ctor_set_uint8(x_12, sizeof(void*)*1, x_8); +lean_ctor_set_uint8(x_12, sizeof(void*)*1 + 1, x_10); lean_ctor_set_uint8(x_12, sizeof(void*)*1 + 2, x_11); x_13 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_13, 0, x_4); @@ -12364,18 +12364,18 @@ return x_15; { if (lean_obj_tag(x_5) == 2) { -x_8 = x_17; -x_9 = x_21; -x_10 = x_18; -x_11 = x_20; +x_8 = x_19; +x_9 = x_20; +x_10 = x_21; +x_11 = x_17; goto block_16; } else { -x_8 = x_17; -x_9 = x_21; -x_10 = x_18; -x_11 = x_19; +x_8 = x_19; +x_9 = x_20; +x_10 = x_21; +x_11 = x_18; goto block_16; } } diff --git a/stage0/stdlib/Lean/Server/Test/Runner.c b/stage0/stdlib/Lean/Server/Test/Runner.c index b5a41a32c7eb..f99288301f68 100644 --- a/stage0/stdlib/Lean/Server/Test/Runner.c +++ b/stage0/stdlib/Lean/Server/Test/Runner.c @@ -20485,8 +20485,8 @@ x_62 = 97; x_63 = lean_uint32_dec_le(x_62, x_59); if (x_63 == 0) { -lean_dec_ref(x_60); -x_55 = x_61; +lean_dec_ref(x_61); +x_55 = x_60; goto block_58; } else @@ -20496,15 +20496,15 @@ x_64 = 122; x_65 = lean_uint32_dec_le(x_59, x_64); if (x_65 == 0) { -lean_dec_ref(x_60); -x_55 = x_61; +lean_dec_ref(x_61); +x_55 = x_60; goto block_58; } else { -lean_dec_ref(x_61); +lean_dec_ref(x_60); lean_dec(x_9); -x_3 = x_60; +x_3 = x_61; x_4 = x_59; goto block_7; } @@ -20551,8 +20551,8 @@ x_78 = lean_uint32_dec_le(x_77, x_74); if (x_78 == 0) { x_59 = x_74; -x_60 = x_76; -x_61 = x_68; +x_60 = x_68; +x_61 = x_76; goto block_66; } else @@ -20563,8 +20563,8 @@ x_80 = lean_uint32_dec_le(x_74, x_79); if (x_80 == 0) { x_59 = x_74; -x_60 = x_76; -x_61 = x_68; +x_60 = x_68; +x_61 = x_76; goto block_66; } else @@ -23115,9 +23115,9 @@ goto block_190; block_182: { lean_object* x_179; lean_object* x_180; lean_object* x_181; -x_179 = lean_nat_add(x_176, x_178); +x_179 = lean_nat_add(x_177, x_178); lean_dec(x_178); -lean_dec(x_176); +lean_dec(x_177); if (lean_is_scalar(x_172)) { x_180 = lean_alloc_ctor(0, 5, 0); } else { @@ -23136,7 +23136,7 @@ if (lean_is_scalar(x_162)) { lean_ctor_set(x_181, 0, x_175); lean_ctor_set(x_181, 1, x_165); lean_ctor_set(x_181, 2, x_166); -lean_ctor_set(x_181, 3, x_177); +lean_ctor_set(x_181, 3, x_176); lean_ctor_set(x_181, 4, x_180); return x_181; } @@ -23162,8 +23162,8 @@ if (lean_obj_tag(x_168) == 0) lean_object* x_188; x_188 = lean_ctor_get(x_168, 0); lean_inc(x_188); -x_176 = x_187; -x_177 = x_186; +x_176 = x_186; +x_177 = x_187; x_178 = x_188; goto block_182; } @@ -23171,8 +23171,8 @@ else { lean_object* x_189; x_189 = lean_unsigned_to_nat(0u); -x_176 = x_187; -x_177 = x_186; +x_176 = x_186; +x_177 = x_187; x_178 = x_189; goto block_182; } @@ -28533,7 +28533,7 @@ goto block_138; { lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; x_98 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_98, 0, x_95); +lean_ctor_set(x_98, 0, x_96); lean_ctor_set(x_98, 1, x_97); x_99 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__11; x_100 = lean_alloc_ctor(3, 1, 0); @@ -28554,13 +28554,13 @@ lean_dec(x_91); x_107 = l_List_appendTR___redArg(x_104, x_106); x_108 = l_Lean_Json_mkObj(x_107); x_109 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_109, 0, x_96); +lean_ctor_set(x_109, 0, x_94); lean_ctor_set(x_109, 1, x_108); x_110 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_110, 0, x_109); lean_ctor_set(x_110, 1, x_102); x_111 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_111, 0, x_94); +lean_ctor_set(x_111, 0, x_95); lean_ctor_set(x_111, 1, x_110); x_112 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_112, 0, x_93); @@ -28594,9 +28594,9 @@ case 0: { lean_object* x_126; x_126 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__18; -x_94 = x_123; -x_95 = x_125; -x_96 = x_124; +x_94 = x_124; +x_95 = x_123; +x_96 = x_125; x_97 = x_126; goto block_120; } @@ -28604,9 +28604,9 @@ case 1: { lean_object* x_127; x_127 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__22; -x_94 = x_123; -x_95 = x_125; -x_96 = x_124; +x_94 = x_124; +x_95 = x_123; +x_96 = x_125; x_97 = x_127; goto block_120; } @@ -28614,9 +28614,9 @@ case 2: { lean_object* x_128; x_128 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__26; -x_94 = x_123; -x_95 = x_125; -x_96 = x_124; +x_94 = x_124; +x_95 = x_123; +x_96 = x_125; x_97 = x_128; goto block_120; } @@ -28624,9 +28624,9 @@ case 3: { lean_object* x_129; x_129 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__30; -x_94 = x_123; -x_95 = x_125; -x_96 = x_124; +x_94 = x_124; +x_95 = x_123; +x_96 = x_125; x_97 = x_129; goto block_120; } @@ -28634,9 +28634,9 @@ case 4: { lean_object* x_130; x_130 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__34; -x_94 = x_123; -x_95 = x_125; -x_96 = x_124; +x_94 = x_124; +x_95 = x_123; +x_96 = x_125; x_97 = x_130; goto block_120; } @@ -28644,9 +28644,9 @@ case 5: { lean_object* x_131; x_131 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__38; -x_94 = x_123; -x_95 = x_125; -x_96 = x_124; +x_94 = x_124; +x_95 = x_123; +x_96 = x_125; x_97 = x_131; goto block_120; } @@ -28654,9 +28654,9 @@ case 6: { lean_object* x_132; x_132 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__42; -x_94 = x_123; -x_95 = x_125; -x_96 = x_124; +x_94 = x_124; +x_95 = x_123; +x_96 = x_125; x_97 = x_132; goto block_120; } @@ -28664,9 +28664,9 @@ case 7: { lean_object* x_133; x_133 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__46; -x_94 = x_123; -x_95 = x_125; -x_96 = x_124; +x_94 = x_124; +x_95 = x_123; +x_96 = x_125; x_97 = x_133; goto block_120; } @@ -28674,9 +28674,9 @@ case 8: { lean_object* x_134; x_134 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__50; -x_94 = x_123; -x_95 = x_125; -x_96 = x_124; +x_94 = x_124; +x_95 = x_123; +x_96 = x_125; x_97 = x_134; goto block_120; } @@ -28684,9 +28684,9 @@ case 9: { lean_object* x_135; x_135 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__54; -x_94 = x_123; -x_95 = x_125; -x_96 = x_124; +x_94 = x_124; +x_95 = x_123; +x_96 = x_125; x_97 = x_135; goto block_120; } @@ -28694,9 +28694,9 @@ case 10: { lean_object* x_136; x_136 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__58; -x_94 = x_123; -x_95 = x_125; -x_96 = x_124; +x_94 = x_124; +x_95 = x_123; +x_96 = x_125; x_97 = x_136; goto block_120; } @@ -28704,9 +28704,9 @@ goto block_120; { lean_object* x_137; x_137 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__62; -x_94 = x_123; -x_95 = x_125; -x_96 = x_124; +x_94 = x_124; +x_95 = x_123; +x_96 = x_125; x_97 = x_137; goto block_120; } @@ -32589,7 +32589,7 @@ goto block_138; { lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; x_98 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_98, 0, x_94); +lean_ctor_set(x_98, 0, x_96); lean_ctor_set(x_98, 1, x_97); x_99 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__11; x_100 = lean_alloc_ctor(3, 1, 0); @@ -32616,7 +32616,7 @@ x_110 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_110, 0, x_109); lean_ctor_set(x_110, 1, x_102); x_111 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_111, 0, x_96); +lean_ctor_set(x_111, 0, x_94); lean_ctor_set(x_111, 1, x_110); x_112 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_112, 0, x_93); @@ -32650,9 +32650,9 @@ case 0: { lean_object* x_126; x_126 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__18; -x_94 = x_125; +x_94 = x_123; x_95 = x_124; -x_96 = x_123; +x_96 = x_125; x_97 = x_126; goto block_120; } @@ -32660,9 +32660,9 @@ case 1: { lean_object* x_127; x_127 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__22; -x_94 = x_125; +x_94 = x_123; x_95 = x_124; -x_96 = x_123; +x_96 = x_125; x_97 = x_127; goto block_120; } @@ -32670,9 +32670,9 @@ case 2: { lean_object* x_128; x_128 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__26; -x_94 = x_125; +x_94 = x_123; x_95 = x_124; -x_96 = x_123; +x_96 = x_125; x_97 = x_128; goto block_120; } @@ -32680,9 +32680,9 @@ case 3: { lean_object* x_129; x_129 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__30; -x_94 = x_125; +x_94 = x_123; x_95 = x_124; -x_96 = x_123; +x_96 = x_125; x_97 = x_129; goto block_120; } @@ -32690,9 +32690,9 @@ case 4: { lean_object* x_130; x_130 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__34; -x_94 = x_125; +x_94 = x_123; x_95 = x_124; -x_96 = x_123; +x_96 = x_125; x_97 = x_130; goto block_120; } @@ -32700,9 +32700,9 @@ case 5: { lean_object* x_131; x_131 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__38; -x_94 = x_125; +x_94 = x_123; x_95 = x_124; -x_96 = x_123; +x_96 = x_125; x_97 = x_131; goto block_120; } @@ -32710,9 +32710,9 @@ case 6: { lean_object* x_132; x_132 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__42; -x_94 = x_125; +x_94 = x_123; x_95 = x_124; -x_96 = x_123; +x_96 = x_125; x_97 = x_132; goto block_120; } @@ -32720,9 +32720,9 @@ case 7: { lean_object* x_133; x_133 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__46; -x_94 = x_125; +x_94 = x_123; x_95 = x_124; -x_96 = x_123; +x_96 = x_125; x_97 = x_133; goto block_120; } @@ -32730,9 +32730,9 @@ case 8: { lean_object* x_134; x_134 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__50; -x_94 = x_125; +x_94 = x_123; x_95 = x_124; -x_96 = x_123; +x_96 = x_125; x_97 = x_134; goto block_120; } @@ -32740,9 +32740,9 @@ case 9: { lean_object* x_135; x_135 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__54; -x_94 = x_125; +x_94 = x_123; x_95 = x_124; -x_96 = x_123; +x_96 = x_125; x_97 = x_135; goto block_120; } @@ -32750,9 +32750,9 @@ case 10: { lean_object* x_136; x_136 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__58; -x_94 = x_125; +x_94 = x_123; x_95 = x_124; -x_96 = x_123; +x_96 = x_125; x_97 = x_136; goto block_120; } @@ -32760,9 +32760,9 @@ goto block_120; { lean_object* x_137; x_137 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__62; -x_94 = x_125; +x_94 = x_123; x_95 = x_124; -x_96 = x_123; +x_96 = x_125; x_97 = x_137; goto block_120; } @@ -38636,7 +38636,7 @@ goto block_138; { lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; x_98 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 0, x_94); lean_ctor_set(x_98, 1, x_97); x_99 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__11; x_100 = lean_alloc_ctor(3, 1, 0); @@ -38663,7 +38663,7 @@ x_110 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_110, 0, x_109); lean_ctor_set(x_110, 1, x_102); x_111 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_111, 0, x_94); +lean_ctor_set(x_111, 0, x_96); lean_ctor_set(x_111, 1, x_110); x_112 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_112, 0, x_93); @@ -38697,9 +38697,9 @@ case 0: { lean_object* x_126; x_126 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__18; -x_94 = x_123; +x_94 = x_125; x_95 = x_124; -x_96 = x_125; +x_96 = x_123; x_97 = x_126; goto block_120; } @@ -38707,9 +38707,9 @@ case 1: { lean_object* x_127; x_127 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__22; -x_94 = x_123; +x_94 = x_125; x_95 = x_124; -x_96 = x_125; +x_96 = x_123; x_97 = x_127; goto block_120; } @@ -38717,9 +38717,9 @@ case 2: { lean_object* x_128; x_128 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__26; -x_94 = x_123; +x_94 = x_125; x_95 = x_124; -x_96 = x_125; +x_96 = x_123; x_97 = x_128; goto block_120; } @@ -38727,9 +38727,9 @@ case 3: { lean_object* x_129; x_129 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__30; -x_94 = x_123; +x_94 = x_125; x_95 = x_124; -x_96 = x_125; +x_96 = x_123; x_97 = x_129; goto block_120; } @@ -38737,9 +38737,9 @@ case 4: { lean_object* x_130; x_130 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__34; -x_94 = x_123; +x_94 = x_125; x_95 = x_124; -x_96 = x_125; +x_96 = x_123; x_97 = x_130; goto block_120; } @@ -38747,9 +38747,9 @@ case 5: { lean_object* x_131; x_131 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__38; -x_94 = x_123; +x_94 = x_125; x_95 = x_124; -x_96 = x_125; +x_96 = x_123; x_97 = x_131; goto block_120; } @@ -38757,9 +38757,9 @@ case 6: { lean_object* x_132; x_132 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__42; -x_94 = x_123; +x_94 = x_125; x_95 = x_124; -x_96 = x_125; +x_96 = x_123; x_97 = x_132; goto block_120; } @@ -38767,9 +38767,9 @@ case 7: { lean_object* x_133; x_133 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__46; -x_94 = x_123; +x_94 = x_125; x_95 = x_124; -x_96 = x_125; +x_96 = x_123; x_97 = x_133; goto block_120; } @@ -38777,9 +38777,9 @@ case 8: { lean_object* x_134; x_134 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__50; -x_94 = x_123; +x_94 = x_125; x_95 = x_124; -x_96 = x_125; +x_96 = x_123; x_97 = x_134; goto block_120; } @@ -38787,9 +38787,9 @@ case 9: { lean_object* x_135; x_135 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__54; -x_94 = x_123; +x_94 = x_125; x_95 = x_124; -x_96 = x_125; +x_96 = x_123; x_97 = x_135; goto block_120; } @@ -38797,9 +38797,9 @@ case 10: { lean_object* x_136; x_136 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__58; -x_94 = x_123; +x_94 = x_125; x_95 = x_124; -x_96 = x_125; +x_96 = x_123; x_97 = x_136; goto block_120; } @@ -38807,9 +38807,9 @@ goto block_120; { lean_object* x_137; x_137 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__62; -x_94 = x_123; +x_94 = x_125; x_95 = x_124; -x_96 = x_125; +x_96 = x_123; x_97 = x_137; goto block_120; } @@ -42033,13 +42033,13 @@ lean_dec(x_91); x_107 = l_List_appendTR___redArg(x_104, x_106); x_108 = l_Lean_Json_mkObj(x_107); x_109 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_109, 0, x_96); +lean_ctor_set(x_109, 0, x_94); lean_ctor_set(x_109, 1, x_108); x_110 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_110, 0, x_109); lean_ctor_set(x_110, 1, x_102); x_111 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_111, 0, x_94); +lean_ctor_set(x_111, 0, x_96); lean_ctor_set(x_111, 1, x_110); x_112 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_112, 0, x_93); @@ -42073,9 +42073,9 @@ case 0: { lean_object* x_126; x_126 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__18; -x_94 = x_123; +x_94 = x_124; x_95 = x_125; -x_96 = x_124; +x_96 = x_123; x_97 = x_126; goto block_120; } @@ -42083,9 +42083,9 @@ case 1: { lean_object* x_127; x_127 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__22; -x_94 = x_123; +x_94 = x_124; x_95 = x_125; -x_96 = x_124; +x_96 = x_123; x_97 = x_127; goto block_120; } @@ -42093,9 +42093,9 @@ case 2: { lean_object* x_128; x_128 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__26; -x_94 = x_123; +x_94 = x_124; x_95 = x_125; -x_96 = x_124; +x_96 = x_123; x_97 = x_128; goto block_120; } @@ -42103,9 +42103,9 @@ case 3: { lean_object* x_129; x_129 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__30; -x_94 = x_123; +x_94 = x_124; x_95 = x_125; -x_96 = x_124; +x_96 = x_123; x_97 = x_129; goto block_120; } @@ -42113,9 +42113,9 @@ case 4: { lean_object* x_130; x_130 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__34; -x_94 = x_123; +x_94 = x_124; x_95 = x_125; -x_96 = x_124; +x_96 = x_123; x_97 = x_130; goto block_120; } @@ -42123,9 +42123,9 @@ case 5: { lean_object* x_131; x_131 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__38; -x_94 = x_123; +x_94 = x_124; x_95 = x_125; -x_96 = x_124; +x_96 = x_123; x_97 = x_131; goto block_120; } @@ -42133,9 +42133,9 @@ case 6: { lean_object* x_132; x_132 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__42; -x_94 = x_123; +x_94 = x_124; x_95 = x_125; -x_96 = x_124; +x_96 = x_123; x_97 = x_132; goto block_120; } @@ -42143,9 +42143,9 @@ case 7: { lean_object* x_133; x_133 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__46; -x_94 = x_123; +x_94 = x_124; x_95 = x_125; -x_96 = x_124; +x_96 = x_123; x_97 = x_133; goto block_120; } @@ -42153,9 +42153,9 @@ case 8: { lean_object* x_134; x_134 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__50; -x_94 = x_123; +x_94 = x_124; x_95 = x_125; -x_96 = x_124; +x_96 = x_123; x_97 = x_134; goto block_120; } @@ -42163,9 +42163,9 @@ case 9: { lean_object* x_135; x_135 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__54; -x_94 = x_123; +x_94 = x_124; x_95 = x_125; -x_96 = x_124; +x_96 = x_123; x_97 = x_135; goto block_120; } @@ -42173,9 +42173,9 @@ case 10: { lean_object* x_136; x_136 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__58; -x_94 = x_123; +x_94 = x_124; x_95 = x_125; -x_96 = x_124; +x_96 = x_123; x_97 = x_136; goto block_120; } @@ -42183,9 +42183,9 @@ goto block_120; { lean_object* x_137; x_137 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__62; -x_94 = x_123; +x_94 = x_124; x_95 = x_125; -x_96 = x_124; +x_96 = x_123; x_97 = x_137; goto block_120; } @@ -43466,7 +43466,7 @@ goto block_138; { lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; x_98 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 0, x_95); lean_ctor_set(x_98, 1, x_97); x_99 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__11; x_100 = lean_alloc_ctor(3, 1, 0); @@ -43487,13 +43487,13 @@ lean_dec(x_91); x_107 = l_List_appendTR___redArg(x_104, x_106); x_108 = l_Lean_Json_mkObj(x_107); x_109 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_109, 0, x_94); +lean_ctor_set(x_109, 0, x_96); lean_ctor_set(x_109, 1, x_108); x_110 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_110, 0, x_109); lean_ctor_set(x_110, 1, x_102); x_111 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_111, 0, x_95); +lean_ctor_set(x_111, 0, x_94); lean_ctor_set(x_111, 1, x_110); x_112 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_112, 0, x_93); @@ -43527,9 +43527,9 @@ case 0: { lean_object* x_126; x_126 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__18; -x_94 = x_124; -x_95 = x_123; -x_96 = x_125; +x_94 = x_123; +x_95 = x_125; +x_96 = x_124; x_97 = x_126; goto block_120; } @@ -43537,9 +43537,9 @@ case 1: { lean_object* x_127; x_127 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__22; -x_94 = x_124; -x_95 = x_123; -x_96 = x_125; +x_94 = x_123; +x_95 = x_125; +x_96 = x_124; x_97 = x_127; goto block_120; } @@ -43547,9 +43547,9 @@ case 2: { lean_object* x_128; x_128 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__26; -x_94 = x_124; -x_95 = x_123; -x_96 = x_125; +x_94 = x_123; +x_95 = x_125; +x_96 = x_124; x_97 = x_128; goto block_120; } @@ -43557,9 +43557,9 @@ case 3: { lean_object* x_129; x_129 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__30; -x_94 = x_124; -x_95 = x_123; -x_96 = x_125; +x_94 = x_123; +x_95 = x_125; +x_96 = x_124; x_97 = x_129; goto block_120; } @@ -43567,9 +43567,9 @@ case 4: { lean_object* x_130; x_130 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__34; -x_94 = x_124; -x_95 = x_123; -x_96 = x_125; +x_94 = x_123; +x_95 = x_125; +x_96 = x_124; x_97 = x_130; goto block_120; } @@ -43577,9 +43577,9 @@ case 5: { lean_object* x_131; x_131 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__38; -x_94 = x_124; -x_95 = x_123; -x_96 = x_125; +x_94 = x_123; +x_95 = x_125; +x_96 = x_124; x_97 = x_131; goto block_120; } @@ -43587,9 +43587,9 @@ case 6: { lean_object* x_132; x_132 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__42; -x_94 = x_124; -x_95 = x_123; -x_96 = x_125; +x_94 = x_123; +x_95 = x_125; +x_96 = x_124; x_97 = x_132; goto block_120; } @@ -43597,9 +43597,9 @@ case 7: { lean_object* x_133; x_133 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__46; -x_94 = x_124; -x_95 = x_123; -x_96 = x_125; +x_94 = x_123; +x_95 = x_125; +x_96 = x_124; x_97 = x_133; goto block_120; } @@ -43607,9 +43607,9 @@ case 8: { lean_object* x_134; x_134 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__50; -x_94 = x_124; -x_95 = x_123; -x_96 = x_125; +x_94 = x_123; +x_95 = x_125; +x_96 = x_124; x_97 = x_134; goto block_120; } @@ -43617,9 +43617,9 @@ case 9: { lean_object* x_135; x_135 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__54; -x_94 = x_124; -x_95 = x_123; -x_96 = x_125; +x_94 = x_123; +x_95 = x_125; +x_96 = x_124; x_97 = x_135; goto block_120; } @@ -43627,9 +43627,9 @@ case 10: { lean_object* x_136; x_136 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__58; -x_94 = x_124; -x_95 = x_123; -x_96 = x_125; +x_94 = x_123; +x_95 = x_125; +x_96 = x_124; x_97 = x_136; goto block_120; } @@ -43637,9 +43637,9 @@ goto block_120; { lean_object* x_137; x_137 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__62; -x_94 = x_124; -x_95 = x_123; -x_96 = x_125; +x_94 = x_123; +x_95 = x_125; +x_96 = x_124; x_97 = x_137; goto block_120; } @@ -44653,9 +44653,9 @@ x_52 = lean_nat_dec_lt(x_50, x_51); if (x_52 == 0) { lean_dec(x_36); -x_12 = x_43; +x_12 = x_44; x_13 = lean_box(0); -x_14 = x_44; +x_14 = x_43; x_15 = x_49; goto block_31; } @@ -44668,9 +44668,9 @@ if (x_53 == 0) if (x_52 == 0) { lean_dec(x_36); -x_12 = x_43; +x_12 = x_44; x_13 = lean_box(0); -x_14 = x_44; +x_14 = x_43; x_15 = x_49; goto block_31; } @@ -44681,9 +44681,9 @@ x_54 = 0; x_55 = lean_usize_of_nat(x_51); x_56 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Server_Test_Runner_processGoals_spec__5(x_36, x_54, x_55, x_49); lean_dec(x_36); -x_12 = x_43; +x_12 = x_44; x_13 = lean_box(0); -x_14 = x_44; +x_14 = x_43; x_15 = x_56; goto block_31; } @@ -44695,9 +44695,9 @@ x_57 = 0; x_58 = lean_usize_of_nat(x_51); x_59 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Server_Test_Runner_processGoals_spec__5(x_36, x_57, x_58, x_49); lean_dec(x_36); -x_12 = x_43; +x_12 = x_44; x_13 = lean_box(0); -x_14 = x_44; +x_14 = x_43; x_15 = x_59; goto block_31; } @@ -44764,7 +44764,7 @@ lean_object* x_16; size_t x_17; size_t x_18; lean_object* x_19; x_16 = lean_box(0); x_17 = lean_array_size(x_15); x_18 = 0; -x_19 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Server_Test_Runner_processGoals_spec__4(x_8, x_9, x_11, x_16, x_15, x_17, x_18, x_16, x_12, x_14); +x_19 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Server_Test_Runner_processGoals_spec__4(x_8, x_9, x_11, x_16, x_15, x_17, x_18, x_16, x_14, x_12); lean_dec_ref(x_15); if (lean_obj_tag(x_19) == 0) { @@ -45617,7 +45617,7 @@ goto block_138; { lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; x_98 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_98, 0, x_94); +lean_ctor_set(x_98, 0, x_96); lean_ctor_set(x_98, 1, x_97); x_99 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__11; x_100 = lean_alloc_ctor(3, 1, 0); @@ -45638,7 +45638,7 @@ lean_dec(x_91); x_107 = l_List_appendTR___redArg(x_104, x_106); x_108 = l_Lean_Json_mkObj(x_107); x_109 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_109, 0, x_96); +lean_ctor_set(x_109, 0, x_94); lean_ctor_set(x_109, 1, x_108); x_110 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_110, 0, x_109); @@ -45678,9 +45678,9 @@ case 0: { lean_object* x_126; x_126 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__18; -x_94 = x_125; +x_94 = x_124; x_95 = x_123; -x_96 = x_124; +x_96 = x_125; x_97 = x_126; goto block_120; } @@ -45688,9 +45688,9 @@ case 1: { lean_object* x_127; x_127 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__22; -x_94 = x_125; +x_94 = x_124; x_95 = x_123; -x_96 = x_124; +x_96 = x_125; x_97 = x_127; goto block_120; } @@ -45698,9 +45698,9 @@ case 2: { lean_object* x_128; x_128 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__26; -x_94 = x_125; +x_94 = x_124; x_95 = x_123; -x_96 = x_124; +x_96 = x_125; x_97 = x_128; goto block_120; } @@ -45708,9 +45708,9 @@ case 3: { lean_object* x_129; x_129 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__30; -x_94 = x_125; +x_94 = x_124; x_95 = x_123; -x_96 = x_124; +x_96 = x_125; x_97 = x_129; goto block_120; } @@ -45718,9 +45718,9 @@ case 4: { lean_object* x_130; x_130 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__34; -x_94 = x_125; +x_94 = x_124; x_95 = x_123; -x_96 = x_124; +x_96 = x_125; x_97 = x_130; goto block_120; } @@ -45728,9 +45728,9 @@ case 5: { lean_object* x_131; x_131 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__38; -x_94 = x_125; +x_94 = x_124; x_95 = x_123; -x_96 = x_124; +x_96 = x_125; x_97 = x_131; goto block_120; } @@ -45738,9 +45738,9 @@ case 6: { lean_object* x_132; x_132 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__42; -x_94 = x_125; +x_94 = x_124; x_95 = x_123; -x_96 = x_124; +x_96 = x_125; x_97 = x_132; goto block_120; } @@ -45748,9 +45748,9 @@ case 7: { lean_object* x_133; x_133 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__46; -x_94 = x_125; +x_94 = x_124; x_95 = x_123; -x_96 = x_124; +x_96 = x_125; x_97 = x_133; goto block_120; } @@ -45758,9 +45758,9 @@ case 8: { lean_object* x_134; x_134 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__50; -x_94 = x_125; +x_94 = x_124; x_95 = x_123; -x_96 = x_124; +x_96 = x_125; x_97 = x_134; goto block_120; } @@ -45768,9 +45768,9 @@ case 9: { lean_object* x_135; x_135 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__54; -x_94 = x_125; +x_94 = x_124; x_95 = x_123; -x_96 = x_124; +x_96 = x_125; x_97 = x_135; goto block_120; } @@ -45778,9 +45778,9 @@ case 10: { lean_object* x_136; x_136 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__58; -x_94 = x_125; +x_94 = x_124; x_95 = x_123; -x_96 = x_124; +x_96 = x_125; x_97 = x_136; goto block_120; } @@ -45788,9 +45788,9 @@ goto block_120; { lean_object* x_137; x_137 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__62; -x_94 = x_125; +x_94 = x_124; x_95 = x_123; -x_96 = x_124; +x_96 = x_125; x_97 = x_137; goto block_120; } @@ -47386,13 +47386,13 @@ lean_dec(x_91); x_107 = l_List_appendTR___redArg(x_104, x_106); x_108 = l_Lean_Json_mkObj(x_107); x_109 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_109, 0, x_95); +lean_ctor_set(x_109, 0, x_96); lean_ctor_set(x_109, 1, x_108); x_110 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_110, 0, x_109); lean_ctor_set(x_110, 1, x_102); x_111 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_111, 0, x_96); +lean_ctor_set(x_111, 0, x_95); lean_ctor_set(x_111, 1, x_110); x_112 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_112, 0, x_93); @@ -47427,8 +47427,8 @@ case 0: lean_object* x_126; x_126 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__18; x_94 = x_125; -x_95 = x_124; -x_96 = x_123; +x_95 = x_123; +x_96 = x_124; x_97 = x_126; goto block_120; } @@ -47437,8 +47437,8 @@ case 1: lean_object* x_127; x_127 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__22; x_94 = x_125; -x_95 = x_124; -x_96 = x_123; +x_95 = x_123; +x_96 = x_124; x_97 = x_127; goto block_120; } @@ -47447,8 +47447,8 @@ case 2: lean_object* x_128; x_128 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__26; x_94 = x_125; -x_95 = x_124; -x_96 = x_123; +x_95 = x_123; +x_96 = x_124; x_97 = x_128; goto block_120; } @@ -47457,8 +47457,8 @@ case 3: lean_object* x_129; x_129 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__30; x_94 = x_125; -x_95 = x_124; -x_96 = x_123; +x_95 = x_123; +x_96 = x_124; x_97 = x_129; goto block_120; } @@ -47467,8 +47467,8 @@ case 4: lean_object* x_130; x_130 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__34; x_94 = x_125; -x_95 = x_124; -x_96 = x_123; +x_95 = x_123; +x_96 = x_124; x_97 = x_130; goto block_120; } @@ -47477,8 +47477,8 @@ case 5: lean_object* x_131; x_131 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__38; x_94 = x_125; -x_95 = x_124; -x_96 = x_123; +x_95 = x_123; +x_96 = x_124; x_97 = x_131; goto block_120; } @@ -47487,8 +47487,8 @@ case 6: lean_object* x_132; x_132 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__42; x_94 = x_125; -x_95 = x_124; -x_96 = x_123; +x_95 = x_123; +x_96 = x_124; x_97 = x_132; goto block_120; } @@ -47497,8 +47497,8 @@ case 7: lean_object* x_133; x_133 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__46; x_94 = x_125; -x_95 = x_124; -x_96 = x_123; +x_95 = x_123; +x_96 = x_124; x_97 = x_133; goto block_120; } @@ -47507,8 +47507,8 @@ case 8: lean_object* x_134; x_134 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__50; x_94 = x_125; -x_95 = x_124; -x_96 = x_123; +x_95 = x_123; +x_96 = x_124; x_97 = x_134; goto block_120; } @@ -47517,8 +47517,8 @@ case 9: lean_object* x_135; x_135 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__54; x_94 = x_125; -x_95 = x_124; -x_96 = x_123; +x_95 = x_123; +x_96 = x_124; x_97 = x_135; goto block_120; } @@ -47527,8 +47527,8 @@ case 10: lean_object* x_136; x_136 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__58; x_94 = x_125; -x_95 = x_124; -x_96 = x_123; +x_95 = x_123; +x_96 = x_124; x_97 = x_136; goto block_120; } @@ -47537,8 +47537,8 @@ goto block_120; lean_object* x_137; x_137 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__62; x_94 = x_125; -x_95 = x_124; -x_96 = x_123; +x_95 = x_123; +x_96 = x_124; x_97 = x_137; goto block_120; } @@ -49230,7 +49230,7 @@ goto block_138; { lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; x_98 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_98, 0, x_94); +lean_ctor_set(x_98, 0, x_96); lean_ctor_set(x_98, 1, x_97); x_99 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__11; x_100 = lean_alloc_ctor(3, 1, 0); @@ -49257,7 +49257,7 @@ x_110 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_110, 0, x_109); lean_ctor_set(x_110, 1, x_102); x_111 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_111, 0, x_96); +lean_ctor_set(x_111, 0, x_94); lean_ctor_set(x_111, 1, x_110); x_112 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_112, 0, x_93); @@ -49291,9 +49291,9 @@ case 0: { lean_object* x_126; x_126 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__18; -x_94 = x_125; +x_94 = x_123; x_95 = x_124; -x_96 = x_123; +x_96 = x_125; x_97 = x_126; goto block_120; } @@ -49301,9 +49301,9 @@ case 1: { lean_object* x_127; x_127 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__22; -x_94 = x_125; +x_94 = x_123; x_95 = x_124; -x_96 = x_123; +x_96 = x_125; x_97 = x_127; goto block_120; } @@ -49311,9 +49311,9 @@ case 2: { lean_object* x_128; x_128 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__26; -x_94 = x_125; +x_94 = x_123; x_95 = x_124; -x_96 = x_123; +x_96 = x_125; x_97 = x_128; goto block_120; } @@ -49321,9 +49321,9 @@ case 3: { lean_object* x_129; x_129 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__30; -x_94 = x_125; +x_94 = x_123; x_95 = x_124; -x_96 = x_123; +x_96 = x_125; x_97 = x_129; goto block_120; } @@ -49331,9 +49331,9 @@ case 4: { lean_object* x_130; x_130 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__34; -x_94 = x_125; +x_94 = x_123; x_95 = x_124; -x_96 = x_123; +x_96 = x_125; x_97 = x_130; goto block_120; } @@ -49341,9 +49341,9 @@ case 5: { lean_object* x_131; x_131 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__38; -x_94 = x_125; +x_94 = x_123; x_95 = x_124; -x_96 = x_123; +x_96 = x_125; x_97 = x_131; goto block_120; } @@ -49351,9 +49351,9 @@ case 6: { lean_object* x_132; x_132 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__42; -x_94 = x_125; +x_94 = x_123; x_95 = x_124; -x_96 = x_123; +x_96 = x_125; x_97 = x_132; goto block_120; } @@ -49361,9 +49361,9 @@ case 7: { lean_object* x_133; x_133 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__46; -x_94 = x_125; +x_94 = x_123; x_95 = x_124; -x_96 = x_123; +x_96 = x_125; x_97 = x_133; goto block_120; } @@ -49371,9 +49371,9 @@ case 8: { lean_object* x_134; x_134 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__50; -x_94 = x_125; +x_94 = x_123; x_95 = x_124; -x_96 = x_123; +x_96 = x_125; x_97 = x_134; goto block_120; } @@ -49381,9 +49381,9 @@ case 9: { lean_object* x_135; x_135 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__54; -x_94 = x_125; +x_94 = x_123; x_95 = x_124; -x_96 = x_123; +x_96 = x_125; x_97 = x_135; goto block_120; } @@ -49391,9 +49391,9 @@ case 10: { lean_object* x_136; x_136 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__58; -x_94 = x_125; +x_94 = x_123; x_95 = x_124; -x_96 = x_123; +x_96 = x_125; x_97 = x_136; goto block_120; } @@ -49401,9 +49401,9 @@ goto block_120; { lean_object* x_137; x_137 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__62; -x_94 = x_125; +x_94 = x_123; x_95 = x_124; -x_96 = x_123; +x_96 = x_125; x_97 = x_137; goto block_120; } @@ -54014,108 +54014,108 @@ x_12 = lean_ctor_get(x_4, 5); lean_dec(x_12); x_13 = lean_ctor_get(x_4, 4); lean_dec(x_13); -lean_inc_ref(x_8); +lean_inc_ref(x_7); lean_ctor_set(x_4, 6, x_9); -lean_ctor_set(x_4, 5, x_8); -lean_ctor_set(x_4, 4, x_7); +lean_ctor_set(x_4, 5, x_7); +lean_ctor_set(x_4, 4, x_8); x_14 = l_Lean_Server_Test_Runner_processEdit___closed__6; -x_15 = lean_string_dec_eq(x_8, x_14); +x_15 = lean_string_dec_eq(x_7, x_14); if (x_15 == 0) { lean_object* x_16; uint8_t x_17; x_16 = l_Lean_Server_Test_Runner_processEdit___closed__7; -x_17 = lean_string_dec_eq(x_8, x_16); +x_17 = lean_string_dec_eq(x_7, x_16); if (x_17 == 0) { lean_object* x_18; uint8_t x_19; x_18 = l_Lean_Server_Test_Runner_processEdit___closed__8; -x_19 = lean_string_dec_eq(x_8, x_18); +x_19 = lean_string_dec_eq(x_7, x_18); if (x_19 == 0) { lean_object* x_20; uint8_t x_21; x_20 = l_Lean_Server_Test_Runner_processDirective___closed__0; -x_21 = lean_string_dec_eq(x_8, x_20); +x_21 = lean_string_dec_eq(x_7, x_20); if (x_21 == 0) { lean_object* x_22; uint8_t x_23; x_22 = l_Lean_Server_Test_Runner_processDirective___closed__1; -x_23 = lean_string_dec_eq(x_8, x_22); +x_23 = lean_string_dec_eq(x_7, x_22); if (x_23 == 0) { lean_object* x_24; uint8_t x_25; x_24 = l_Lean_Server_Test_Runner_processDirective___closed__2; -x_25 = lean_string_dec_eq(x_8, x_24); +x_25 = lean_string_dec_eq(x_7, x_24); if (x_25 == 0) { lean_object* x_26; uint8_t x_27; x_26 = l_Lean_Server_Test_Runner_processDirective___closed__3; -x_27 = lean_string_dec_eq(x_8, x_26); +x_27 = lean_string_dec_eq(x_7, x_26); if (x_27 == 0) { lean_object* x_28; uint8_t x_29; x_28 = l_Lean_Server_Test_Runner_processDirective___closed__4; -x_29 = lean_string_dec_eq(x_8, x_28); +x_29 = lean_string_dec_eq(x_7, x_28); if (x_29 == 0) { lean_object* x_30; uint8_t x_31; x_30 = l_Lean_Server_Test_Runner_processDirective___closed__5; -x_31 = lean_string_dec_eq(x_8, x_30); +x_31 = lean_string_dec_eq(x_7, x_30); if (x_31 == 0) { lean_object* x_32; uint8_t x_33; x_32 = l_Lean_Server_Test_Runner_Client_instFromJsonInteractiveGoals_fromJson___closed__0; -x_33 = lean_string_dec_eq(x_8, x_32); +x_33 = lean_string_dec_eq(x_7, x_32); if (x_33 == 0) { lean_object* x_34; uint8_t x_35; x_34 = l_Lean_Server_Test_Runner_processDirective___closed__6; -x_35 = lean_string_dec_eq(x_8, x_34); +x_35 = lean_string_dec_eq(x_7, x_34); if (x_35 == 0) { lean_object* x_36; uint8_t x_37; x_36 = l_Lean_Server_Test_Runner_instFromJsonGetWidgetsResponse_fromJson___closed__0; -x_37 = lean_string_dec_eq(x_8, x_36); +x_37 = lean_string_dec_eq(x_7, x_36); if (x_37 == 0) { lean_object* x_38; uint8_t x_39; x_38 = l_Lean_Server_Test_Runner_processDirective___closed__7; -x_39 = lean_string_dec_eq(x_8, x_38); +x_39 = lean_string_dec_eq(x_7, x_38); if (x_39 == 0) { lean_object* x_40; uint8_t x_41; x_40 = l_Lean_Server_Test_Runner_processDirective___closed__8; -x_41 = lean_string_dec_eq(x_8, x_40); +x_41 = lean_string_dec_eq(x_7, x_40); if (x_41 == 0) { lean_object* x_42; uint8_t x_43; x_42 = l_Lean_Server_Test_Runner_processDirective___closed__9; -x_43 = lean_string_dec_eq(x_8, x_42); +x_43 = lean_string_dec_eq(x_7, x_42); if (x_43 == 0) { lean_object* x_44; uint8_t x_45; x_44 = l_Lean_Server_Test_Runner_processDirective___closed__10; -x_45 = lean_string_dec_eq(x_8, x_44); +x_45 = lean_string_dec_eq(x_7, x_44); if (x_45 == 0) { lean_object* x_46; uint8_t x_47; x_46 = l_Lean_Server_Test_Runner_processDirective___closed__11; -x_47 = lean_string_dec_eq(x_8, x_46); +x_47 = lean_string_dec_eq(x_7, x_46); if (x_47 == 0) { lean_object* x_48; uint8_t x_49; x_48 = l_Lean_Server_Test_Runner_processDirective___closed__12; -x_49 = lean_string_dec_eq(x_8, x_48); +x_49 = lean_string_dec_eq(x_7, x_48); if (x_49 == 0) { lean_object* x_50; uint8_t x_51; x_50 = l_Lean_Server_Test_Runner_processDirective___closed__13; -x_51 = lean_string_dec_eq(x_8, x_50); +x_51 = lean_string_dec_eq(x_7, x_50); if (x_51 == 0) { lean_object* x_52; uint8_t x_53; x_52 = l_Lean_Server_Test_Runner_processDirective___closed__14; -x_53 = lean_string_dec_eq(x_8, x_52); -lean_dec_ref(x_8); +x_53 = lean_string_dec_eq(x_7, x_52); +lean_dec_ref(x_7); if (x_53 == 0) { lean_object* x_54; @@ -54132,7 +54132,7 @@ return x_55; else { lean_object* x_56; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_56 = l_Lean_Server_Test_Runner_processModuleHierarchyImportedBy(x_4, x_5); return x_56; } @@ -54140,7 +54140,7 @@ return x_56; else { lean_object* x_57; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_57 = l_Lean_Server_Test_Runner_processModuleHierarchyImports(x_4, x_5); return x_57; } @@ -54148,7 +54148,7 @@ return x_57; else { lean_object* x_58; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_58 = l_Lean_Server_Test_Runner_processSymbols(x_4, x_5); return x_58; } @@ -54156,7 +54156,7 @@ return x_58; else { lean_object* x_59; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_59 = l_Lean_Server_Test_Runner_processReferences(x_4, x_5); return x_59; } @@ -54164,7 +54164,7 @@ return x_59; else { lean_object* x_60; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_60 = l_Lean_Server_Test_Runner_processOutgoingCallHierarchy(x_4, x_5); return x_60; } @@ -54172,7 +54172,7 @@ return x_60; else { lean_object* x_61; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_61 = l_Lean_Server_Test_Runner_processIncomingCallHierarchy(x_4, x_5); return x_61; } @@ -54180,7 +54180,7 @@ return x_61; else { lean_object* x_62; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_62 = l_Lean_Server_Test_Runner_processCompletion(x_4, x_5); return x_62; } @@ -54188,7 +54188,7 @@ return x_62; else { lean_object* x_63; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_63 = l_Lean_Server_Test_Runner_processWidgets(x_4, x_5); return x_63; } @@ -54196,7 +54196,7 @@ return x_63; else { lean_object* x_64; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_64 = l_Lean_Server_Test_Runner_processTermGoal(x_4, x_5); return x_64; } @@ -54204,7 +54204,7 @@ return x_64; else { lean_object* x_65; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_65 = l_Lean_Server_Test_Runner_processGoals(x_4, x_5); return x_65; } @@ -54212,7 +54212,7 @@ return x_65; else { lean_object* x_66; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_66 = l_Lean_Server_Test_Runner_processInteractiveDiagnostics(x_4, x_5); return x_66; } @@ -54220,7 +54220,7 @@ return x_66; else { lean_object* x_67; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_67 = l_Lean_Server_Test_Runner_processCodeAction(x_4, x_5); return x_67; } @@ -54228,7 +54228,7 @@ return x_67; else { lean_object* x_68; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_68 = l_Lean_Server_Test_Runner_processWaitFor(x_4, x_5); return x_68; } @@ -54236,7 +54236,7 @@ return x_68; else { lean_object* x_69; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_69 = l_Lean_Server_Test_Runner_processSync(x_4, x_5); return x_69; } @@ -54244,7 +54244,7 @@ return x_69; else { lean_object* x_70; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_70 = l_Lean_Server_Test_Runner_processWaitForILeans(x_4, x_5); return x_70; } @@ -54252,7 +54252,7 @@ return x_70; else { lean_object* x_71; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_71 = l_Lean_Server_Test_Runner_processCollectDiagnostics(x_4, x_5); return x_71; } @@ -54260,7 +54260,7 @@ return x_71; else { lean_object* x_72; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_72 = l_Lean_Server_Test_Runner_processEdit(x_4, x_5); return x_72; } @@ -54268,7 +54268,7 @@ return x_72; else { lean_object* x_73; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_73 = l_Lean_Server_Test_Runner_processEdit(x_4, x_5); return x_73; } @@ -54276,7 +54276,7 @@ return x_73; else { lean_object* x_74; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_74 = l_Lean_Server_Test_Runner_processEdit(x_4, x_5); return x_74; } @@ -54298,116 +54298,116 @@ lean_inc(x_78); lean_inc(x_77); lean_inc(x_75); lean_dec(x_4); -lean_inc_ref(x_8); +lean_inc_ref(x_7); x_82 = lean_alloc_ctor(0, 9, 1); lean_ctor_set(x_82, 0, x_75); lean_ctor_set(x_82, 1, x_77); lean_ctor_set(x_82, 2, x_78); lean_ctor_set(x_82, 3, x_79); -lean_ctor_set(x_82, 4, x_7); -lean_ctor_set(x_82, 5, x_8); +lean_ctor_set(x_82, 4, x_8); +lean_ctor_set(x_82, 5, x_7); lean_ctor_set(x_82, 6, x_9); lean_ctor_set(x_82, 7, x_80); lean_ctor_set(x_82, 8, x_81); lean_ctor_set_uint8(x_82, sizeof(void*)*9, x_76); x_83 = l_Lean_Server_Test_Runner_processEdit___closed__6; -x_84 = lean_string_dec_eq(x_8, x_83); +x_84 = lean_string_dec_eq(x_7, x_83); if (x_84 == 0) { lean_object* x_85; uint8_t x_86; x_85 = l_Lean_Server_Test_Runner_processEdit___closed__7; -x_86 = lean_string_dec_eq(x_8, x_85); +x_86 = lean_string_dec_eq(x_7, x_85); if (x_86 == 0) { lean_object* x_87; uint8_t x_88; x_87 = l_Lean_Server_Test_Runner_processEdit___closed__8; -x_88 = lean_string_dec_eq(x_8, x_87); +x_88 = lean_string_dec_eq(x_7, x_87); if (x_88 == 0) { lean_object* x_89; uint8_t x_90; x_89 = l_Lean_Server_Test_Runner_processDirective___closed__0; -x_90 = lean_string_dec_eq(x_8, x_89); +x_90 = lean_string_dec_eq(x_7, x_89); if (x_90 == 0) { lean_object* x_91; uint8_t x_92; x_91 = l_Lean_Server_Test_Runner_processDirective___closed__1; -x_92 = lean_string_dec_eq(x_8, x_91); +x_92 = lean_string_dec_eq(x_7, x_91); if (x_92 == 0) { lean_object* x_93; uint8_t x_94; x_93 = l_Lean_Server_Test_Runner_processDirective___closed__2; -x_94 = lean_string_dec_eq(x_8, x_93); +x_94 = lean_string_dec_eq(x_7, x_93); if (x_94 == 0) { lean_object* x_95; uint8_t x_96; x_95 = l_Lean_Server_Test_Runner_processDirective___closed__3; -x_96 = lean_string_dec_eq(x_8, x_95); +x_96 = lean_string_dec_eq(x_7, x_95); if (x_96 == 0) { lean_object* x_97; uint8_t x_98; x_97 = l_Lean_Server_Test_Runner_processDirective___closed__4; -x_98 = lean_string_dec_eq(x_8, x_97); +x_98 = lean_string_dec_eq(x_7, x_97); if (x_98 == 0) { lean_object* x_99; uint8_t x_100; x_99 = l_Lean_Server_Test_Runner_processDirective___closed__5; -x_100 = lean_string_dec_eq(x_8, x_99); +x_100 = lean_string_dec_eq(x_7, x_99); if (x_100 == 0) { lean_object* x_101; uint8_t x_102; x_101 = l_Lean_Server_Test_Runner_Client_instFromJsonInteractiveGoals_fromJson___closed__0; -x_102 = lean_string_dec_eq(x_8, x_101); +x_102 = lean_string_dec_eq(x_7, x_101); if (x_102 == 0) { lean_object* x_103; uint8_t x_104; x_103 = l_Lean_Server_Test_Runner_processDirective___closed__6; -x_104 = lean_string_dec_eq(x_8, x_103); +x_104 = lean_string_dec_eq(x_7, x_103); if (x_104 == 0) { lean_object* x_105; uint8_t x_106; x_105 = l_Lean_Server_Test_Runner_instFromJsonGetWidgetsResponse_fromJson___closed__0; -x_106 = lean_string_dec_eq(x_8, x_105); +x_106 = lean_string_dec_eq(x_7, x_105); if (x_106 == 0) { lean_object* x_107; uint8_t x_108; x_107 = l_Lean_Server_Test_Runner_processDirective___closed__7; -x_108 = lean_string_dec_eq(x_8, x_107); +x_108 = lean_string_dec_eq(x_7, x_107); if (x_108 == 0) { lean_object* x_109; uint8_t x_110; x_109 = l_Lean_Server_Test_Runner_processDirective___closed__8; -x_110 = lean_string_dec_eq(x_8, x_109); +x_110 = lean_string_dec_eq(x_7, x_109); if (x_110 == 0) { lean_object* x_111; uint8_t x_112; x_111 = l_Lean_Server_Test_Runner_processDirective___closed__9; -x_112 = lean_string_dec_eq(x_8, x_111); +x_112 = lean_string_dec_eq(x_7, x_111); if (x_112 == 0) { lean_object* x_113; uint8_t x_114; x_113 = l_Lean_Server_Test_Runner_processDirective___closed__10; -x_114 = lean_string_dec_eq(x_8, x_113); +x_114 = lean_string_dec_eq(x_7, x_113); if (x_114 == 0) { lean_object* x_115; uint8_t x_116; x_115 = l_Lean_Server_Test_Runner_processDirective___closed__11; -x_116 = lean_string_dec_eq(x_8, x_115); +x_116 = lean_string_dec_eq(x_7, x_115); if (x_116 == 0) { lean_object* x_117; uint8_t x_118; x_117 = l_Lean_Server_Test_Runner_processDirective___closed__12; -x_118 = lean_string_dec_eq(x_8, x_117); +x_118 = lean_string_dec_eq(x_7, x_117); if (x_118 == 0) { lean_object* x_119; uint8_t x_120; x_119 = l_Lean_Server_Test_Runner_processDirective___closed__13; -x_120 = lean_string_dec_eq(x_8, x_119); +x_120 = lean_string_dec_eq(x_7, x_119); if (x_120 == 0) { lean_object* x_121; uint8_t x_122; x_121 = l_Lean_Server_Test_Runner_processDirective___closed__14; -x_122 = lean_string_dec_eq(x_8, x_121); -lean_dec_ref(x_8); +x_122 = lean_string_dec_eq(x_7, x_121); +lean_dec_ref(x_7); if (x_122 == 0) { lean_object* x_123; @@ -54424,7 +54424,7 @@ return x_124; else { lean_object* x_125; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_125 = l_Lean_Server_Test_Runner_processModuleHierarchyImportedBy(x_82, x_5); return x_125; } @@ -54432,7 +54432,7 @@ return x_125; else { lean_object* x_126; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_126 = l_Lean_Server_Test_Runner_processModuleHierarchyImports(x_82, x_5); return x_126; } @@ -54440,7 +54440,7 @@ return x_126; else { lean_object* x_127; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_127 = l_Lean_Server_Test_Runner_processSymbols(x_82, x_5); return x_127; } @@ -54448,7 +54448,7 @@ return x_127; else { lean_object* x_128; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_128 = l_Lean_Server_Test_Runner_processReferences(x_82, x_5); return x_128; } @@ -54456,7 +54456,7 @@ return x_128; else { lean_object* x_129; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_129 = l_Lean_Server_Test_Runner_processOutgoingCallHierarchy(x_82, x_5); return x_129; } @@ -54464,7 +54464,7 @@ return x_129; else { lean_object* x_130; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_130 = l_Lean_Server_Test_Runner_processIncomingCallHierarchy(x_82, x_5); return x_130; } @@ -54472,7 +54472,7 @@ return x_130; else { lean_object* x_131; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_131 = l_Lean_Server_Test_Runner_processCompletion(x_82, x_5); return x_131; } @@ -54480,7 +54480,7 @@ return x_131; else { lean_object* x_132; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_132 = l_Lean_Server_Test_Runner_processWidgets(x_82, x_5); return x_132; } @@ -54488,7 +54488,7 @@ return x_132; else { lean_object* x_133; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_133 = l_Lean_Server_Test_Runner_processTermGoal(x_82, x_5); return x_133; } @@ -54496,7 +54496,7 @@ return x_133; else { lean_object* x_134; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_134 = l_Lean_Server_Test_Runner_processGoals(x_82, x_5); return x_134; } @@ -54504,7 +54504,7 @@ return x_134; else { lean_object* x_135; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_135 = l_Lean_Server_Test_Runner_processInteractiveDiagnostics(x_82, x_5); return x_135; } @@ -54512,7 +54512,7 @@ return x_135; else { lean_object* x_136; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_136 = l_Lean_Server_Test_Runner_processCodeAction(x_82, x_5); return x_136; } @@ -54520,7 +54520,7 @@ return x_136; else { lean_object* x_137; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_137 = l_Lean_Server_Test_Runner_processWaitFor(x_82, x_5); return x_137; } @@ -54528,7 +54528,7 @@ return x_137; else { lean_object* x_138; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_138 = l_Lean_Server_Test_Runner_processSync(x_82, x_5); return x_138; } @@ -54536,7 +54536,7 @@ return x_138; else { lean_object* x_139; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_139 = l_Lean_Server_Test_Runner_processWaitForILeans(x_82, x_5); return x_139; } @@ -54544,7 +54544,7 @@ return x_139; else { lean_object* x_140; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_140 = l_Lean_Server_Test_Runner_processCollectDiagnostics(x_82, x_5); return x_140; } @@ -54552,7 +54552,7 @@ return x_140; else { lean_object* x_141; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_141 = l_Lean_Server_Test_Runner_processEdit(x_82, x_5); return x_141; } @@ -54560,7 +54560,7 @@ return x_141; else { lean_object* x_142; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_142 = l_Lean_Server_Test_Runner_processEdit(x_82, x_5); return x_142; } @@ -54568,7 +54568,7 @@ return x_142; else { lean_object* x_143; -lean_dec_ref(x_8); +lean_dec_ref(x_7); x_143 = l_Lean_Server_Test_Runner_processEdit(x_82, x_5); return x_143; } @@ -54633,8 +54633,8 @@ x_172 = lean_string_utf8_extract(x_169, x_170, x_171); lean_dec(x_171); lean_dec(x_170); lean_dec_ref(x_169); -x_7 = x_162; -x_8 = x_158; +x_7 = x_158; +x_8 = x_162; x_9 = x_172; goto block_144; } @@ -54646,8 +54646,8 @@ lean_dec(x_151); lean_dec(x_149); lean_dec_ref(x_2); x_173 = l_Lean_Server_Test_Runner_processDirective___closed__17; -x_7 = x_162; -x_8 = x_158; +x_7 = x_158; +x_8 = x_162; x_9 = x_173; goto block_144; } @@ -54702,8 +54702,8 @@ x_192 = lean_string_utf8_extract(x_189, x_190, x_191); lean_dec(x_191); lean_dec(x_190); lean_dec_ref(x_189); -x_7 = x_181; -x_8 = x_177; +x_7 = x_177; +x_8 = x_181; x_9 = x_192; goto block_144; } @@ -54714,8 +54714,8 @@ lean_dec(x_151); lean_dec(x_149); lean_dec_ref(x_2); x_193 = l_Lean_Server_Test_Runner_processDirective___closed__17; -x_7 = x_181; -x_8 = x_177; +x_7 = x_177; +x_8 = x_181; x_9 = x_193; goto block_144; } @@ -57804,7 +57804,7 @@ goto block_138; { lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; x_98 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 0, x_95); lean_ctor_set(x_98, 1, x_97); x_99 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__11; x_100 = lean_alloc_ctor(3, 1, 0); @@ -57825,13 +57825,13 @@ lean_dec(x_91); x_107 = l_List_appendTR___redArg(x_104, x_106); x_108 = l_Lean_Json_mkObj(x_107); x_109 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_109, 0, x_95); +lean_ctor_set(x_109, 0, x_94); lean_ctor_set(x_109, 1, x_108); x_110 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_110, 0, x_109); lean_ctor_set(x_110, 1, x_102); x_111 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_111, 0, x_94); +lean_ctor_set(x_111, 0, x_96); lean_ctor_set(x_111, 1, x_110); x_112 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_112, 0, x_93); @@ -57865,9 +57865,9 @@ case 0: { lean_object* x_126; x_126 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__18; -x_94 = x_123; -x_95 = x_124; -x_96 = x_125; +x_94 = x_124; +x_95 = x_125; +x_96 = x_123; x_97 = x_126; goto block_120; } @@ -57875,9 +57875,9 @@ case 1: { lean_object* x_127; x_127 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__22; -x_94 = x_123; -x_95 = x_124; -x_96 = x_125; +x_94 = x_124; +x_95 = x_125; +x_96 = x_123; x_97 = x_127; goto block_120; } @@ -57885,9 +57885,9 @@ case 2: { lean_object* x_128; x_128 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__26; -x_94 = x_123; -x_95 = x_124; -x_96 = x_125; +x_94 = x_124; +x_95 = x_125; +x_96 = x_123; x_97 = x_128; goto block_120; } @@ -57895,9 +57895,9 @@ case 3: { lean_object* x_129; x_129 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__30; -x_94 = x_123; -x_95 = x_124; -x_96 = x_125; +x_94 = x_124; +x_95 = x_125; +x_96 = x_123; x_97 = x_129; goto block_120; } @@ -57905,9 +57905,9 @@ case 4: { lean_object* x_130; x_130 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__34; -x_94 = x_123; -x_95 = x_124; -x_96 = x_125; +x_94 = x_124; +x_95 = x_125; +x_96 = x_123; x_97 = x_130; goto block_120; } @@ -57915,9 +57915,9 @@ case 5: { lean_object* x_131; x_131 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__38; -x_94 = x_123; -x_95 = x_124; -x_96 = x_125; +x_94 = x_124; +x_95 = x_125; +x_96 = x_123; x_97 = x_131; goto block_120; } @@ -57925,9 +57925,9 @@ case 6: { lean_object* x_132; x_132 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__42; -x_94 = x_123; -x_95 = x_124; -x_96 = x_125; +x_94 = x_124; +x_95 = x_125; +x_96 = x_123; x_97 = x_132; goto block_120; } @@ -57935,9 +57935,9 @@ case 7: { lean_object* x_133; x_133 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__46; -x_94 = x_123; -x_95 = x_124; -x_96 = x_125; +x_94 = x_124; +x_95 = x_125; +x_96 = x_123; x_97 = x_133; goto block_120; } @@ -57945,9 +57945,9 @@ case 8: { lean_object* x_134; x_134 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__50; -x_94 = x_123; -x_95 = x_124; -x_96 = x_125; +x_94 = x_124; +x_95 = x_125; +x_96 = x_123; x_97 = x_134; goto block_120; } @@ -57955,9 +57955,9 @@ case 9: { lean_object* x_135; x_135 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__54; -x_94 = x_123; -x_95 = x_124; -x_96 = x_125; +x_94 = x_124; +x_95 = x_125; +x_96 = x_123; x_97 = x_135; goto block_120; } @@ -57965,9 +57965,9 @@ case 10: { lean_object* x_136; x_136 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__58; -x_94 = x_123; -x_95 = x_124; -x_96 = x_125; +x_94 = x_124; +x_95 = x_125; +x_96 = x_123; x_97 = x_136; goto block_120; } @@ -57975,9 +57975,9 @@ goto block_120; { lean_object* x_137; x_137 = l_Lean_Lsp_Ipc_readResponseAs___at___00Lean_Server_Test_Runner_request___at___00Lean_Server_Test_Runner_rpcRequest___at___00Lean_Server_Test_Runner_expandTraces_spec__2_spec__3_spec__5___closed__62; -x_94 = x_123; -x_95 = x_124; -x_96 = x_125; +x_94 = x_124; +x_95 = x_125; +x_96 = x_123; x_97 = x_137; goto block_120; } @@ -58468,8 +58468,8 @@ lean_closure_set(x_28, 2, x_9); lean_closure_set(x_28, 3, x_4); lean_closure_set(x_28, 4, x_11); lean_closure_set(x_28, 5, x_27); -lean_closure_set(x_28, 6, x_7); -x_29 = l_Lean_Lsp_Ipc_runWith___redArg(x_6, x_5, x_28); +lean_closure_set(x_28, 6, x_6); +x_29 = l_Lean_Lsp_Ipc_runWith___redArg(x_7, x_5, x_28); return x_29; } block_41: @@ -58486,8 +58486,8 @@ x_37 = l_Lean_Server_Test_Runner_word___closed__0; x_38 = lean_array_get(x_37, x_3, x_34); lean_dec_ref(x_3); x_5 = x_32; -x_6 = x_31; -x_7 = x_34; +x_6 = x_34; +x_7 = x_31; x_8 = x_36; x_9 = x_38; goto block_30; @@ -58499,8 +58499,8 @@ x_39 = l_Lean_Server_Test_Runner_word___closed__0; x_40 = lean_array_get(x_39, x_3, x_4); lean_dec_ref(x_3); x_5 = x_32; -x_6 = x_31; -x_7 = x_34; +x_6 = x_34; +x_7 = x_31; x_8 = x_36; x_9 = x_40; goto block_30; diff --git a/stage0/stdlib/Lean/Widget/InteractiveDiagnostic.c b/stage0/stdlib/Lean/Widget/InteractiveDiagnostic.c index ad756757d7f2..1fdf5b564b82 100644 --- a/stage0/stdlib/Lean/Widget/InteractiveDiagnostic.c +++ b/stage0/stdlib/Lean/Widget/InteractiveDiagnostic.c @@ -12433,7 +12433,7 @@ return x_2; LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_32; uint8_t x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_63; lean_object* x_64; uint8_t x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_76; lean_object* x_77; uint8_t x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_85; lean_object* x_86; lean_object* x_87; uint8_t x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_101; lean_object* x_102; lean_object* x_103; uint8_t x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; uint8_t x_119; lean_object* x_120; lean_object* x_121; uint8_t x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_129; uint8_t x_130; double x_131; lean_object* x_132; lean_object* x_133; double x_134; uint8_t x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; uint8_t x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_85; uint8_t x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_101; uint8_t x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_119; uint8_t x_120; lean_object* x_121; lean_object* x_122; uint8_t x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; double x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; double x_137; uint8_t x_138; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; x_317 = l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_go___closed__6; if (lean_obj_tag(x_2) == 0) { @@ -12617,10 +12617,10 @@ x_366 = lean_ctor_get(x_3, 0); lean_inc_ref(x_366); lean_dec_ref(x_3); x_367 = lean_box(0); -x_318 = x_4; +x_318 = lean_box(0); x_319 = x_2; -x_320 = lean_box(0); -x_321 = x_366; +x_320 = x_366; +x_321 = x_4; x_322 = x_367; goto block_329; } @@ -13053,10 +13053,10 @@ lean_dec_ref(x_3); x_456 = l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_mkPPContext(x_1, x_454); x_457 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_457, 0, x_456); -x_318 = x_4; +x_318 = lean_box(0); x_319 = x_2; -x_320 = lean_box(0); -x_321 = x_455; +x_320 = x_455; +x_321 = x_4; x_322 = x_457; goto block_329; } @@ -13149,10 +13149,10 @@ return x_11; { lean_object* x_38; lean_object* x_39; x_38 = lean_alloc_ctor(3, 3, 1); -lean_ctor_set(x_38, 0, x_32); +lean_ctor_set(x_38, 0, x_33); lean_ctor_set(x_38, 1, x_34); lean_ctor_set(x_38, 2, x_35); -lean_ctor_set_uint8(x_38, sizeof(void*)*3, x_33); +lean_ctor_set_uint8(x_38, sizeof(void*)*3, x_32); x_39 = l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_pushEmbed(x_38, x_36); if (lean_obj_tag(x_39) == 0) { @@ -13251,18 +13251,18 @@ return x_61; { lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; x_70 = lean_unsigned_to_nat(0u); -x_71 = lean_array_get_size(x_67); -x_72 = l_Array_toSubarray___redArg(x_67, x_70, x_71); -lean_inc(x_64); -x_73 = l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_chopUpChildren(x_64, x_69, x_72); +x_71 = lean_array_get_size(x_68); +x_72 = l_Array_toSubarray___redArg(x_68, x_70, x_71); +lean_inc(x_65); +x_73 = l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_chopUpChildren(x_65, x_69, x_72); lean_dec(x_69); x_74 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_74, 0, x_73); -x_32 = x_64; +x_32 = x_63; x_33 = x_65; -x_34 = x_68; +x_34 = x_67; x_35 = x_74; -x_36 = x_63; +x_36 = x_64; x_37 = lean_box(0); goto block_62; } @@ -13284,25 +13284,25 @@ goto block_75; block_100: { size_t x_92; size_t x_93; lean_object* x_94; -x_92 = lean_array_size(x_89); +x_92 = lean_array_size(x_88); x_93 = 0; -x_94 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_go_spec__1(x_86, x_1, x_92, x_93, x_89); -if (lean_obj_tag(x_86) == 0) +x_94 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_go_spec__1(x_85, x_1, x_92, x_93, x_88); +if (lean_obj_tag(x_85) == 0) { -x_76 = x_85; +x_76 = x_86; x_77 = x_87; -x_78 = x_88; +x_78 = x_89; x_79 = lean_box(0); -x_80 = x_94; -x_81 = x_91; +x_80 = x_91; +x_81 = x_94; goto block_84; } else { lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; -x_95 = lean_ctor_get(x_86, 0); +x_95 = lean_ctor_get(x_85, 0); lean_inc(x_95); -lean_dec_ref(x_86); +lean_dec_ref(x_85); x_96 = lean_ctor_get(x_95, 3); lean_inc(x_96); lean_dec(x_95); @@ -13311,12 +13311,12 @@ x_98 = l_Lean_Option_get_x3f___at___00__private_Lean_Widget_InteractiveDiagnosti lean_dec(x_96); if (lean_obj_tag(x_98) == 0) { -x_76 = x_85; +x_76 = x_86; x_77 = x_87; -x_78 = x_88; +x_78 = x_89; x_79 = lean_box(0); -x_80 = x_94; -x_81 = x_91; +x_80 = x_91; +x_81 = x_94; goto block_84; } else @@ -13325,12 +13325,12 @@ lean_object* x_99; x_99 = lean_ctor_get(x_98, 0); lean_inc(x_99); lean_dec_ref(x_98); -x_63 = x_85; +x_63 = x_86; x_64 = x_87; -x_65 = x_88; +x_65 = x_89; x_66 = lean_box(0); -x_67 = x_94; -x_68 = x_91; +x_67 = x_91; +x_68 = x_94; x_69 = x_99; goto block_75; } @@ -13339,9 +13339,9 @@ goto block_75; block_118: { size_t x_108; size_t x_109; lean_object* x_110; -x_108 = lean_array_size(x_105); +x_108 = lean_array_size(x_104); x_109 = 0; -x_110 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_go_spec__0(x_1, x_102, x_108, x_109, x_105, x_101); +x_110 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_go_spec__0(x_1, x_101, x_108, x_109, x_104, x_103); if (lean_obj_tag(x_110) == 0) { lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; @@ -13355,8 +13355,8 @@ lean_inc(x_113); lean_dec(x_111); x_114 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_114, 0, x_112); -x_32 = x_103; -x_33 = x_104; +x_32 = x_102; +x_33 = x_105; x_34 = x_107; x_35 = x_114; x_36 = x_113; @@ -13367,7 +13367,7 @@ else { uint8_t x_115; lean_dec(x_107); -lean_dec(x_103); +lean_dec(x_105); x_115 = !lean_is_exclusive(x_110); if (x_115 == 0) { @@ -13387,13 +13387,13 @@ return x_117; } block_128: { -if (x_122 == 0) +if (x_120 == 0) { -x_101 = x_125; +x_101 = x_119; x_102 = x_120; -x_103 = x_121; -x_104 = x_122; -x_105 = x_123; +x_103 = x_125; +x_104 = x_121; +x_105 = x_122; x_106 = lean_box(0); x_107 = x_124; goto block_118; @@ -13401,38 +13401,38 @@ goto block_118; else { uint8_t x_127; -x_127 = l_Array_isEmpty___redArg(x_123); +x_127 = l_Array_isEmpty___redArg(x_121); if (x_127 == 0) { -x_85 = x_125; +x_85 = x_119; x_86 = x_120; -x_87 = x_121; -x_88 = x_122; -x_89 = x_123; +x_87 = x_125; +x_88 = x_121; +x_89 = x_122; x_90 = lean_box(0); x_91 = x_124; goto block_100; } else { -if (x_119 == 0) +if (x_123 == 0) { -x_101 = x_125; +x_101 = x_119; x_102 = x_120; -x_103 = x_121; -x_104 = x_122; -x_105 = x_123; +x_103 = x_125; +x_104 = x_121; +x_105 = x_122; x_106 = lean_box(0); x_107 = x_124; goto block_118; } else { -x_85 = x_125; +x_85 = x_119; x_86 = x_120; -x_87 = x_121; -x_88 = x_122; -x_89 = x_123; +x_87 = x_125; +x_88 = x_121; +x_89 = x_122; x_90 = lean_box(0); x_91 = x_124; goto block_100; @@ -13444,7 +13444,7 @@ goto block_100; { lean_object* x_139; double x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; x_139 = l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_go___closed__2; -x_140 = lean_float_sub(x_134, x_131); +x_140 = lean_float_sub(x_137, x_133); x_141 = lean_float_to_string(x_140); x_142 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_142, 0, x_141); @@ -13457,14 +13457,14 @@ lean_ctor_set(x_145, 0, x_143); lean_ctor_set(x_145, 1, x_144); x_146 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_146, 0, x_145); -lean_ctor_set(x_146, 1, x_137); -x_119 = x_130; -x_120 = x_129; -x_121 = x_132; +lean_ctor_set(x_146, 1, x_136); +x_119 = x_129; +x_120 = x_132; +x_121 = x_134; x_122 = x_135; -x_123 = x_136; +x_123 = x_138; x_124 = x_146; -x_125 = x_138; +x_125 = x_131; x_126 = lean_box(0); goto block_128; } @@ -13505,26 +13505,26 @@ x_166 = lean_float_beq(x_155, x_165); if (x_166 == 0) { x_129 = x_148; -x_130 = x_158; -x_131 = x_155; -x_132 = x_154; -x_133 = lean_box(0); -x_134 = x_156; -x_135 = x_157; -x_136 = x_151; -x_137 = x_160; -x_138 = x_163; +x_130 = lean_box(0); +x_131 = x_163; +x_132 = x_157; +x_133 = x_155; +x_134 = x_151; +x_135 = x_154; +x_136 = x_160; +x_137 = x_156; +x_138 = x_158; goto block_147; } else { if (x_158 == 0) { -x_119 = x_158; -x_120 = x_148; -x_121 = x_154; -x_122 = x_157; -x_123 = x_151; +x_119 = x_148; +x_120 = x_157; +x_121 = x_151; +x_122 = x_154; +x_123 = x_158; x_124 = x_160; x_125 = x_163; x_126 = lean_box(0); @@ -13533,15 +13533,15 @@ goto block_128; else { x_129 = x_148; -x_130 = x_158; -x_131 = x_155; -x_132 = x_154; -x_133 = lean_box(0); -x_134 = x_156; -x_135 = x_157; -x_136 = x_151; -x_137 = x_160; -x_138 = x_163; +x_130 = lean_box(0); +x_131 = x_163; +x_132 = x_157; +x_133 = x_155; +x_134 = x_151; +x_135 = x_154; +x_136 = x_160; +x_137 = x_156; +x_138 = x_158; goto block_147; } } @@ -13563,26 +13563,26 @@ x_172 = lean_float_beq(x_155, x_171); if (x_172 == 0) { x_129 = x_148; -x_130 = x_158; -x_131 = x_155; -x_132 = x_154; -x_133 = lean_box(0); -x_134 = x_156; -x_135 = x_157; -x_136 = x_151; -x_137 = x_170; -x_138 = x_168; +x_130 = lean_box(0); +x_131 = x_168; +x_132 = x_157; +x_133 = x_155; +x_134 = x_151; +x_135 = x_154; +x_136 = x_170; +x_137 = x_156; +x_138 = x_158; goto block_147; } else { if (x_158 == 0) { -x_119 = x_158; -x_120 = x_148; -x_121 = x_154; -x_122 = x_157; -x_123 = x_151; +x_119 = x_148; +x_120 = x_157; +x_121 = x_151; +x_122 = x_154; +x_123 = x_158; x_124 = x_170; x_125 = x_168; x_126 = lean_box(0); @@ -13591,15 +13591,15 @@ goto block_128; else { x_129 = x_148; -x_130 = x_158; -x_131 = x_155; -x_132 = x_154; -x_133 = lean_box(0); -x_134 = x_156; -x_135 = x_157; -x_136 = x_151; -x_137 = x_170; -x_138 = x_168; +x_130 = lean_box(0); +x_131 = x_168; +x_132 = x_157; +x_133 = x_155; +x_134 = x_151; +x_135 = x_154; +x_136 = x_170; +x_137 = x_156; +x_138 = x_158; goto block_147; } } @@ -14160,7 +14160,7 @@ return x_296; block_329: { lean_object* x_323; lean_object* x_324; -x_323 = lean_apply_2(x_321, x_322, lean_box(0)); +x_323 = lean_apply_2(x_320, x_322, lean_box(0)); x_324 = l___private_Init_Dynamic_0__Dynamic_get_x3fImpl___redArg(x_323, x_317); lean_dec(x_323); if (lean_obj_tag(x_324) == 1) @@ -14171,15 +14171,15 @@ lean_inc(x_325); lean_dec_ref(x_324); x_2 = x_319; x_3 = x_325; -x_4 = x_318; +x_4 = x_321; goto _start; } else { lean_object* x_327; lean_object* x_328; lean_dec(x_324); +lean_dec_ref(x_321); lean_dec(x_319); -lean_dec_ref(x_318); lean_dec_ref(x_1); x_327 = l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_go___closed__8; x_328 = lean_alloc_ctor(1, 1, 0); @@ -15648,7 +15648,7 @@ return x_3; LEAN_EXPORT lean_object* l_Lean_Widget_msgToInteractiveDiagnostic(lean_object* x_1, lean_object* x_2, uint8_t x_3) { _start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_19; lean_object* x_20; uint8_t x_21; uint8_t x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; lean_object* x_53; lean_object* x_54; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; lean_object* x_80; lean_object* x_81; lean_object* x_88; lean_object* x_89; lean_object* x_90; uint8_t x_91; lean_object* x_92; lean_object* x_100; lean_object* x_101; uint8_t x_102; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_118; lean_object* x_119; lean_object* x_122; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_19; lean_object* x_20; uint8_t x_21; uint8_t x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_47; lean_object* x_48; uint8_t x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; lean_object* x_80; lean_object* x_81; lean_object* x_88; uint8_t x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_100; lean_object* x_101; uint8_t x_102; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_118; lean_object* x_119; lean_object* x_122; x_19 = lean_ctor_get(x_2, 1); lean_inc_ref(x_19); x_20 = lean_ctor_get(x_2, 2); @@ -15682,18 +15682,18 @@ goto block_141; { lean_object* x_15; lean_object* x_16; lean_object* x_17; x_15 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_15, 0, x_11); +lean_ctor_set(x_15, 0, x_10); x_16 = lean_box(0); x_17 = lean_alloc_ctor(0, 11, 0); -lean_ctor_set(x_17, 0, x_8); +lean_ctor_set(x_17, 0, x_5); lean_ctor_set(x_17, 1, x_15); -lean_ctor_set(x_17, 2, x_5); -lean_ctor_set(x_17, 3, x_13); +lean_ctor_set(x_17, 2, x_9); +lean_ctor_set(x_17, 3, x_12); lean_ctor_set(x_17, 4, x_14); -lean_ctor_set(x_17, 5, x_6); -lean_ctor_set(x_17, 6, x_9); +lean_ctor_set(x_17, 5, x_8); +lean_ctor_set(x_17, 6, x_6); lean_ctor_set(x_17, 7, x_7); -lean_ctor_set(x_17, 8, x_12); +lean_ctor_set(x_17, 8, x_11); lean_ctor_set(x_17, 9, x_16); lean_ctor_set(x_17, 10, x_16); return x_17; @@ -15710,14 +15710,14 @@ if (lean_obj_tag(x_36) == 0) lean_object* x_37; x_37 = lean_box(0); x_5 = x_25; -x_6 = x_26; -x_7 = x_28; -x_8 = x_27; -x_9 = x_34; -x_10 = lean_box(0); -x_11 = x_29; -x_12 = x_33; -x_13 = x_32; +x_6 = x_34; +x_7 = x_26; +x_8 = x_30; +x_9 = x_29; +x_10 = x_28; +x_11 = x_31; +x_12 = x_32; +x_13 = lean_box(0); x_14 = x_37; goto block_18; } @@ -15729,19 +15729,19 @@ if (x_38 == 0) { lean_object* x_39; lean_object* x_40; lean_object* x_41; x_39 = lean_ctor_get(x_36, 0); -x_40 = l_Lean_Name_toString(x_39, x_31); +x_40 = l_Lean_Name_toString(x_39, x_27); x_41 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_41, 0, x_40); lean_ctor_set(x_36, 0, x_41); x_5 = x_25; -x_6 = x_26; -x_7 = x_28; -x_8 = x_27; -x_9 = x_34; -x_10 = lean_box(0); -x_11 = x_29; -x_12 = x_33; -x_13 = x_32; +x_6 = x_34; +x_7 = x_26; +x_8 = x_30; +x_9 = x_29; +x_10 = x_28; +x_11 = x_31; +x_12 = x_32; +x_13 = lean_box(0); x_14 = x_36; goto block_18; } @@ -15751,20 +15751,20 @@ lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; x_42 = lean_ctor_get(x_36, 0); lean_inc(x_42); lean_dec(x_36); -x_43 = l_Lean_Name_toString(x_42, x_31); +x_43 = l_Lean_Name_toString(x_42, x_27); x_44 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_44, 0, x_43); x_45 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_45, 0, x_44); x_5 = x_25; -x_6 = x_26; -x_7 = x_28; -x_8 = x_27; -x_9 = x_34; -x_10 = lean_box(0); -x_11 = x_29; -x_12 = x_33; -x_13 = x_32; +x_6 = x_34; +x_7 = x_26; +x_8 = x_30; +x_9 = x_29; +x_10 = x_28; +x_11 = x_31; +x_12 = x_32; +x_13 = lean_box(0); x_14 = x_45; goto block_18; } @@ -15787,10 +15787,10 @@ x_26 = x_48; x_27 = x_49; x_28 = x_50; x_29 = x_51; -x_30 = lean_box(0); -x_31 = x_52; +x_30 = x_52; +x_31 = x_54; x_32 = x_53; -x_33 = x_54; +x_33 = lean_box(0); x_34 = x_57; goto block_46; } @@ -15815,10 +15815,10 @@ x_26 = x_48; x_27 = x_49; x_28 = x_50; x_29 = x_51; -x_30 = lean_box(0); -x_31 = x_52; +x_30 = x_52; +x_31 = x_54; x_32 = x_53; -x_33 = x_54; +x_33 = lean_box(0); x_34 = x_56; goto block_46; } @@ -15841,10 +15841,10 @@ x_26 = x_48; x_27 = x_49; x_28 = x_50; x_29 = x_51; -x_30 = lean_box(0); -x_31 = x_52; +x_30 = x_52; +x_31 = x_54; x_32 = x_53; -x_33 = x_54; +x_33 = lean_box(0); x_34 = x_71; goto block_46; } @@ -15865,11 +15865,11 @@ if (x_83 == 0) lean_object* x_84; x_84 = lean_box(0); x_47 = x_75; -x_48 = x_76; -x_49 = x_77; -x_50 = x_81; -x_51 = x_78; -x_52 = x_79; +x_48 = x_81; +x_49 = x_79; +x_50 = x_78; +x_51 = x_77; +x_52 = x_76; x_53 = x_80; x_54 = x_84; goto block_72; @@ -15879,11 +15879,11 @@ else lean_object* x_85; x_85 = l_Lean_Widget_msgToInteractiveDiagnostic___closed__4; x_47 = x_75; -x_48 = x_76; -x_49 = x_77; -x_50 = x_81; -x_51 = x_78; -x_52 = x_79; +x_48 = x_81; +x_49 = x_79; +x_50 = x_78; +x_51 = x_77; +x_52 = x_76; x_53 = x_80; x_54 = x_85; goto block_72; @@ -15895,11 +15895,11 @@ lean_object* x_86; lean_dec_ref(x_73); x_86 = l_Lean_Widget_msgToInteractiveDiagnostic___closed__6; x_47 = x_75; -x_48 = x_76; -x_49 = x_77; -x_50 = x_81; -x_51 = x_78; -x_52 = x_79; +x_48 = x_81; +x_49 = x_79; +x_50 = x_78; +x_51 = x_77; +x_52 = x_76; x_53 = x_80; x_54 = x_86; goto block_72; @@ -15922,9 +15922,9 @@ lean_object* x_96; x_96 = lean_box(0); x_75 = x_88; x_76 = x_93; -x_77 = x_89; +x_77 = x_91; x_78 = x_90; -x_79 = x_91; +x_79 = x_89; x_80 = x_92; x_81 = x_96; goto block_87; @@ -15935,9 +15935,9 @@ lean_object* x_97; x_97 = l_Lean_Widget_msgToInteractiveDiagnostic___closed__11; x_75 = x_88; x_76 = x_93; -x_77 = x_89; +x_77 = x_91; x_78 = x_90; -x_79 = x_91; +x_79 = x_89; x_80 = x_92; x_81 = x_97; goto block_87; @@ -15949,9 +15949,9 @@ lean_object* x_98; x_98 = l_Lean_Widget_msgToInteractiveDiagnostic___closed__13; x_75 = x_88; x_76 = x_93; -x_77 = x_89; +x_77 = x_91; x_78 = x_90; -x_79 = x_91; +x_79 = x_89; x_80 = x_92; x_81 = x_98; goto block_87; @@ -15968,10 +15968,10 @@ if (x_23 == 0) { lean_object* x_106; x_106 = lean_box(0); -x_88 = x_104; -x_89 = x_100; +x_88 = x_100; +x_89 = x_105; x_90 = x_101; -x_91 = x_105; +x_91 = x_104; x_92 = x_106; goto block_99; } @@ -15979,10 +15979,10 @@ else { lean_object* x_107; x_107 = l_Lean_Widget_msgToInteractiveDiagnostic___closed__14; -x_88 = x_104; -x_89 = x_100; +x_88 = x_100; +x_89 = x_105; x_90 = x_101; -x_91 = x_105; +x_91 = x_104; x_92 = x_107; goto block_99; } diff --git a/tests/bench/big_do.lean b/tests/bench/big_do.lean index 41319f1aaa66..c3c3f75b0680 100644 --- a/tests/bench/big_do.lean +++ b/tests/bench/big_do.lean @@ -425,34 +425,6 @@ def addALot (x: Nat) : StateM Nat Nat := do modifyGet (λ y => ((), y + x)) modifyGet (λ y => ((), y + x)) modifyGet (λ y => ((), y + x)) - modifyGet (λ y => ((), y + x)) - modifyGet (λ y => ((), y + x)) - modifyGet (λ y => ((), y + x)) - modifyGet (λ y => ((), y + x)) - modifyGet (λ y => ((), y + x)) - modifyGet (λ y => ((), y + x)) - modifyGet (λ y => ((), y + x)) - modifyGet (λ y => ((), y + x)) - modifyGet (λ y => ((), y + x)) - modifyGet (λ y => ((), y + x)) - modifyGet (λ y => ((), y + x)) - modifyGet (λ y => ((), y + x)) - modifyGet (λ y => ((), y + x)) - modifyGet (λ y => ((), y + x)) - modifyGet (λ y => ((), y + x)) - modifyGet (λ y => ((), y + x)) - modifyGet (λ y => ((), y + x)) - modifyGet (λ y => ((), y + x)) - modifyGet (λ y => ((), y + x)) - modifyGet (λ y => ((), y + x)) - modifyGet (λ y => ((), y + x)) - modifyGet (λ y => ((), y + x)) - modifyGet (λ y => ((), y + x)) - modifyGet (λ y => ((), y + x)) - modifyGet (λ y => ((), y + x)) - modifyGet (λ y => ((), y + x)) - modifyGet (λ y => ((), y + x)) - modifyGet (λ y => ((), y + x)) let y <- get pure y diff --git a/tests/compiler/unboxing.lean b/tests/compiler/unboxing.lean new file mode 100644 index 000000000000..1c0340086960 --- /dev/null +++ b/tests/compiler/unboxing.lean @@ -0,0 +1,54 @@ +@[noinline] +def discardIO (a : α) : BaseIO Unit := + go (a, pure ()) +where + go (x : α × BaseIO Unit) : BaseIO Unit := x.2 + +def myStructConst : UInt8 × UInt16 × UInt32 × UInt64 × USize × Float × Float32 × Nat := + (1, 2, 3, 4, 5, 6, 7, 8) + +@[unbox] +inductive SpecialValue (α : Type) where + | thing (a b : Nat) (c : UInt8) (d : UInt32) (e : USize) + | nothing + | something (x : α) +deriving Repr + +@[noinline] def special1 (_ : Unit) : SpecialValue Nat := .thing 1 2 3 4 5 +@[noinline] def special2 (_ : Unit) : SpecialValue (Nat × Nat) := .nothing +@[noinline] def special3 (_ : Unit) : SpecialValue (Nat × Nat) := .something (2, 3) +@[noinline] def special4 (_ : Unit) : SpecialValue Prop := .something True + +instance : Repr Prop where + reprPrec _ _ := "_" + + +@[noinline] +def myFunc (x : α × (Nat × Nat)) : Nat := x.2.1 + x.2.2 + +@[noinline] +def myOtherFunc (x : α) : (Nat × Nat) × α := ((1, 2), x) + +@[noinline] +def reshapeTest (x : Nat × Nat) : Nat := + myFunc (myOtherFunc x) + +set_option compiler.extract_closed false in +def main : IO Unit := do + -- Access a persistent builtin `struct` declaration + let a := stdRange + discardIO a.1 + -- Ensure that struct constructors with scalar values work + IO.println (repr myStructConst) + -- Construct different values of the `SpecialValue` type + IO.println (repr (special1 ())) + IO.println (repr (special2 ())) + IO.println (repr (special3 ())) + IO.println (repr (special4 ())) + -- box them + IO.println (repr [special1 ()]) + IO.println (repr [special2 ()]) + IO.println (repr [special3 ()]) + IO.println (repr [special4 ()]) + -- try a reshape + IO.println (reshapeTest (1, 2)) diff --git a/tests/compiler/unboxing.lean.expected.out b/tests/compiler/unboxing.lean.expected.out new file mode 100644 index 000000000000..b69efb1d3bf8 --- /dev/null +++ b/tests/compiler/unboxing.lean.expected.out @@ -0,0 +1,10 @@ +(1, 2, 3, 4, 5, 6.000000, 7.000000, 8) +SpecialValue.thing 1 2 3 4 5 +SpecialValue.nothing +SpecialValue.something (2, 3) +SpecialValue.something _ +[SpecialValue.thing 1 2 3 4 5] +[SpecialValue.nothing] +[SpecialValue.something (2, 3)] +[SpecialValue.something _] +3 diff --git a/tests/lean/4089.lean b/tests/lean/4089.lean index f205995323b6..3b34c9ad107b 100644 --- a/tests/lean/4089.lean +++ b/tests/lean/4089.lean @@ -1,10 +1,18 @@ +structure BoxedProd (α β : Type) where + fst : α + snd : β + +structure BoxedSigma (α : Type) (β : α → Type) where + fst : α + snd : β fst + set_option trace.compiler.ir.reset_reuse true -def f : Nat × Nat → Nat × Nat - | (a, b) => (b, a) +def f : BoxedProd Nat Nat → BoxedProd Nat Nat + | ⟨a, b⟩ => ⟨b, a⟩ -def Sigma.toProd : (_ : α) × β → α × β - | ⟨a, b⟩ => (a, b) +def Sigma.toProd : BoxedSigma α (fun _ => β) → BoxedProd α β + | ⟨a, b⟩ => ⟨a, b⟩ def foo : List (Nat × Nat) → List Nat | [] => [] diff --git a/tests/lean/4089.lean.expected.out b/tests/lean/4089.lean.expected.out index 94f27f05d37d..1fb89ea1cbcb 100644 --- a/tests/lean/4089.lean.expected.out +++ b/tests/lean/4089.lean.expected.out @@ -1,11 +1,11 @@ [Compiler.IR] [reset_reuse] def f (x_1 : obj) : obj := case x_1 : obj of - Prod.mk → - let x_2 : tobj := proj[0] x_1; - let x_3 : tobj := proj[1] x_1; + BoxedProd.mk → + let x_2 : tobj := proj[0, 0] x_1; + let x_3 : tobj := proj[0, 1] x_1; let x_5 : tobj := reset[2] x_1; - let x_4 : obj := reuse x_5 in ctor_0[Prod.mk] x_3 x_2; + let x_4 : obj := reuse x_5 in ctor_0[BoxedProd.mk] x_3 x_2; ret x_4 [Compiler.IR] [reset_reuse] def Sigma.toProd (x_1 : ◾) (x_2 : ◾) (x_3 : obj) : obj := @@ -13,11 +13,11 @@ ret x_4 def Sigma.toProd._redArg (x_1 : obj) : obj := case x_1 : obj of - Sigma.mk → - let x_2 : tobj := proj[0] x_1; - let x_3 : tobj := proj[1] x_1; + BoxedSigma.mk → + let x_2 : tobj := proj[0, 0] x_1; + let x_3 : tobj := proj[0, 1] x_1; let x_5 : tobj := reset[2] x_1; - let x_4 : obj := reuse x_5 in ctor_0[Prod.mk] x_2 x_3; + let x_4 : obj := reuse x_5 in ctor_0[BoxedProd.mk] x_2 x_3; ret x_4 [Compiler.IR] [reset_reuse] def foo (x_1 : tobj) : tobj := @@ -26,13 +26,12 @@ let x_2 : tagged := ctor_0[List.nil]; ret x_2 List.cons → - let x_3 : tobj := proj[0] x_1; - case x_3 : obj of + let x_3 : tobj := proj[1, 0] x_1; + case x_3 : tobj of Prod.mk → - let x_4 : tobj := proj[1] x_1; + let x_4 : tobj := proj[1, 1] x_1; let x_10 : tobj := reset[2] x_1; - let x_5 : tobj := proj[0] x_3; - let x_6 : tobj := proj[1] x_3; + let x_5 : tobj := proj[0, 0] x_3; let x_7 : tobj := 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/computedFieldsCode.lean.expected.out b/tests/lean/computedFieldsCode.lean.expected.out index 55b538059e91..a76e1d00334f 100644 --- a/tests/lean/computedFieldsCode.lean.expected.out +++ b/tests/lean/computedFieldsCode.lean.expected.out @@ -33,15 +33,15 @@ def Exp.ctorElim._redArg (x_1 : tobj) (x_2 : tobj) : tobj := case x_1 : tobj of Exp.var → - let x_3 : u32 := sproj[0, 0] x_1; + let x_3 : u32 := sproj[0, 0, 0] x_1; dec x_1; let x_4 : tobj := box x_3; let x_5 : tobj := app x_2 x_4; ret x_5 Exp.app → - let x_6 : tobj := proj[0] x_1; + let x_6 : tobj := proj[1, 0] x_1; inc x_6; - let x_7 : tobj := proj[1] x_1; + let x_7 : tobj := proj[1, 1] x_1; inc x_7; dec x_1; let x_8 : tobj := app x_2 x_6 x_7; @@ -49,7 +49,7 @@ default → dec x_1; ret x_2 - def Exp.ctorElim._boxed (x_1 : tobj) (x_2 : tobj) (x_3 : tobj) (x_4 : tobj) (x_5 : tobj) : tobj := + def Exp.ctorElim._boxed (x_1 : tagged) (x_2 : tobj) (x_3 : tobj) (x_4 : tagged) (x_5 : tobj) : tobj := let x_6 : tobj := Exp.ctorElim x_1 x_2 x_3 x_4 x_5; dec x_2; ret x_6 @@ -107,16 +107,16 @@ case x_2 : tobj of Exp.var._impl → dec x_4; - let x_10 : u32 := sproj[0, 8] x_2; + let x_10 : u32 := sproj[0, 0, 8] x_2; dec x_2; let x_11 : tobj := box x_10; let x_12 : tobj := app x_3 x_11; ret x_12 Exp.app._impl → dec x_3; - let x_13 : tobj := proj[0] x_2; + let x_13 : tobj := proj[1, 0] x_2; inc x_13; - let x_14 : tobj := proj[1] x_2; + let x_14 : tobj := proj[1, 1] x_2; inc x_14; dec x_2; let x_15 : tobj := app x_4 x_13 x_14; @@ -150,16 +150,16 @@ case x_1 : tobj of Exp.var._impl → dec x_3; - let x_9 : u32 := sproj[0, 8] x_1; + let x_9 : u32 := sproj[0, 0, 8] x_1; dec x_1; let x_10 : tobj := box x_9; let x_11 : tobj := app x_2 x_10; ret x_11 Exp.app._impl → dec x_2; - let x_12 : tobj := proj[0] x_1; + let x_12 : tobj := proj[1, 0] x_1; inc x_12; - let x_13 : tobj := proj[1] x_1; + let x_13 : tobj := proj[1, 1] x_1; inc x_13; dec x_1; let x_14 : tobj := app x_3 x_12 x_13; @@ -189,7 +189,7 @@ dec x_2; inc x_8; ret x_8 - def Exp.casesOn._override._boxed (x_1 : tobj) (x_2 : tobj) (x_3 : tobj) (x_4 : tobj) (x_5 : tobj) (x_6 : tobj) (x_7 : tobj) (x_8 : tobj) (x_9 : tobj) : tobj := + def Exp.casesOn._override._boxed (x_1 : tagged) (x_2 : tobj) (x_3 : tobj) (x_4 : tobj) (x_5 : tobj) (x_6 : tobj) (x_7 : tobj) (x_8 : tobj) (x_9 : tobj) : tobj := let x_10 : tobj := 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; @@ -226,23 +226,23 @@ 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; + sset x_6[1, 2, 0] : u64 := x_5; ret x_6 def Exp.var._override (x_1 : u32) : tobj := 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; + sset x_5[0, 0, 0] : u64 := x_4; + sset x_5[0, 0, 8] : u32 := x_1; ret x_5 def Exp.hash._override (x_1 : @& tobj) : u64 := case x_1 : tobj of Exp.var._impl → - let x_2 : u64 := sproj[0, 0] x_1; + let x_2 : u64 := sproj[0, 0, 0] x_1; ret x_2 Exp.app._impl → - let x_3 : u64 := sproj[2, 0] x_1; + let x_3 : u64 := sproj[1, 2, 0] x_1; ret x_3 default → let x_4 : u64 := 42; @@ -293,12 +293,12 @@ def hash' (x_1 : @& tobj) : tobj := case x_1 : tobj of Exp.var._impl → - let x_2 : u32 := sproj[0, 8] x_1; + let x_2 : u32 := sproj[0, 0, 8] x_1; let x_3 : tobj := UInt32.toNat x_2; ret x_3 Exp.app._impl → - let x_4 : tobj := proj[0] x_1; - let x_5 : tobj := proj[1] x_1; + let x_4 : tobj := proj[1, 0] x_1; + let x_5 : tobj := proj[1, 1] x_1; let x_6 : tobj := hash' x_4; let x_7 : tobj := hash' x_5; let x_8 : tobj := Nat.add x_6 x_7; @@ -316,7 +316,7 @@ def getAppFn (x_1 : @& tobj) : tobj := case x_1 : tobj of Exp.app._impl → - let x_2 : tobj := proj[0] x_1; + let x_2 : tobj := proj[1, 0] x_1; let x_3 : tobj := getAppFn x_2; ret x_3 default → diff --git a/tests/lean/doubleReset.lean.expected.out b/tests/lean/doubleReset.lean.expected.out index f6d2154c546d..a57b7a660a21 100644 --- a/tests/lean/doubleReset.lean.expected.out +++ b/tests/lean/doubleReset.lean.expected.out @@ -17,14 +17,14 @@ ret x_4 Bool.true → let x_6 : tobj := Array.uget ◾ x_4 x_3 ◾; - case x_6 : obj of + case x_6 : tobj of Prod.mk → - let x_7 : tobj := proj[0] x_6; + let x_7 : tobj := proj[0, 0] x_6; let x_19 : tobj := reset[2] x_6; - case x_7 : obj of + case x_7 : tobj of Prod.mk → - let x_8 : tobj := proj[0] x_7; - let x_9 : tobj := proj[1] x_7; + let x_8 : tobj := proj[0, 0] x_7; + let x_9 : tobj := proj[0, 1] x_7; let x_18 : tobj := reset[2] x_7; let x_10 : tagged := 0; let x_11 : obj := Array.uset ◾ x_4 x_3 x_10 ◾; diff --git a/tests/lean/run/4278.lean b/tests/lean/run/4278.lean index fc91c897c2c9..eefff6c5983f 100644 --- a/tests/lean/run/4278.lean +++ b/tests/lean/run/4278.lean @@ -8,7 +8,7 @@ structure S where /-- trace: [Compiler.IR] [result] def get_unboxed (x_1 : obj) : u64 := - let x_2 : u64 := sproj[0, 0] x_1; + let x_2 : u64 := sproj[0, 0, 0] x_1; dec x_1; ret x_2 def get_unboxed._boxed (x_1 : obj) : tobj := diff --git a/tests/lean/run/struct_rc.lean b/tests/lean/run/struct_rc.lean new file mode 100644 index 000000000000..1bad5db69b7c --- /dev/null +++ b/tests/lean/run/struct_rc.lean @@ -0,0 +1,98 @@ +/-! +# Tests the `struct_rc` compiler pass +-/ + +@[extern] +opaque fn (a b : Nat) : Nat + +/-! +The `struct_rc` pass correctly removes `inc` and `dec` instructions when possible. +-/ + +/-- +trace: [Compiler.IR] [struct_rc] + def test1 (x_1 : {tobj, tobj}) : tobj := + let x_2 : tobj := proj[0, 0] x_1; + let x_3 : tobj := proj[0, 1] x_1; + let x_4 : tobj := fn x_2 x_3; + ret x_4 + def test1._boxed (x_1 : obj) : tobj := + let x_2 : {tobj, tobj} := unbox x_1; + inc x_2; + dec x_1; + let x_3 : tobj := test1 x_2; + ret x_3 +-/ +#guard_msgs in +set_option trace.compiler.ir.struct_rc true in +@[noinline] +def test1 (x : Nat × Nat) : Nat := + fn x.1 x.2 + +/-! +The `struct_rc` introduces new `inc` or `dec` instructions when needed. +-/ + +/-- +trace: [Compiler.IR] [struct_rc] + def test2 (x_1 : {tobj, tobj}) : tobj := + let x_2 : tobj := proj[0, 0] x_1; + let x_4 : tobj := proj[0, 1] x_1; + dec x_4; + inc x_2; + let x_3 : tobj := fn x_2 x_2; + ret x_3 + def test2._boxed (x_1 : obj) : tobj := + let x_2 : {tobj, tobj} := unbox x_1; + inc x_2; + dec x_1; + let x_3 : tobj := test2 x_2; + ret x_3 +-/ +#guard_msgs in +set_option trace.compiler.ir.struct_rc true in +def test2 (x : Nat × Nat) : Nat := + fn x.1 x.1 + +/-! +The `struct_rc` pass may introduce double-decrements (which are otherwise impossible). +-/ + +def borrowedFn (x : Nat × Nat) : Nat := + x.1 + x.2 + +/-- +trace: [Compiler.IR] [struct_rc] + def test3 (x_1 : tobj) : tobj := + inc x_1; + let x_2 : {tobj, tobj} := ctor_0[Prod.mk] x_1 x_1; + let x_3 : tobj := borrowedFn x_2; + dec[2] x_1; + ret x_3 +-/ +#guard_msgs in +set_option trace.compiler.ir.struct_rc true in +def test3 (a : Nat) : Nat := + borrowedFn (a, a) + +/-! +The `struct_rc` pass doesn't apply reference counting instructions to persistent values (and +removes existing such RC instructions). +-/ + +@[noinline] +def persistentValue : Nat := 2^100 + +/-- +trace: [Compiler.IR] [struct_rc] + def test4 (x_1 : tagged) : tobj := + let x_2 : tobj := persistentValue; + let x_3 : {tobj, tobj} := ctor_0[Prod.mk] x_2 x_2; + let x_4 : tobj := borrowedFn x_3; + ret x_4 +-/ +#guard_msgs in +set_option trace.compiler.ir.struct_rc true in +set_option compiler.extract_closed false in +def test4 (_ : Unit) : Nat := + borrowedFn (persistentValue, persistentValue) diff --git a/tests/lean/unboxStruct.lean.expected.out b/tests/lean/unboxStruct.lean.expected.out index 82d0e2989c53..921c128749f0 100644 --- a/tests/lean/unboxStruct.lean.expected.out +++ b/tests/lean/unboxStruct.lean.expected.out @@ -1,9 +1,13 @@ [Compiler.IR] [result] - def test2 (x_1 : u32) (x_2 : void) : obj := + def test2 (x_1 : u32) (x_2 : void) : ({tagged} | {tobj}) := let x_3 : obj := foo x_1 ◾; - ret x_3 + let x_4 : ({tagged} | {tobj}) := unbox x_3; + inc x_4; + dec x_3; + ret x_4 def test2._boxed (x_1 : tobj) (x_2 : tagged) : obj := let x_3 : u32 := unbox x_1; dec x_1; - let x_4 : obj := test2 x_3 x_2; - ret x_4 + let x_4 : ({tagged} | {tobj}) := test2 x_3 x_2; + let x_5 : obj := box x_4; + ret x_5 diff --git a/tests/lean/updateExprIssue.lean.expected.out b/tests/lean/updateExprIssue.lean.expected.out index 6bc808fe89db..3e3ffe443a81 100644 --- a/tests/lean/updateExprIssue.lean.expected.out +++ b/tests/lean/updateExprIssue.lean.expected.out @@ -2,9 +2,9 @@ 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; + let x_3 : u64 := sproj[5, 2, 0] x_1; + let x_4 : obj := proj[5, 0] x_1; + let x_5 : obj := proj[5, 1] x_1; block_6 (x_7 : u8) := case x_7 : u8 of Bool.false →